:: CIRCCOMB semantic presentation

definition
let c1 be ManySortedSign ;
mode Gate is Element of the OperSymbols of a1;
end;

registration
let c1 be set ;
let c2 be Function;
cluster a1 --> a2 -> Function-yielding ;
coherence
c1 --> c2 is Function-yielding
proof end;
end;

E1: now
let c1 be Nat;
let c2 be non empty set ;
thus product ((Seg c1) --> c2) = product (c1 |-> c2) by FINSEQ_2:def 2
.= c1 -tuples_on c2 by FUNCT_6:9 ;
end;

E2: now
let c1 be non void ManySortedSign ;
let c2 be Gate of c1;
not the OperSymbols of c1 is empty by MSUALG_1:def 5;
hence c2 in the OperSymbols of c1 ;
end;

registration
let c1, c2 be non-empty Function;
cluster a1 +* a2 -> non-empty ;
coherence
c1 +* c2 is non-empty
proof end;
end;

definition
let c1, c2 be set ;
let c3 be ManySortedSet of c1;
let c4 be ManySortedSet of c2;
redefine func +* as c3 +* c4 -> ManySortedSet of a1 \/ a2;
coherence
c3 +* c4 is ManySortedSet of c1 \/ c2
proof end;
end;

theorem Th1: :: CIRCCOMB:1
canceled;

theorem Th2: :: CIRCCOMB:2
for b1, b2, b3 being Function holds
( rng b3 c= dom b1 & rng b3 c= dom b2 & b1 tolerates b2 implies b1 * b3 = b2 * b3 )
proof end;

theorem Th3: :: CIRCCOMB:3
for b1, b2 being set
for b3 being ManySortedSet of b1
for b4 being ManySortedSet of b2 holds
( b3 c= b4 implies b3 # c= b4 # )
proof end;

theorem Th4: :: CIRCCOMB:4
for b1, b2, b3, b4 being set holds
( b1 --> b3 tolerates b2 --> b4 iff ( b3 = b4 or b1 misses b2 ) )
proof end;

theorem Th5: :: CIRCCOMB:5
for b1, b2, b3 being Function holds
( b1 tolerates b2 & b2 tolerates b3 & b3 tolerates b1 implies b1 +* b2 tolerates b3 )
proof end;

theorem Th6: :: CIRCCOMB:6
for b1 being set
for b2 being non empty set
for b3 being FinSequence of b1 holds ((b1 --> b2) # ) . b3 = (len b3) -tuples_on b2
proof end;

definition
let c1 be set ;
let c2, c3 be V3 ManySortedSet of c1;
let c4 be set ;
let c5, c6 be V3 ManySortedSet of c4;
let c7 be ManySortedFunction of c2,c3;
let c8 be ManySortedFunction of c5,c6;
redefine func +* as c7 +* c8 -> ManySortedFunction of a2 +* a5,a3 +* a6;
coherence
c7 +* c8 is ManySortedFunction of c2 +* c5,c3 +* c6
proof end;
end;

definition
let c1, c2 be ManySortedSign ;
pred c1 tolerates c2 means :: CIRCCOMB:def 1
( the Arity of a1 tolerates the Arity of a2 & the ResultSort of a1 tolerates the ResultSort of a2 );
reflexivity
for b1 being ManySortedSign holds
( the Arity of b1 tolerates the Arity of b1 & the ResultSort of b1 tolerates the ResultSort of b1 )
;
symmetry
for b1, b2 being ManySortedSign holds
( the Arity of b1 tolerates the Arity of b2 & the ResultSort of b1 tolerates the ResultSort of b2 implies ( the Arity of b2 tolerates the Arity of b1 & the ResultSort of b2 tolerates the ResultSort of b1 ) )
;
end;

:: deftheorem Def1 defines tolerates CIRCCOMB:def 1 :
for b1, b2 being ManySortedSign holds
( b1 tolerates b2 iff ( the Arity of b1 tolerates the Arity of b2 & the ResultSort of b1 tolerates the ResultSort of b2 ) );

definition
let c1, c2 be non empty ManySortedSign ;
func c1 +* c2 -> non empty strict ManySortedSign means :Def2: :: CIRCCOMB:def 2
( the carrier of a3 = the carrier of a1 \/ the carrier of a2 & the OperSymbols of a3 = the OperSymbols of a1 \/ the OperSymbols of a2 & the Arity of a3 = the Arity of a1 +* the Arity of a2 & the ResultSort of a3 = the ResultSort of a1 +* the ResultSort of a2 );
existence
ex b1 being non empty strict ManySortedSign st
( the carrier of b1 = the carrier of c1 \/ the carrier of c2 & the OperSymbols of b1 = the OperSymbols of c1 \/ the OperSymbols of c2 & the Arity of b1 = the Arity of c1 +* the Arity of c2 & the ResultSort of b1 = the ResultSort of c1 +* the ResultSort of c2 )
proof end;
uniqueness
for b1, b2 being non empty strict ManySortedSign holds
( the carrier of b1 = the carrier of c1 \/ the carrier of c2 & the OperSymbols of b1 = the OperSymbols of c1 \/ the OperSymbols of c2 & the Arity of b1 = the Arity of c1 +* the Arity of c2 & the ResultSort of b1 = the ResultSort of c1 +* the ResultSort of c2 & the carrier of b2 = the carrier of c1 \/ the carrier of c2 & the OperSymbols of b2 = the OperSymbols of c1 \/ the OperSymbols of c2 & the Arity of b2 = the Arity of c1 +* the Arity of c2 & the ResultSort of b2 = the ResultSort of c1 +* the ResultSort of c2 implies b1 = b2 )
;
end;

:: deftheorem Def2 defines +* CIRCCOMB:def 2 :
for b1, b2 being non empty ManySortedSign
for b3 being non empty strict ManySortedSign holds
( b3 = b1 +* b2 iff ( the carrier of b3 = the carrier of b1 \/ the carrier of b2 & the OperSymbols of b3 = the OperSymbols of b1 \/ the OperSymbols of b2 & the Arity of b3 = the Arity of b1 +* the Arity of b2 & the ResultSort of b3 = the ResultSort of b1 +* the ResultSort of b2 ) );

theorem Th7: :: CIRCCOMB:7
for b1, b2, b3 being non empty ManySortedSign holds
( b1 tolerates b2 & b2 tolerates b3 & b3 tolerates b1 implies b1 +* b2 tolerates b3 )
proof end;

theorem Th8: :: CIRCCOMB:8
for b1 being non empty ManySortedSign holds b1 +* b1 = ManySortedSign(# the carrier of b1,the OperSymbols of b1,the Arity of b1,the ResultSort of b1 #)
proof end;

theorem Th9: :: CIRCCOMB:9
for b1, b2 being non empty ManySortedSign holds
( b1 tolerates b2 implies b1 +* b2 = b2 +* b1 )
proof end;

theorem Th10: :: CIRCCOMB:10
for b1, b2, b3 being non empty ManySortedSign holds (b1 +* b2) +* b3 = b1 +* (b2 +* b3)
proof end;

registration
cluster one-to-one set ;
existence
ex b1 being Function st b1 is one-to-one
proof end;
end;

theorem Th11: :: CIRCCOMB:11
for b1 being one-to-one Function
for b2, b3 being non empty Circuit-like ManySortedSign holds
( the ResultSort of b2 c= b1 & the ResultSort of b3 c= b1 implies b2 +* b3 is Circuit-like )
proof end;

theorem Th12: :: CIRCCOMB:12
for b1, b2 being non empty Circuit-like ManySortedSign holds
( InnerVertices b1 misses InnerVertices b2 implies b1 +* b2 is Circuit-like )
proof end;

theorem Th13: :: CIRCCOMB:13
for b1, b2 being non empty ManySortedSign holds
not ( not ( b1 is void & b2 is void ) & b1 +* b2 is void )
proof end;

theorem Th14: :: CIRCCOMB:14
for b1, b2 being non empty finite ManySortedSign holds b1 +* b2 is finite
proof end;

registration
let c1 be non empty non void ManySortedSign ;
let c2 be non empty ManySortedSign ;
cluster a1 +* a2 -> non empty strict non void ;
coherence
not c1 +* c2 is void
by Th13;
cluster a2 +* a1 -> non empty strict non void ;
coherence
not c2 +* c1 is void
by Th13;
end;

theorem Th15: :: CIRCCOMB:15
for b1, b2 being non empty ManySortedSign holds
( b1 tolerates b2 implies ( InnerVertices (b1 +* b2) = (InnerVertices b1) \/ (InnerVertices b2) & InputVertices (b1 +* b2) c= (InputVertices b1) \/ (InputVertices b2) ) )
proof end;

theorem Th16: :: CIRCCOMB:16
for b1, b2 being non empty ManySortedSign
for b3 being Vertex of b2 holds
( b3 in InputVertices (b1 +* b2) implies b3 in InputVertices b2 )
proof end;

theorem Th17: :: CIRCCOMB:17
for b1, b2 being non empty ManySortedSign holds
( b1 tolerates b2 implies for b3 being Vertex of b1 holds
( b3 in InputVertices (b1 +* b2) implies b3 in InputVertices b1 ) )
proof end;

theorem Th18: :: CIRCCOMB:18
for b1 being non empty ManySortedSign
for b2 being non empty non void ManySortedSign
for b3 being OperSymbol of b2
for b4 being OperSymbol of (b1 +* b2) holds
( b3 = b4 implies ( the_arity_of b4 = the_arity_of b3 & the_result_sort_of b4 = the_result_sort_of b3 ) )
proof end;

theorem Th19: :: CIRCCOMB:19
for b1 being non empty ManySortedSign
for b2, b3 being non empty non void Circuit-like ManySortedSign holds
( b3 = b1 +* b2 implies for b4 being Vertex of b2 holds
( b4 in InnerVertices b2 implies for b5 being Vertex of b3 holds
( b4 = b5 implies ( b5 in InnerVertices b3 & action_at b5 = action_at b4 ) ) ) )
proof end;

theorem Th20: :: CIRCCOMB:20
for b1 being non empty non void ManySortedSign
for b2 being non empty ManySortedSign holds
( b1 tolerates b2 implies for b3 being OperSymbol of b1
for b4 being OperSymbol of (b1 +* b2) holds
( b3 = b4 implies ( the_arity_of b4 = the_arity_of b3 & the_result_sort_of b4 = the_result_sort_of b3 ) ) )
proof end;

theorem Th21: :: CIRCCOMB:21
for b1, b2 being non empty non void Circuit-like ManySortedSign
for b3 being non empty ManySortedSign holds
( b1 tolerates b3 & b2 = b1 +* b3 implies for b4 being Vertex of b1 holds
( b4 in InnerVertices b1 implies for b5 being Vertex of b2 holds
( b4 = b5 implies ( b5 in InnerVertices b2 & action_at b5 = action_at b4 ) ) ) )
proof end;

definition
let c1, c2 be non empty ManySortedSign ;
let c3 be MSAlgebra of c1;
let c4 be MSAlgebra of c2;
pred c3 tolerates c4 means :Def3: :: CIRCCOMB:def 3
( a1 tolerates a2 & the Sorts of a3 tolerates the Sorts of a4 & the Charact of a3 tolerates the Charact of a4 );
end;

:: deftheorem Def3 defines tolerates CIRCCOMB:def 3 :
for b1, b2 being non empty ManySortedSign
for b3 being MSAlgebra of b1
for b4 being MSAlgebra of b2 holds
( b3 tolerates b4 iff ( b1 tolerates b2 & the Sorts of b3 tolerates the Sorts of b4 & the Charact of b3 tolerates the Charact of b4 ) );

definition
let c1, c2 be non empty ManySortedSign ;
let c3 be non-empty MSAlgebra of c1;
let c4 be non-empty MSAlgebra of c2;
assume E21: the Sorts of c3 tolerates the Sorts of c4 ;
func c3 +* c4 -> strict non-empty MSAlgebra of a1 +* a2 means :Def4: :: CIRCCOMB:def 4
( the Sorts of a5 = the Sorts of a3 +* the Sorts of a4 & the Charact of a5 = the Charact of a3 +* the Charact of a4 );
uniqueness
for b1, b2 being strict non-empty MSAlgebra of c1 +* c2 holds
( the Sorts of b1 = the Sorts of c3 +* the Sorts of c4 & the Charact of b1 = the Charact of c3 +* the Charact of c4 & the Sorts of b2 = the Sorts of c3 +* the Sorts of c4 & the Charact of b2 = the Charact of c3 +* the Charact of c4 implies b1 = b2 )
;
existence
ex b1 being strict non-empty MSAlgebra of c1 +* c2 st
( the Sorts of b1 = the Sorts of c3 +* the Sorts of c4 & the Charact of b1 = the Charact of c3 +* the Charact of c4 )
proof end;
end;

:: deftheorem Def4 defines +* CIRCCOMB:def 4 :
for b1, b2 being non empty ManySortedSign
for b3 being non-empty MSAlgebra of b1
for b4 being non-empty MSAlgebra of b2 holds
( the Sorts of b3 tolerates the Sorts of b4 implies for b5 being strict non-empty MSAlgebra of b1 +* b2 holds
( b5 = b3 +* b4 iff ( the Sorts of b5 = the Sorts of b3 +* the Sorts of b4 & the Charact of b5 = the Charact of b3 +* the Charact of b4 ) ) );

theorem Th22: :: CIRCCOMB:22
for b1 being non empty non void ManySortedSign
for b2 being MSAlgebra of b1 holds b2 tolerates b2
proof end;

theorem Th23: :: CIRCCOMB:23
for b1, b2 being non empty non void ManySortedSign
for b3 being MSAlgebra of b1
for b4 being MSAlgebra of b2 holds
( b3 tolerates b4 implies b4 tolerates b3 )
proof end;

theorem Th24: :: CIRCCOMB:24
for b1, b2, b3 being non empty ManySortedSign
for b4 being non-empty MSAlgebra of b1
for b5 being non-empty MSAlgebra of b2
for b6 being MSAlgebra of b3 holds
( b4 tolerates b5 & b5 tolerates b6 & b6 tolerates b4 implies b4 +* b5 tolerates b6 )
proof end;

theorem Th25: :: CIRCCOMB:25
for b1 being non empty strict ManySortedSign
for b2 being non-empty MSAlgebra of b1 holds b2 +* b2 = MSAlgebra(# the Sorts of b2,the Charact of b2 #)
proof end;

theorem Th26: :: CIRCCOMB:26
for b1, b2 being non empty ManySortedSign
for b3 being non-empty MSAlgebra of b1
for b4 being non-empty MSAlgebra of b2 holds
( b3 tolerates b4 implies b3 +* b4 = b4 +* b3 )
proof end;

theorem Th27: :: CIRCCOMB:27
for b1, b2, b3 being non empty ManySortedSign
for b4 being non-empty MSAlgebra of b1
for b5 being non-empty MSAlgebra of b2
for b6 being non-empty MSAlgebra of b3 holds
( the Sorts of b4 tolerates the Sorts of b5 & the Sorts of b5 tolerates the Sorts of b6 & the Sorts of b6 tolerates the Sorts of b4 implies (b4 +* b5) +* b6 = b4 +* (b5 +* b6) )
proof end;

theorem Th28: :: CIRCCOMB:28
for b1, b2 being non empty ManySortedSign
for b3 being non-empty locally-finite MSAlgebra of b1
for b4 being non-empty locally-finite MSAlgebra of b2 holds
( the Sorts of b3 tolerates the Sorts of b4 implies b3 +* b4 is locally-finite )
proof end;

theorem Th29: :: CIRCCOMB:29
for b1, b2 being non-empty Function
for b3 being Element of product b1
for b4 being Element of product b2 holds b3 +* b4 in product (b1 +* b2)
proof end;

theorem Th30: :: CIRCCOMB:30
for b1, b2 being non-empty Function
for b3 being Element of product (b1 +* b2) holds b3 | (dom b2) in product b2
proof end;

theorem Th31: :: CIRCCOMB:31
for b1, b2 being non-empty Function holds
( b1 tolerates b2 implies for b3 being Element of product (b1 +* b2) holds b3 | (dom b1) in product b1 )
proof end;

theorem Th32: :: CIRCCOMB:32
for b1, b2 being non empty ManySortedSign
for b3 being non-empty MSAlgebra of b1
for b4 being Element of product the Sorts of b3
for b5 being non-empty MSAlgebra of b2
for b6 being Element of product the Sorts of b5 holds
( the Sorts of b3 tolerates the Sorts of b5 implies b4 +* b6 in product the Sorts of (b3 +* b5) )
proof end;

theorem Th33: :: CIRCCOMB:33
for b1, b2 being non empty ManySortedSign
for b3 being non-empty MSAlgebra of b1
for b4 being non-empty MSAlgebra of b2 holds
( the Sorts of b3 tolerates the Sorts of b4 implies for b5 being Element of product the Sorts of (b3 +* b4) holds
( b5 | the carrier of b1 in product the Sorts of b3 & b5 | the carrier of b2 in product the Sorts of b4 ) )
proof end;

theorem Th34: :: CIRCCOMB:34
for b1, b2 being non empty non void ManySortedSign
for b3 being non-empty MSAlgebra of b1
for b4 being non-empty MSAlgebra of b2 holds
( the Sorts of b3 tolerates the Sorts of b4 implies for b5 being OperSymbol of (b1 +* b2)
for b6 being OperSymbol of b2 holds
( b5 = b6 implies Den b5,(b3 +* b4) = Den b6,b4 ) )
proof end;

theorem Th35: :: CIRCCOMB:35
for b1, b2 being non empty non void ManySortedSign
for b3 being non-empty MSAlgebra of b1
for b4 being non-empty MSAlgebra of b2 holds
( the Sorts of b3 tolerates the Sorts of b4 & the Charact of b3 tolerates the Charact of b4 implies for b5 being OperSymbol of (b1 +* b2)
for b6 being OperSymbol of b1 holds
( b5 = b6 implies Den b5,(b3 +* b4) = Den b6,b3 ) )
proof end;

theorem Th36: :: CIRCCOMB:36
for b1, b2, b3 being non empty non void Circuit-like ManySortedSign holds
( b3 = b1 +* b2 implies for b4 being non-empty Circuit of b1
for b5 being non-empty Circuit of b2
for b6 being non-empty Circuit of b3
for b7 being State of b6
for b8 being State of b5 holds
( b8 = b7 | the carrier of b2 implies for b9 being Gate of b3
for b10 being Gate of b2 holds
( b9 = b10 implies b9 depends_on_in b7 = b10 depends_on_in b8 ) ) )
proof end;

theorem Th37: :: CIRCCOMB:37
for b1, b2, b3 being non empty non void Circuit-like ManySortedSign holds
( b3 = b1 +* b2 & b1 tolerates b2 implies for b4 being non-empty Circuit of b1
for b5 being non-empty Circuit of b2
for b6 being non-empty Circuit of b3
for b7 being State of b6
for b8 being State of b4 holds
( b8 = b7 | the carrier of b1 implies for b9 being Gate of b3
for b10 being Gate of b1 holds
( b9 = b10 implies b9 depends_on_in b7 = b10 depends_on_in b8 ) ) )
proof end;

theorem Th38: :: CIRCCOMB:38
for b1, b2, b3 being non empty non void Circuit-like ManySortedSign holds
( b3 = b1 +* b2 implies for b4 being non-empty Circuit of b1
for b5 being non-empty Circuit of b2
for b6 being non-empty Circuit of b3 holds
( b4 tolerates b5 & b6 = b4 +* b5 implies for b7 being State of b6
for b8 being Vertex of b3 holds
( ( for b9 being State of b4 holds
( b9 = b7 | the carrier of b1 & ( b8 in InnerVertices b1 or ( b8 in the carrier of b1 & b8 in InputVertices b3 ) ) implies (Following b7) . b8 = (Following b9) . b8 ) ) & ( for b9 being State of b5 holds
( b9 = b7 | the carrier of b2 & ( b8 in InnerVertices b2 or ( b8 in the carrier of b2 & b8 in InputVertices b3 ) ) implies (Following b7) . b8 = (Following b9) . b8 ) ) ) ) )
proof end;

theorem Th39: :: CIRCCOMB:39
for b1, b2, b3 being non empty non void Circuit-like ManySortedSign holds
( InnerVertices b1 misses InputVertices b2 & b3 = b1 +* b2 implies for b4 being non-empty Circuit of b1
for b5 being non-empty Circuit of b2
for b6 being non-empty Circuit of b3 holds
( b4 tolerates b5 & b6 = b4 +* b5 implies for b7 being State of b6
for b8 being State of b4
for b9 being State of b5 holds
( b8 = b7 | the carrier of b1 & b9 = b7 | the carrier of b2 implies Following b7 = (Following b8) +* (Following b9) ) ) )
proof end;

theorem Th40: :: CIRCCOMB:40
for b1, b2, b3 being non empty non void Circuit-like ManySortedSign holds
( InnerVertices b2 misses InputVertices b1 & b3 = b1 +* b2 implies for b4 being non-empty Circuit of b1
for b5 being non-empty Circuit of b2
for b6 being non-empty Circuit of b3 holds
( b4 tolerates b5 & b6 = b4 +* b5 implies for b7 being State of b6
for b8 being State of b4
for b9 being State of b5 holds
( b8 = b7 | the carrier of b1 & b9 = b7 | the carrier of b2 implies Following b7 = (Following b9) +* (Following b8) ) ) )
proof end;

theorem Th41: :: CIRCCOMB:41
for b1, b2, b3 being non empty non void Circuit-like ManySortedSign holds
( InputVertices b1 c= InputVertices b2 & b3 = b1 +* b2 implies for b4 being non-empty Circuit of b1
for b5 being non-empty Circuit of b2
for b6 being non-empty Circuit of b3 holds
( b4 tolerates b5 & b6 = b4 +* b5 implies for b7 being State of b6
for b8 being State of b4
for b9 being State of b5 holds
( b8 = b7 | the carrier of b1 & b9 = b7 | the carrier of b2 implies Following b7 = (Following b9) +* (Following b8) ) ) )
proof end;

theorem Th42: :: CIRCCOMB:42
for b1, b2, b3 being non empty non void Circuit-like ManySortedSign holds
( InputVertices b2 c= InputVertices b1 & b3 = b1 +* b2 implies for b4 being non-empty Circuit of b1
for b5 being non-empty Circuit of b2
for b6 being non-empty Circuit of b3 holds
( b4 tolerates b5 & b6 = b4 +* b5 implies for b7 being State of b6
for b8 being State of b4
for b9 being State of b5 holds
( b8 = b7 | the carrier of b1 & b9 = b7 | the carrier of b2 implies Following b7 = (Following b8) +* (Following b9) ) ) )
proof end;

definition
let c1, c2 be non empty set ;
let c3 be Element of c1;
redefine func --> as c2 --> c3 -> Function of a2,a1;
coherence
c2 --> c3 is Function of c2,c1
by FUNCOP_1:57;
end;

definition
let c1 be set ;
let c2 be FinSequence;
let c3 be set ;
func 1GateCircStr c2,c1,c3 -> strict non void ManySortedSign means :Def5: :: CIRCCOMB:def 5
( the carrier of a4 = (rng a2) \/ {a3} & the OperSymbols of a4 = {[a2,a1]} & the Arity of a4 . [a2,a1] = a2 & the ResultSort of a4 . [a2,a1] = a3 );
existence
ex b1 being strict non void ManySortedSign st
( the carrier of b1 = (rng c2) \/ {c3} & the OperSymbols of b1 = {[c2,c1]} & the Arity of b1 . [c2,c1] = c2 & the ResultSort of b1 . [c2,c1] = c3 )
proof end;
uniqueness
for b1, b2 being strict non void ManySortedSign holds
( the carrier of b1 = (rng c2) \/ {c3} & the OperSymbols of b1 = {[c2,c1]} & the Arity of b1 . [c2,c1] = c2 & the ResultSort of b1 . [c2,c1] = c3 & the carrier of b2 = (rng c2) \/ {c3} & the OperSymbols of b2 = {[c2,c1]} & the Arity of b2 . [c2,c1] = c2 & the ResultSort of b2 . [c2,c1] = c3 implies b1 = b2 )
proof end;
end;

:: deftheorem Def5 defines 1GateCircStr CIRCCOMB:def 5 :
for b1 being set
for b2 being FinSequence
for b3 being set
for b4 being strict non void ManySortedSign holds
( b4 = 1GateCircStr b2,b1,b3 iff ( the carrier of b4 = (rng b2) \/ {b3} & the OperSymbols of b4 = {[b2,b1]} & the Arity of b4 . [b2,b1] = b2 & the ResultSort of b4 . [b2,b1] = b3 ) );

registration
let c1 be set ;
let c2 be FinSequence;
let c3 be set ;
cluster 1GateCircStr a2,a1,a3 -> non empty strict non void ;
coherence
not 1GateCircStr c2,c1,c3 is empty
proof end;
end;

theorem Th43: :: CIRCCOMB:43
for b1, b2 being set
for b3 being FinSequence holds
( the Arity of (1GateCircStr b3,b1,b2) = {[b3,b1]} --> b3 & the ResultSort of (1GateCircStr b3,b1,b2) = {[b3,b1]} --> b2 )
proof end;

theorem Th44: :: CIRCCOMB:44
for b1, b2 being set
for b3 being FinSequence
for b4 being Gate of (1GateCircStr b3,b1,b2) holds
( b4 = [b3,b1] & the_arity_of b4 = b3 & the_result_sort_of b4 = b2 )
proof end;

theorem Th45: :: CIRCCOMB:45
for b1, b2 being set
for b3 being FinSequence holds
( InputVertices (1GateCircStr b3,b1,b2) = (rng b3) \ {b2} & InnerVertices (1GateCircStr b3,b1,b2) = {b2} )
proof end;

definition
let c1 be set ;
let c2 be FinSequence;
func 1GateCircStr c2,c1 -> strict non void ManySortedSign means :Def6: :: CIRCCOMB:def 6
( the carrier of a3 = (rng a2) \/ {[a2,a1]} & the OperSymbols of a3 = {[a2,a1]} & the Arity of a3 . [a2,a1] = a2 & the ResultSort of a3 . [a2,a1] = [a2,a1] );
existence
ex b1 being strict non void ManySortedSign st
( the carrier of b1 = (rng c2) \/ {[c2,c1]} & the OperSymbols of b1 = {[c2,c1]} & the Arity of b1 . [c2,c1] = c2 & the ResultSort of b1 . [c2,c1] = [c2,c1] )
proof end;
uniqueness
for b1, b2 being strict non void ManySortedSign holds
( the carrier of b1 = (rng c2) \/ {[c2,c1]} & the OperSymbols of b1 = {[c2,c1]} & the Arity of b1 . [c2,c1] = c2 & the ResultSort of b1 . [c2,c1] = [c2,c1] & the carrier of b2 = (rng c2) \/ {[c2,c1]} & the OperSymbols of b2 = {[c2,c1]} & the Arity of b2 . [c2,c1] = c2 & the ResultSort of b2 . [c2,c1] = [c2,c1] implies b1 = b2 )
proof end;
end;

:: deftheorem Def6 defines 1GateCircStr CIRCCOMB:def 6 :
for b1 being set
for b2 being FinSequence
for b3 being strict non void ManySortedSign holds
( b3 = 1GateCircStr b2,b1 iff ( the carrier of b3 = (rng b2) \/ {[b2,b1]} & the OperSymbols of b3 = {[b2,b1]} & the Arity of b3 . [b2,b1] = b2 & the ResultSort of b3 . [b2,b1] = [b2,b1] ) );

registration
let c1 be set ;
let c2 be FinSequence;
cluster 1GateCircStr a2,a1 -> non empty strict non void ;
coherence
not 1GateCircStr c2,c1 is empty
proof end;
end;

theorem Th46: :: CIRCCOMB:46
for b1 being set
for b2 being FinSequence holds 1GateCircStr b2,b1 = 1GateCircStr b2,b1,[b2,b1]
proof end;

theorem Th47: :: CIRCCOMB:47
for b1 being set
for b2 being FinSequence holds
( the Arity of (1GateCircStr b2,b1) = {[b2,b1]} --> b2 & the ResultSort of (1GateCircStr b2,b1) = {[b2,b1]} --> [b2,b1] )
proof end;

theorem Th48: :: CIRCCOMB:48
for b1 being set
for b2 being FinSequence
for b3 being Gate of (1GateCircStr b2,b1) holds
( b3 = [b2,b1] & the_arity_of b3 = b2 & the_result_sort_of b3 = b3 )
proof end;

theorem Th49: :: CIRCCOMB:49
for b1 being set
for b2 being FinSequence holds
( InputVertices (1GateCircStr b2,b1) = rng b2 & InnerVertices (1GateCircStr b2,b1) = {[b2,b1]} )
proof end;

theorem Th50: :: CIRCCOMB:50
for b1 being set
for b2 being FinSequence
for b3 being set holds
( b3 in rng b2 implies the_rank_of b3 in the_rank_of [b2,b1] )
proof end;

theorem Th51: :: CIRCCOMB:51
for b1 being set
for b2, b3 being FinSequence holds 1GateCircStr b2,b1 tolerates 1GateCircStr b3,b1
proof end;

definition
let c1 be ManySortedSign ;
attr a1 is unsplit means :Def7: :: CIRCCOMB:def 7
the ResultSort of a1 = id the OperSymbols of a1;
attr a1 is gate`1=arity means :Def8: :: CIRCCOMB:def 8
for b1 being set holds
( b1 in the OperSymbols of a1 implies b1 = [(the Arity of a1 . b1),(b1 `2 )] );
attr a1 is gate`2isBoolean means :Def9: :: CIRCCOMB:def 9
for b1 being set holds
( b1 in the OperSymbols of a1 implies for b2 being FinSequence holds
not ( b2 = the Arity of a1 . b1 & ( for b3 being Function of (len b2) -tuples_on BOOLEAN , BOOLEAN holds
not b1 = [(b1 `1 ),b3] ) ) );
end;

:: deftheorem Def7 defines unsplit CIRCCOMB:def 7 :
for b1 being ManySortedSign holds
( b1 is unsplit iff the ResultSort of b1 = id the OperSymbols of b1 );

:: deftheorem Def8 defines gate`1=arity CIRCCOMB:def 8 :
for b1 being ManySortedSign holds
( b1 is gate`1=arity iff for b2 being set holds
( b2 in the OperSymbols of b1 implies b2 = [(the Arity of b1 . b2),(b2 `2 )] ) );

:: deftheorem Def9 defines gate`2isBoolean CIRCCOMB:def 9 :
for b1 being ManySortedSign holds
( b1 is gate`2isBoolean iff for b2 being set holds
( b2 in the OperSymbols of b1 implies for b3 being FinSequence holds
not ( b3 = the Arity of b1 . b2 & ( for b4 being Function of (len b3) -tuples_on BOOLEAN , BOOLEAN holds
not b2 = [(b2 `1 ),b4] ) ) ) );

definition
let c1 be non empty ManySortedSign ;
let c2 be MSAlgebra of c1;
attr a2 is gate`2=den means :Def10: :: CIRCCOMB:def 10
for b1 being set holds
( b1 in the OperSymbols of a1 implies b1 = [(b1 `1 ),(the Charact of a2 . b1)] );
end;

:: deftheorem Def10 defines gate`2=den CIRCCOMB:def 10 :
for b1 being non empty ManySortedSign
for b2 being MSAlgebra of b1 holds
( b2 is gate`2=den iff for b3 being set holds
( b3 in the OperSymbols of b1 implies b3 = [(b3 `1 ),(the Charact of b2 . b3)] ) );

definition
let c1 be non empty ManySortedSign ;
attr a1 is gate`2=den means :: CIRCCOMB:def 11
ex b1 being MSAlgebra of a1 st b1 is gate`2=den;
end;

:: deftheorem Def11 defines gate`2=den CIRCCOMB:def 11 :
for b1 being non empty ManySortedSign holds
( b1 is gate`2=den iff ex b2 being MSAlgebra of b1 st b2 is gate`2=den );

scheme :: CIRCCOMB:sch 1
s1{ F1() -> set , F2() -> set , F3() -> Function of F1(),F2(), F4( set , set ) -> set } :
ex b1 being ManySortedSet of F1() st
for b2 being set
for b3 being Element of F2() holds
( b2 in F1() & b3 = F3() . b2 implies b1 . b2 = F4(b2,b3) )
proof end;

scheme :: CIRCCOMB:sch 2
s2{ F1() -> non empty ManySortedSign , F2( set , set ) -> set } :
ex b1 being strict MSAlgebra of F1() st
( the Sorts of b1 = the carrier of F1() --> BOOLEAN & ( for b2 being set
for b3 being Element of the carrier of F1() * holds
( b2 in the OperSymbols of F1() & b3 = the Arity of F1() . b2 implies the Charact of b1 . b2 = F2(b2,b3) ) ) )
provided
E45: for b1 being set
for b2 being Element of the carrier of F1() * holds
( b1 in the OperSymbols of F1() & b2 = the Arity of F1() . b1 implies F2(b1,b2) is Function of (len b2) -tuples_on BOOLEAN , BOOLEAN )
proof end;

registration
cluster non empty gate`2isBoolean -> non empty gate`2=den ManySortedSign ;
coherence
for b1 being non empty ManySortedSign holds
( b1 is gate`2isBoolean implies b1 is gate`2=den )
proof end;
end;

theorem Th52: :: CIRCCOMB:52
for b1 being non empty ManySortedSign holds
( b1 is unsplit iff for b2 being set holds
( b2 in the OperSymbols of b1 implies the ResultSort of b1 . b2 = b2 ) )
proof end;

theorem Th53: :: CIRCCOMB:53
for b1 being non empty ManySortedSign holds
( b1 is unsplit implies the OperSymbols of b1 c= the carrier of b1 )
proof end;

registration
cluster non empty unsplit -> non empty Circuit-like ManySortedSign ;
coherence
for b1 being non empty ManySortedSign holds
( b1 is unsplit implies b1 is Circuit-like )
proof end;
end;

theorem Th54: :: CIRCCOMB:54
for b1 being set
for b2 being FinSequence holds
( 1GateCircStr b2,b1 is unsplit & 1GateCircStr b2,b1 is gate`1=arity )
proof end;

registration
let c1 be set ;
let c2 be FinSequence;
cluster 1GateCircStr a2,a1 -> non empty strict non void Circuit-like unsplit gate`1=arity ;
coherence
( 1GateCircStr c2,c1 is unsplit & 1GateCircStr c2,c1 is gate`1=arity )
by Th54;
end;

registration
cluster non empty strict non void Circuit-like unsplit gate`1=arity ManySortedSign ;
existence
ex b1 being ManySortedSign st
( b1 is unsplit & b1 is gate`1=arity & not b1 is void & b1 is strict & not b1 is empty )
proof end;
end;

theorem Th55: :: CIRCCOMB:55
for b1, b2 being non empty unsplit gate`1=arity ManySortedSign holds b1 tolerates b2
proof end;

theorem Th56: :: CIRCCOMB:56
for b1, b2 being non empty ManySortedSign
for b3 being MSAlgebra of b1
for b4 being MSAlgebra of b2 holds
( b3 is gate`2=den & b4 is gate`2=den implies the Charact of b3 tolerates the Charact of b4 )
proof end;

theorem Th57: :: CIRCCOMB:57
for b1, b2 being non empty unsplit ManySortedSign holds b1 +* b2 is unsplit
proof end;

registration
let c1, c2 be non empty unsplit ManySortedSign ;
cluster a1 +* a2 -> non empty strict Circuit-like unsplit ;
coherence
c1 +* c2 is unsplit
by Th57;
end;

theorem Th58: :: CIRCCOMB:58
for b1, b2 being non empty gate`1=arity ManySortedSign holds b1 +* b2 is gate`1=arity
proof end;

registration
let c1, c2 be non empty gate`1=arity ManySortedSign ;
cluster a1 +* a2 -> non empty strict gate`1=arity ;
coherence
c1 +* c2 is gate`1=arity
by Th58;
end;

theorem Th59: :: CIRCCOMB:59
for b1, b2 being non empty ManySortedSign holds
( b1 is gate`2isBoolean & b2 is gate`2isBoolean implies b1 +* b2 is gate`2isBoolean )
proof end;

definition
let c1 be Nat;
mode FinSeqLen of c1 -> FinSequence means :Def12: :: CIRCCOMB:def 12
len a2 = a1;
existence
ex b1 being FinSequence st len b1 = c1
proof end;
end;

:: deftheorem Def12 defines FinSeqLen CIRCCOMB:def 12 :
for b1 being Nat
for b2 being FinSequence holds
( b2 is FinSeqLen of b1 iff len b2 = b1 );

definition
let c1 be Nat;
let c2, c3 be non empty set ;
let c4 be Function of c1 -tuples_on c2,c3;
let c5 be FinSeqLen of c1;
let c6 be set ;
assume E53: ( c6 in rng c5 implies c2 = c3 ) ;
func 1GateCircuit c5,c4,c6 -> strict non-empty MSAlgebra of 1GateCircStr a5,a4,a6 means :: CIRCCOMB:def 13
( the Sorts of a7 = ((rng a5) --> a2) +* ({a6} --> a3) & the Charact of a7 . [a5,a4] = a4 );
existence
ex b1 being strict non-empty MSAlgebra of 1GateCircStr c5,c4,c6 st
( the Sorts of b1 = ((rng c5) --> c2) +* ({c6} --> c3) & the Charact of b1 . [c5,c4] = c4 )
proof end;
uniqueness
for b1, b2 being strict non-empty MSAlgebra of 1GateCircStr c5,c4,c6 holds
( the Sorts of b1 = ((rng c5) --> c2) +* ({c6} --> c3) & the Charact of b1 . [c5,c4] = c4 & the Sorts of b2 = ((rng c5) --> c2) +* ({c6} --> c3) & the Charact of b2 . [c5,c4] = c4 implies b1 = b2 )
proof end;
end;

:: deftheorem Def13 defines 1GateCircuit CIRCCOMB:def 13 :
for b1 being Nat
for b2, b3 being non empty set
for b4 being Function of b1 -tuples_on b2,b3
for b5 being FinSeqLen of b1
for b6 being set holds
( ( b6 in rng b5 implies b2 = b3 ) implies for b7 being strict non-empty MSAlgebra of 1GateCircStr b5,b4,b6 holds
( b7 = 1GateCircuit b5,b4,b6 iff ( the Sorts of b7 = ((rng b5) --> b2) +* ({b6} --> b3) & the Charact of b7 . [b5,b4] = b4 ) ) );

definition
let c1 be Nat;
let c2 be non empty set ;
let c3 be Function of c1 -tuples_on c2,c2;
let c4 be FinSeqLen of c1;
func 1GateCircuit c4,c3 -> strict non-empty MSAlgebra of 1GateCircStr a4,a3 means :Def14: :: CIRCCOMB:def 14
( the Sorts of a5 = the carrier of (1GateCircStr a4,a3) --> a2 & the Charact of a5 . [a4,a3] = a3 );
existence
ex b1 being strict non-empty MSAlgebra of 1GateCircStr c4,c3 st
( the Sorts of b1 = the carrier of (1GateCircStr c4,c3) --> c2 & the Charact of b1 . [c4,c3] = c3 )
proof end;
uniqueness
for b1, b2 being strict non-empty MSAlgebra of 1GateCircStr c4,c3 holds
( the Sorts of b1 = the carrier of (1GateCircStr c4,c3) --> c2 & the Charact of b1 . [c4,c3] = c3 & the Sorts of b2 = the carrier of (1GateCircStr c4,c3) --> c2 & the Charact of b2 . [c4,c3] = c3 implies b1 = b2 )
proof end;
end;

:: deftheorem Def14 defines 1GateCircuit CIRCCOMB:def 14 :
for b1 being Nat
for b2 being non empty set
for b3 being Function of b1 -tuples_on b2,b2
for b4 being FinSeqLen of b1
for b5 being strict non-empty MSAlgebra of 1GateCircStr b4,b3 holds
( b5 = 1GateCircuit b4,b3 iff ( the Sorts of b5 = the carrier of (1GateCircStr b4,b3) --> b2 & the Charact of b5 . [b4,b3] = b3 ) );

theorem Th60: :: CIRCCOMB:60
for b1 being Nat
for b2 being non empty set
for b3 being Function of b1 -tuples_on b2,b2
for b4 being FinSeqLen of b1 holds
( 1GateCircuit b4,b3 is gate`2=den & 1GateCircStr b4,b3 is gate`2=den )
proof end;

registration
let c1 be Nat;
let c2 be non empty set ;
let c3 be Function of c1 -tuples_on c2,c2;
let c4 be FinSeqLen of c1;
cluster 1GateCircuit a4,a3 -> strict non-empty gate`2=den ;
coherence
1GateCircuit c4,c3 is gate`2=den
by Th60;
cluster 1GateCircStr a4,a3 -> non empty strict non void Circuit-like unsplit gate`1=arity gate`2=den ;
coherence
1GateCircStr c4,c3 is gate`2=den
by Th60;
end;

theorem Th61: :: CIRCCOMB:61
for b1 being Nat
for b2 being FinSeqLen of b1
for b3 being Function of b1 -tuples_on BOOLEAN , BOOLEAN holds 1GateCircStr b2,b3 is gate`2isBoolean
proof end;

registration
let c1 be Nat;
let c2 be Function of c1 -tuples_on BOOLEAN , BOOLEAN ;
let c3 be FinSeqLen of c1;
cluster 1GateCircStr a3,a2 -> non empty strict non void Circuit-like unsplit gate`1=arity gate`2isBoolean gate`2=den ;
coherence
1GateCircStr c3,c2 is gate`2isBoolean
by Th61;
end;

registration
cluster non empty gate`2isBoolean gate`2=den ManySortedSign ;
existence
ex b1 being ManySortedSign st
( b1 is gate`2isBoolean & not b1 is empty )
proof end;
end;

registration
let c1, c2 be non empty gate`2isBoolean ManySortedSign ;
cluster a1 +* a2 -> non empty strict gate`2isBoolean gate`2=den ;
coherence
c1 +* c2 is gate`2isBoolean
by Th59;
end;

theorem Th62: :: CIRCCOMB:62
for b1 being Nat
for b2 being non empty set
for b3 being Function of b1 -tuples_on b2,b2
for b4 being FinSeqLen of b1 holds
( the Charact of (1GateCircuit b4,b3) = {[b4,b3]} --> b3 & ( for b5 being Vertex of (1GateCircStr b4,b3) holds the Sorts of (1GateCircuit b4,b3) . b5 = b2 ) )
proof end;

registration
let c1 be Nat;
let c2 be non empty finite set ;
let c3 be Function of c1 -tuples_on c2,c2;
let c4 be FinSeqLen of c1;
cluster 1GateCircuit a4,a3 -> strict non-empty locally-finite gate`2=den ;
coherence
1GateCircuit c4,c3 is locally-finite
proof end;
end;

theorem Th63: :: CIRCCOMB:63
for b1 being Nat
for b2 being non empty set
for b3 being Function of b1 -tuples_on b2,b2
for b4, b5 being FinSeqLen of b1 holds 1GateCircuit b4,b3 tolerates 1GateCircuit b5,b3
proof end;

theorem Th64: :: CIRCCOMB:64
for b1 being Nat
for b2 being non empty finite set
for b3 being Function of b1 -tuples_on b2,b2
for b4 being FinSeqLen of b1
for b5 being State of (1GateCircuit b4,b3) holds (Following b5) . [b4,b3] = b3 . (b5 * b4)
proof end;

definition
redefine func BOOLEAN as BOOLEAN -> non empty finite Subset of NAT ;
coherence
BOOLEAN is non empty finite Subset of NAT
by MARGREL1:def 12, ZFMISC_1:38;
end;

definition
let c1 be non empty ManySortedSign ;
let c2 be MSAlgebra of c1;
attr a2 is Boolean means :Def15: :: CIRCCOMB:def 15
for b1 being Vertex of a1 holds the Sorts of a2 . b1 = BOOLEAN ;
end;

:: deftheorem Def15 defines Boolean CIRCCOMB:def 15 :
for b1 being non empty ManySortedSign
for b2 being MSAlgebra of b1 holds
( b2 is Boolean iff for b3 being Vertex of b1 holds the Sorts of b2 . b3 = BOOLEAN );

theorem Th65: :: CIRCCOMB:65
for b1 being non empty ManySortedSign
for b2 being MSAlgebra of b1 holds
( b2 is Boolean iff the Sorts of b2 = the carrier of b1 --> BOOLEAN )
proof end;

registration
let c1 be non empty ManySortedSign ;
cluster Boolean -> non-empty locally-finite MSAlgebra of a1;
coherence
for b1 being MSAlgebra of c1 holds
( b1 is Boolean implies ( b1 is non-empty & b1 is locally-finite ) )
proof end;
end;

theorem Th66: :: CIRCCOMB:66
for b1 being non empty ManySortedSign
for b2 being MSAlgebra of b1 holds
( b2 is Boolean iff rng the Sorts of b2 c= {BOOLEAN } )
proof end;

theorem Th67: :: CIRCCOMB:67
for b1, b2 being non empty ManySortedSign
for b3 being MSAlgebra of b1
for b4 being MSAlgebra of b2 holds
( b3 is Boolean & b4 is Boolean implies the Sorts of b3 tolerates the Sorts of b4 )
proof end;

theorem Th68: :: CIRCCOMB:68
for b1, b2 being non empty unsplit gate`1=arity ManySortedSign
for b3 being MSAlgebra of b1
for b4 being MSAlgebra of b2 holds
( b3 is Boolean & b3 is gate`2=den & b4 is Boolean & b4 is gate`2=den implies b3 tolerates b4 )
proof end;

registration
let c1 be non empty ManySortedSign ;
cluster strict non-empty locally-finite Boolean MSAlgebra of a1;
existence
ex b1 being strict MSAlgebra of c1 st b1 is Boolean
proof end;
end;

theorem Th69: :: CIRCCOMB:69
for b1 being Nat
for b2 being Function of b1 -tuples_on BOOLEAN , BOOLEAN
for b3 being FinSeqLen of b1 holds 1GateCircuit b3,b2 is Boolean
proof end;

theorem Th70: :: CIRCCOMB:70
for b1, b2 being non empty ManySortedSign
for b3 being Boolean MSAlgebra of b1
for b4 being Boolean MSAlgebra of b2 holds b3 +* b4 is Boolean
proof end;

theorem Th71: :: CIRCCOMB:71
for b1, b2 being non empty ManySortedSign
for b3 being non-empty MSAlgebra of b1
for b4 being non-empty MSAlgebra of b2 holds
( b3 is gate`2=den & b4 is gate`2=den & the Sorts of b3 tolerates the Sorts of b4 implies b3 +* b4 is gate`2=den )
proof end;

registration
cluster non empty strict non void Circuit-like unsplit gate`1=arity gate`2isBoolean gate`2=den ManySortedSign ;
existence
ex b1 being non empty ManySortedSign st
( b1 is unsplit & b1 is gate`1=arity & b1 is gate`2=den & b1 is gate`2isBoolean & not b1 is void & b1 is strict )
proof end;
end;

registration
let c1 be non empty gate`2isBoolean ManySortedSign ;
cluster strict non-empty locally-finite gate`2=den Boolean MSAlgebra of a1;
existence
ex b1 being strict MSAlgebra of c1 st
( b1 is Boolean & b1 is gate`2=den )
proof end;
end;

registration
let c1, c2 be non empty non void unsplit gate`2isBoolean ManySortedSign ;
let c3 be gate`2=den Boolean Circuit of c1;
let c4 be gate`2=den Boolean Circuit of c2;
cluster a3 +* a4 -> strict non-empty locally-finite gate`2=den Boolean ;
coherence
( c3 +* c4 is Boolean & c3 +* c4 is gate`2=den )
proof end;
end;

registration
let c1 be Nat;
let c2 be non empty finite set ;
let c3 be Function of c1 -tuples_on c2,c2;
let c4 be FinSeqLen of c1;
cluster strict non-empty gate`2=den MSAlgebra of 1GateCircStr a4,a3;
existence
ex b1 being Circuit of 1GateCircStr c4,c3 st
( b1 is gate`2=den & b1 is strict & b1 is non-empty )
proof end;
end;

registration
let c1 be Nat;
let c2 be non empty finite set ;
let c3 be Function of c1 -tuples_on c2,c2;
let c4 be FinSeqLen of c1;
cluster 1GateCircuit a4,a3 -> strict non-empty locally-finite gate`2=den ;
coherence
1GateCircuit c4,c3 is gate`2=den
;
end;

theorem Th72: :: CIRCCOMB:72
for b1, b2 being non empty non void unsplit gate`1=arity gate`2isBoolean ManySortedSign
for b3 being gate`2=den Boolean Circuit of b1
for b4 being gate`2=den Boolean Circuit of b2
for b5 being State of (b3 +* b4)
for b6 being Vertex of (b1 +* b2) holds
( ( for b7 being State of b3 holds
( b7 = b5 | the carrier of b1 & ( b6 in InnerVertices b1 or ( b6 in the carrier of b1 & b6 in InputVertices (b1 +* b2) ) ) implies (Following b5) . b6 = (Following b7) . b6 ) ) & ( for b7 being State of b4 holds
( b7 = b5 | the carrier of b2 & ( b6 in InnerVertices b2 or ( b6 in the carrier of b2 & b6 in InputVertices (b1 +* b2) ) ) implies (Following b5) . b6 = (Following b7) . b6 ) ) )
proof end;