begin
theorem
canceled;
theorem Th2:
theorem Th3:
theorem Th4:
theorem Th5:
theorem Th6:
:: deftheorem defines tolerates CIRCCOMB:def 1 :
for S1, S2 being ManySortedSign holds
( S1 tolerates S2 iff ( the Arity of S1 tolerates the Arity of S2 & the ResultSort of S1 tolerates the ResultSort of S2 ) );
:: deftheorem Def2 defines +* CIRCCOMB:def 2 :
for S1, S2 being non empty ManySortedSign
for b3 being non empty strict ManySortedSign holds
( b3 = S1 +* S2 iff ( the carrier of b3 = the carrier of S1 \/ the carrier of S2 & the carrier' of b3 = the carrier' of S1 \/ the carrier' of S2 & the Arity of b3 = the Arity of S1 +* the Arity of S2 & the ResultSort of b3 = the ResultSort of S1 +* the ResultSort of S2 ) );
theorem Th7:
theorem
theorem Th9:
theorem
theorem
theorem
theorem Th13:
theorem
theorem Th15:
theorem Th16:
theorem Th17:
theorem Th18:
theorem Th19:
theorem Th20:
theorem Th21:
begin
:: deftheorem Def3 defines tolerates CIRCCOMB:def 3 :
for S1, S2 being non empty ManySortedSign
for A1 being MSAlgebra of S1
for A2 being MSAlgebra of S2 holds
( A1 tolerates A2 iff ( S1 tolerates S2 & the Sorts of A1 tolerates the Sorts of A2 & the Charact of A1 tolerates the Charact of A2 ) );
:: deftheorem Def4 defines +* CIRCCOMB:def 4 :
for S1, S2 being non empty ManySortedSign
for A1 being non-empty MSAlgebra of S1
for A2 being non-empty MSAlgebra of S2 st the Sorts of A1 tolerates the Sorts of A2 holds
for b5 being strict non-empty MSAlgebra of S1 +* S2 holds
( b5 = A1 +* A2 iff ( the Sorts of b5 = the Sorts of A1 +* the Sorts of A2 & the Charact of b5 = the Charact of A1 +* the Charact of A2 ) );
theorem
theorem Th23:
theorem
theorem
theorem Th26:
theorem
theorem
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
theorem
theorem Th34:
theorem Th35:
theorem Th36:
theorem Th37:
theorem Th38:
theorem Th39:
theorem Th40:
theorem
theorem
begin
definition
let f be
set ;
let p be
FinSequence;
let x be
set ;
func 1GateCircStr (
p,
f,
x)
-> non
void strict ManySortedSign means :
Def5:
( the
carrier of
it = (rng p) \/ {x} & the
carrier' of
it = {[p,f]} & the
Arity of
it . [p,f] = p & the
ResultSort of
it . [p,f] = x );
existence
ex b1 being non void strict ManySortedSign st
( the carrier of b1 = (rng p) \/ {x} & the carrier' of b1 = {[p,f]} & the Arity of b1 . [p,f] = p & the ResultSort of b1 . [p,f] = x )
uniqueness
for b1, b2 being non void strict ManySortedSign st the carrier of b1 = (rng p) \/ {x} & the carrier' of b1 = {[p,f]} & the Arity of b1 . [p,f] = p & the ResultSort of b1 . [p,f] = x & the carrier of b2 = (rng p) \/ {x} & the carrier' of b2 = {[p,f]} & the Arity of b2 . [p,f] = p & the ResultSort of b2 . [p,f] = x holds
b1 = b2
end;
:: deftheorem Def5 defines 1GateCircStr CIRCCOMB:def 5 :
for f being set
for p being FinSequence
for x being set
for b4 being non void strict ManySortedSign holds
( b4 = 1GateCircStr (p,f,x) iff ( the carrier of b4 = (rng p) \/ {x} & the carrier' of b4 = {[p,f]} & the Arity of b4 . [p,f] = p & the ResultSort of b4 . [p,f] = x ) );
theorem Th43:
theorem
theorem
definition
let f be
set ;
let p be
FinSequence;
func 1GateCircStr (
p,
f)
-> non
void strict ManySortedSign means :
Def6:
( the
carrier of
it = (rng p) \/ {[p,f]} & the
carrier' of
it = {[p,f]} & the
Arity of
it . [p,f] = p & the
ResultSort of
it . [p,f] = [p,f] );
existence
ex b1 being non void strict ManySortedSign st
( the carrier of b1 = (rng p) \/ {[p,f]} & the carrier' of b1 = {[p,f]} & the Arity of b1 . [p,f] = p & the ResultSort of b1 . [p,f] = [p,f] )
uniqueness
for b1, b2 being non void strict ManySortedSign st the carrier of b1 = (rng p) \/ {[p,f]} & the carrier' of b1 = {[p,f]} & the Arity of b1 . [p,f] = p & the ResultSort of b1 . [p,f] = [p,f] & the carrier of b2 = (rng p) \/ {[p,f]} & the carrier' of b2 = {[p,f]} & the Arity of b2 . [p,f] = p & the ResultSort of b2 . [p,f] = [p,f] holds
b1 = b2
end;
:: deftheorem Def6 defines 1GateCircStr CIRCCOMB:def 6 :
for f being set
for p being FinSequence
for b3 being non void strict ManySortedSign holds
( b3 = 1GateCircStr (p,f) iff ( the carrier of b3 = (rng p) \/ {[p,f]} & the carrier' of b3 = {[p,f]} & the Arity of b3 . [p,f] = p & the ResultSort of b3 . [p,f] = [p,f] ) );
theorem
theorem Th47:
theorem Th48:
theorem Th49:
theorem
canceled;
theorem Th51:
begin
:: deftheorem Def7 defines unsplit CIRCCOMB:def 7 :
for IT being ManySortedSign holds
( IT is unsplit iff the ResultSort of IT = id the carrier' of IT );
:: deftheorem Def8 defines gate`1=arity CIRCCOMB:def 8 :
for IT being ManySortedSign holds
( IT is gate`1=arity iff for g being set st g in the carrier' of IT holds
g = [( the Arity of IT . g),(g `2)] );
:: deftheorem Def9 defines gate`2isBoolean CIRCCOMB:def 9 :
for IT being ManySortedSign holds
( IT is gate`2isBoolean iff for g being set st g in the carrier' of IT holds
for p being FinSequence st p = the Arity of IT . g holds
ex f being Function of ((len p) -tuples_on BOOLEAN),BOOLEAN st g = [(g `1),f] );
:: deftheorem Def10 defines gate`2=den CIRCCOMB:def 10 :
for S being non empty ManySortedSign
for IT being MSAlgebra of S holds
( IT is gate`2=den iff for g being set st g in the carrier' of S holds
g = [(g `1),( the Charact of IT . g)] );
:: deftheorem defines gate`2=den CIRCCOMB:def 11 :
for IT being non empty ManySortedSign holds
( IT is gate`2=den iff ex A being MSAlgebra of IT st A is gate`2=den );
theorem Th52:
theorem
theorem Th54:
theorem Th55:
theorem Th56:
theorem Th57:
theorem Th58:
theorem Th59:
begin
definition
canceled;let n be
Nat;
let X,
Y be non
empty set ;
let f be
Function of
(n -tuples_on X),
Y;
let p be
FinSeqLen of
n;
let x be
set ;
assume A1:
(
x in rng p implies
X = Y )
;
func 1GateCircuit (
p,
f,
x)
-> strict non-empty MSAlgebra of
1GateCircStr (
p,
f,
x)
means
( the
Sorts of
it = ((rng p) --> X) +* (x .--> Y) & the
Charact of
it . [p,f] = f );
existence
ex b1 being strict non-empty MSAlgebra of 1GateCircStr (p,f,x) st
( the Sorts of b1 = ((rng p) --> X) +* (x .--> Y) & the Charact of b1 . [p,f] = f )
uniqueness
for b1, b2 being strict non-empty MSAlgebra of 1GateCircStr (p,f,x) st the Sorts of b1 = ((rng p) --> X) +* (x .--> Y) & the Charact of b1 . [p,f] = f & the Sorts of b2 = ((rng p) --> X) +* (x .--> Y) & the Charact of b2 . [p,f] = f holds
b1 = b2
end;
:: deftheorem CIRCCOMB:def 12 :
canceled;
:: deftheorem defines 1GateCircuit CIRCCOMB:def 13 :
for n being Nat
for X, Y being non empty set
for f being Function of (n -tuples_on X),Y
for p being FinSeqLen of n
for x being set st ( x in rng p implies X = Y ) holds
for b7 being strict non-empty MSAlgebra of 1GateCircStr (p,f,x) holds
( b7 = 1GateCircuit (p,f,x) iff ( the Sorts of b7 = ((rng p) --> X) +* (x .--> Y) & the Charact of b7 . [p,f] = f ) );
definition
let n be
Nat;
let X be non
empty set ;
let f be
Function of
(n -tuples_on X),
X;
let p be
FinSeqLen of
n;
func 1GateCircuit (
p,
f)
-> strict non-empty MSAlgebra of
1GateCircStr (
p,
f)
means :
Def14:
( the
Sorts of
it = the
carrier of
(1GateCircStr (p,f)) --> X & the
Charact of
it . [p,f] = f );
existence
ex b1 being strict non-empty MSAlgebra of 1GateCircStr (p,f) st
( the Sorts of b1 = the carrier of (1GateCircStr (p,f)) --> X & the Charact of b1 . [p,f] = f )
uniqueness
for b1, b2 being strict non-empty MSAlgebra of 1GateCircStr (p,f) st the Sorts of b1 = the carrier of (1GateCircStr (p,f)) --> X & the Charact of b1 . [p,f] = f & the Sorts of b2 = the carrier of (1GateCircStr (p,f)) --> X & the Charact of b2 . [p,f] = f holds
b1 = b2
end;
:: deftheorem Def14 defines 1GateCircuit CIRCCOMB:def 14 :
for n being Nat
for X being non empty set
for f being Function of (n -tuples_on X),X
for p being FinSeqLen of n
for b5 being strict non-empty MSAlgebra of 1GateCircStr (p,f) holds
( b5 = 1GateCircuit (p,f) iff ( the Sorts of b5 = the carrier of (1GateCircStr (p,f)) --> X & the Charact of b5 . [p,f] = f ) );
theorem Th60:
theorem Th61:
theorem Th62:
theorem
theorem
begin
:: deftheorem Def15 defines Boolean CIRCCOMB:def 15 :
for S being non empty ManySortedSign
for IT being MSAlgebra of S holds
( IT is Boolean iff for v being Vertex of S holds the Sorts of IT . v = BOOLEAN );
theorem Th65:
theorem
theorem Th67:
theorem Th68:
theorem
theorem Th70:
theorem Th71:
theorem