begin
theorem Th1:
theorem
canceled;
theorem
:: deftheorem Def1 defines SingleMSS FACIRC_2:def 1 :
for x being set
for b2 being void strict ManySortedSign holds
( b2 = SingleMSS x iff the carrier of b2 = {x} );
:: deftheorem defines SingleMSA FACIRC_2:def 2 :
for x being set
for b2 being strict Boolean MSAlgebra of SingleMSS x holds
( b2 = SingleMSA x iff verum );
theorem
theorem Th5:
theorem
definition
let n be
Nat;
let x,
y be
FinSequence;
A1:
(
1GateCircStr (
<*>,
((0 -tuples_on BOOLEAN) --> FALSE)) is
unsplit &
1GateCircStr (
<*>,
((0 -tuples_on BOOLEAN) --> FALSE)) is
gate`1=arity &
1GateCircStr (
<*>,
((0 -tuples_on BOOLEAN) --> FALSE)) is
gate`2isBoolean & not
1GateCircStr (
<*>,
((0 -tuples_on BOOLEAN) --> FALSE)) is
void & not
1GateCircStr (
<*>,
((0 -tuples_on BOOLEAN) --> FALSE)) is
empty &
1GateCircStr (
<*>,
((0 -tuples_on BOOLEAN) --> FALSE)) is
strict )
;
func n -BitAdderStr (
x,
y)
-> non
empty non
void strict unsplit gate`1=arity gate`2isBoolean ManySortedSign means :
Def3:
ex
f,
g being
ManySortedSet of
NAT st
(
it = f . n &
f . 0 = 1GateCircStr (
<*>,
((0 -tuples_on BOOLEAN) --> FALSE)) &
g . 0 = [<*>,((0 -tuples_on BOOLEAN) --> FALSE)] & ( for
n being
Nat for
S being non
empty ManySortedSign for
z being
set st
S = f . n &
z = g . n holds
(
f . (n + 1) = S +* (BitAdderWithOverflowStr ((x . (n + 1)),(y . (n + 1)),z)) &
g . (n + 1) = MajorityOutput (
(x . (n + 1)),
(y . (n + 1)),
z) ) ) );
uniqueness
for b1, b2 being non empty non void strict unsplit gate`1=arity gate`2isBoolean ManySortedSign st ex f, g being ManySortedSet of NAT st
( b1 = f . n & f . 0 = 1GateCircStr (<*>,((0 -tuples_on BOOLEAN) --> FALSE)) & g . 0 = [<*>,((0 -tuples_on BOOLEAN) --> FALSE)] & ( for n being Nat
for S being non empty ManySortedSign
for z being set st S = f . n & z = g . n holds
( f . (n + 1) = S +* (BitAdderWithOverflowStr ((x . (n + 1)),(y . (n + 1)),z)) & g . (n + 1) = MajorityOutput ((x . (n + 1)),(y . (n + 1)),z) ) ) ) & ex f, g being ManySortedSet of NAT st
( b2 = f . n & f . 0 = 1GateCircStr (<*>,((0 -tuples_on BOOLEAN) --> FALSE)) & g . 0 = [<*>,((0 -tuples_on BOOLEAN) --> FALSE)] & ( for n being Nat
for S being non empty ManySortedSign
for z being set st S = f . n & z = g . n holds
( f . (n + 1) = S +* (BitAdderWithOverflowStr ((x . (n + 1)),(y . (n + 1)),z)) & g . (n + 1) = MajorityOutput ((x . (n + 1)),(y . (n + 1)),z) ) ) ) holds
b1 = b2
existence
ex b1 being non empty non void strict unsplit gate`1=arity gate`2isBoolean ManySortedSign ex f, g being ManySortedSet of NAT st
( b1 = f . n & f . 0 = 1GateCircStr (<*>,((0 -tuples_on BOOLEAN) --> FALSE)) & g . 0 = [<*>,((0 -tuples_on BOOLEAN) --> FALSE)] & ( for n being Nat
for S being non empty ManySortedSign
for z being set st S = f . n & z = g . n holds
( f . (n + 1) = S +* (BitAdderWithOverflowStr ((x . (n + 1)),(y . (n + 1)),z)) & g . (n + 1) = MajorityOutput ((x . (n + 1)),(y . (n + 1)),z) ) ) )
end;
:: deftheorem Def3 defines -BitAdderStr FACIRC_2:def 3 :
for n being Nat
for x, y being FinSequence
for b4 being non empty non void strict unsplit gate`1=arity gate`2isBoolean ManySortedSign holds
( b4 = n -BitAdderStr (x,y) iff ex f, g being ManySortedSet of NAT st
( b4 = f . n & f . 0 = 1GateCircStr (<*>,((0 -tuples_on BOOLEAN) --> FALSE)) & g . 0 = [<*>,((0 -tuples_on BOOLEAN) --> FALSE)] & ( for n being Nat
for S being non empty ManySortedSign
for z being set st S = f . n & z = g . n holds
( f . (n + 1) = S +* (BitAdderWithOverflowStr ((x . (n + 1)),(y . (n + 1)),z)) & g . (n + 1) = MajorityOutput ((x . (n + 1)),(y . (n + 1)),z) ) ) ) );
definition
let n be
Nat;
let x,
y be
FinSequence;
func n -BitAdderCirc (
x,
y)
-> strict gate`2=den Boolean Circuit of
n -BitAdderStr (
x,
y)
means :
Def4:
ex
f,
g,
h being
ManySortedSet of
NAT st
(
n -BitAdderStr (
x,
y)
= f . n &
it = g . n &
f . 0 = 1GateCircStr (
<*>,
((0 -tuples_on BOOLEAN) --> FALSE)) &
g . 0 = 1GateCircuit (
<*>,
((0 -tuples_on BOOLEAN) --> FALSE)) &
h . 0 = [<*>,((0 -tuples_on BOOLEAN) --> FALSE)] & ( for
n being
Nat for
S being non
empty ManySortedSign for
A being
non-empty MSAlgebra of
S for
z being
set st
S = f . n &
A = g . n &
z = h . n holds
(
f . (n + 1) = S +* (BitAdderWithOverflowStr ((x . (n + 1)),(y . (n + 1)),z)) &
g . (n + 1) = A +* (BitAdderWithOverflowCirc ((x . (n + 1)),(y . (n + 1)),z)) &
h . (n + 1) = MajorityOutput (
(x . (n + 1)),
(y . (n + 1)),
z) ) ) );
uniqueness
for b1, b2 being strict gate`2=den Boolean Circuit of n -BitAdderStr (x,y) st ex f, g, h being ManySortedSet of NAT st
( n -BitAdderStr (x,y) = f . n & b1 = g . n & f . 0 = 1GateCircStr (<*>,((0 -tuples_on BOOLEAN) --> FALSE)) & g . 0 = 1GateCircuit (<*>,((0 -tuples_on BOOLEAN) --> FALSE)) & h . 0 = [<*>,((0 -tuples_on BOOLEAN) --> FALSE)] & ( for n being Nat
for S being non empty ManySortedSign
for A being non-empty MSAlgebra of S
for z being set st S = f . n & A = g . n & z = h . n holds
( f . (n + 1) = S +* (BitAdderWithOverflowStr ((x . (n + 1)),(y . (n + 1)),z)) & g . (n + 1) = A +* (BitAdderWithOverflowCirc ((x . (n + 1)),(y . (n + 1)),z)) & h . (n + 1) = MajorityOutput ((x . (n + 1)),(y . (n + 1)),z) ) ) ) & ex f, g, h being ManySortedSet of NAT st
( n -BitAdderStr (x,y) = f . n & b2 = g . n & f . 0 = 1GateCircStr (<*>,((0 -tuples_on BOOLEAN) --> FALSE)) & g . 0 = 1GateCircuit (<*>,((0 -tuples_on BOOLEAN) --> FALSE)) & h . 0 = [<*>,((0 -tuples_on BOOLEAN) --> FALSE)] & ( for n being Nat
for S being non empty ManySortedSign
for A being non-empty MSAlgebra of S
for z being set st S = f . n & A = g . n & z = h . n holds
( f . (n + 1) = S +* (BitAdderWithOverflowStr ((x . (n + 1)),(y . (n + 1)),z)) & g . (n + 1) = A +* (BitAdderWithOverflowCirc ((x . (n + 1)),(y . (n + 1)),z)) & h . (n + 1) = MajorityOutput ((x . (n + 1)),(y . (n + 1)),z) ) ) ) holds
b1 = b2
existence
ex b1 being strict gate`2=den Boolean Circuit of n -BitAdderStr (x,y) ex f, g, h being ManySortedSet of NAT st
( n -BitAdderStr (x,y) = f . n & b1 = g . n & f . 0 = 1GateCircStr (<*>,((0 -tuples_on BOOLEAN) --> FALSE)) & g . 0 = 1GateCircuit (<*>,((0 -tuples_on BOOLEAN) --> FALSE)) & h . 0 = [<*>,((0 -tuples_on BOOLEAN) --> FALSE)] & ( for n being Nat
for S being non empty ManySortedSign
for A being non-empty MSAlgebra of S
for z being set st S = f . n & A = g . n & z = h . n holds
( f . (n + 1) = S +* (BitAdderWithOverflowStr ((x . (n + 1)),(y . (n + 1)),z)) & g . (n + 1) = A +* (BitAdderWithOverflowCirc ((x . (n + 1)),(y . (n + 1)),z)) & h . (n + 1) = MajorityOutput ((x . (n + 1)),(y . (n + 1)),z) ) ) )
end;
:: deftheorem Def4 defines -BitAdderCirc FACIRC_2:def 4 :
for n being Nat
for x, y being FinSequence
for b4 being strict gate`2=den Boolean Circuit of n -BitAdderStr (x,y) holds
( b4 = n -BitAdderCirc (x,y) iff ex f, g, h being ManySortedSet of NAT st
( n -BitAdderStr (x,y) = f . n & b4 = g . n & f . 0 = 1GateCircStr (<*>,((0 -tuples_on BOOLEAN) --> FALSE)) & g . 0 = 1GateCircuit (<*>,((0 -tuples_on BOOLEAN) --> FALSE)) & h . 0 = [<*>,((0 -tuples_on BOOLEAN) --> FALSE)] & ( for n being Nat
for S being non empty ManySortedSign
for A being non-empty MSAlgebra of S
for z being set st S = f . n & A = g . n & z = h . n holds
( f . (n + 1) = S +* (BitAdderWithOverflowStr ((x . (n + 1)),(y . (n + 1)),z)) & g . (n + 1) = A +* (BitAdderWithOverflowCirc ((x . (n + 1)),(y . (n + 1)),z)) & h . (n + 1) = MajorityOutput ((x . (n + 1)),(y . (n + 1)),z) ) ) ) );
definition
let n be
Nat;
let x,
y be
FinSequence;
set c =
[<*>,((0 -tuples_on BOOLEAN) --> FALSE)];
func n -BitMajorityOutput (
x,
y)
-> Element of
InnerVertices (n -BitAdderStr (x,y)) means :
Def5:
ex
h being
ManySortedSet of
NAT st
(
it = h . n &
h . 0 = [<*>,((0 -tuples_on BOOLEAN) --> FALSE)] & ( for
n being
Nat for
z being
set st
z = h . n holds
h . (n + 1) = MajorityOutput (
(x . (n + 1)),
(y . (n + 1)),
z) ) );
uniqueness
for b1, b2 being Element of InnerVertices (n -BitAdderStr (x,y)) st ex h being ManySortedSet of NAT st
( b1 = h . n & h . 0 = [<*>,((0 -tuples_on BOOLEAN) --> FALSE)] & ( for n being Nat
for z being set st z = h . n holds
h . (n + 1) = MajorityOutput ((x . (n + 1)),(y . (n + 1)),z) ) ) & ex h being ManySortedSet of NAT st
( b2 = h . n & h . 0 = [<*>,((0 -tuples_on BOOLEAN) --> FALSE)] & ( for n being Nat
for z being set st z = h . n holds
h . (n + 1) = MajorityOutput ((x . (n + 1)),(y . (n + 1)),z) ) ) holds
b1 = b2
existence
ex b1 being Element of InnerVertices (n -BitAdderStr (x,y)) ex h being ManySortedSet of NAT st
( b1 = h . n & h . 0 = [<*>,((0 -tuples_on BOOLEAN) --> FALSE)] & ( for n being Nat
for z being set st z = h . n holds
h . (n + 1) = MajorityOutput ((x . (n + 1)),(y . (n + 1)),z) ) )
end;
:: deftheorem Def5 defines -BitMajorityOutput FACIRC_2:def 5 :
for n being Nat
for x, y being FinSequence
for b4 being Element of InnerVertices (n -BitAdderStr (x,y)) holds
( b4 = n -BitMajorityOutput (x,y) iff ex h being ManySortedSet of NAT st
( b4 = h . n & h . 0 = [<*>,((0 -tuples_on BOOLEAN) --> FALSE)] & ( for n being Nat
for z being set st z = h . n holds
h . (n + 1) = MajorityOutput ((x . (n + 1)),(y . (n + 1)),z) ) ) );
theorem Th7:
for
x,
y being
FinSequence for
f,
g,
h being
ManySortedSet of
NAT st
f . 0 = 1GateCircStr (
<*>,
((0 -tuples_on BOOLEAN) --> FALSE)) &
g . 0 = 1GateCircuit (
<*>,
((0 -tuples_on BOOLEAN) --> FALSE)) &
h . 0 = [<*>,((0 -tuples_on BOOLEAN) --> FALSE)] & ( for
n being
Nat for
S being non
empty ManySortedSign for
A being
non-empty MSAlgebra of
S for
z being
set st
S = f . n &
A = g . n &
z = h . n holds
(
f . (n + 1) = S +* (BitAdderWithOverflowStr ((x . (n + 1)),(y . (n + 1)),z)) &
g . (n + 1) = A +* (BitAdderWithOverflowCirc ((x . (n + 1)),(y . (n + 1)),z)) &
h . (n + 1) = MajorityOutput (
(x . (n + 1)),
(y . (n + 1)),
z) ) ) holds
for
n being
Nat holds
(
n -BitAdderStr (
x,
y)
= f . n &
n -BitAdderCirc (
x,
y)
= g . n &
n -BitMajorityOutput (
x,
y)
= h . n )
theorem Th8:
theorem Th9:
for
a,
b being
FinSequence for
c being
set st
c = [<*>,((0 -tuples_on BOOLEAN) --> FALSE)] holds
( 1
-BitAdderStr (
a,
b)
= (1GateCircStr (<*>,((0 -tuples_on BOOLEAN) --> FALSE))) +* (BitAdderWithOverflowStr ((a . 1),(b . 1),c)) & 1
-BitAdderCirc (
a,
b)
= (1GateCircuit (<*>,((0 -tuples_on BOOLEAN) --> FALSE))) +* (BitAdderWithOverflowCirc ((a . 1),(b . 1),c)) & 1
-BitMajorityOutput (
a,
b)
= MajorityOutput (
(a . 1),
(b . 1),
c) )
theorem
for
a,
b,
c being
set st
c = [<*>,((0 -tuples_on BOOLEAN) --> FALSE)] holds
( 1
-BitAdderStr (
<*a*>,
<*b*>)
= (1GateCircStr (<*>,((0 -tuples_on BOOLEAN) --> FALSE))) +* (BitAdderWithOverflowStr (a,b,c)) & 1
-BitAdderCirc (
<*a*>,
<*b*>)
= (1GateCircuit (<*>,((0 -tuples_on BOOLEAN) --> FALSE))) +* (BitAdderWithOverflowCirc (a,b,c)) & 1
-BitMajorityOutput (
<*a*>,
<*b*>)
= MajorityOutput (
a,
b,
c) )
theorem Th11:
for
n being
Element of
NAT for
p,
q being
FinSeqLen of
n for
p1,
p2,
q1,
q2 being
FinSequence holds
(
n -BitAdderStr (
(p ^ p1),
(q ^ q1))
= n -BitAdderStr (
(p ^ p2),
(q ^ q2)) &
n -BitAdderCirc (
(p ^ p1),
(q ^ q1))
= n -BitAdderCirc (
(p ^ p2),
(q ^ q2)) &
n -BitMajorityOutput (
(p ^ p1),
(q ^ q1))
= n -BitMajorityOutput (
(p ^ p2),
(q ^ q2)) )
theorem
for
n being
Element of
NAT for
x,
y being
FinSeqLen of
n for
a,
b being
set holds
(
(n + 1) -BitAdderStr (
(x ^ <*a*>),
(y ^ <*b*>))
= (n -BitAdderStr (x,y)) +* (BitAdderWithOverflowStr (a,b,(n -BitMajorityOutput (x,y)))) &
(n + 1) -BitAdderCirc (
(x ^ <*a*>),
(y ^ <*b*>))
= (n -BitAdderCirc (x,y)) +* (BitAdderWithOverflowCirc (a,b,(n -BitMajorityOutput (x,y)))) &
(n + 1) -BitMajorityOutput (
(x ^ <*a*>),
(y ^ <*b*>))
= MajorityOutput (
a,
b,
(n -BitMajorityOutput (x,y))) )
theorem Th13:
for
n being
Element of
NAT for
x,
y being
FinSequence holds
(
(n + 1) -BitAdderStr (
x,
y)
= (n -BitAdderStr (x,y)) +* (BitAdderWithOverflowStr ((x . (n + 1)),(y . (n + 1)),(n -BitMajorityOutput (x,y)))) &
(n + 1) -BitAdderCirc (
x,
y)
= (n -BitAdderCirc (x,y)) +* (BitAdderWithOverflowCirc ((x . (n + 1)),(y . (n + 1)),(n -BitMajorityOutput (x,y)))) &
(n + 1) -BitMajorityOutput (
x,
y)
= MajorityOutput (
(x . (n + 1)),
(y . (n + 1)),
(n -BitMajorityOutput (x,y))) )
theorem Th14:
theorem
definition
let k,
n be
Nat;
assume that A1:
k >= 1
and A2:
k <= n
;
let x,
y be
FinSequence;
func (
k,
n)
-BitAdderOutput (
x,
y)
-> Element of
InnerVertices (n -BitAdderStr (x,y)) means :
Def6:
ex
i being
Element of
NAT st
(
k = i + 1 &
it = BitAdderOutput (
(x . k),
(y . k),
(i -BitMajorityOutput (x,y))) );
uniqueness
for b1, b2 being Element of InnerVertices (n -BitAdderStr (x,y)) st ex i being Element of NAT st
( k = i + 1 & b1 = BitAdderOutput ((x . k),(y . k),(i -BitMajorityOutput (x,y))) ) & ex i being Element of NAT st
( k = i + 1 & b2 = BitAdderOutput ((x . k),(y . k),(i -BitMajorityOutput (x,y))) ) holds
b1 = b2
;
existence
ex b1 being Element of InnerVertices (n -BitAdderStr (x,y)) ex i being Element of NAT st
( k = i + 1 & b1 = BitAdderOutput ((x . k),(y . k),(i -BitMajorityOutput (x,y))) )
end;
:: deftheorem Def6 defines -BitAdderOutput FACIRC_2:def 6 :
for k, n being Nat st k >= 1 & k <= n holds
for x, y being FinSequence
for b5 being Element of InnerVertices (n -BitAdderStr (x,y)) holds
( b5 = (k,n) -BitAdderOutput (x,y) iff ex i being Element of NAT st
( k = i + 1 & b5 = BitAdderOutput ((x . k),(y . k),(i -BitMajorityOutput (x,y))) ) );
theorem
theorem
theorem Th18:
for
x,
y,
c being
set holds
InnerVertices (MajorityIStr (x,y,c)) = {[<*x,y*>,'&'],[<*y,c*>,'&'],[<*c,x*>,'&']}
theorem Th19:
for
x,
y,
c being
set st
x <> [<*y,c*>,'&'] &
y <> [<*c,x*>,'&'] &
c <> [<*x,y*>,'&'] holds
InputVertices (MajorityIStr (x,y,c)) = {x,y,c}
theorem Th20:
for
x,
y,
c being
set holds
InnerVertices (MajorityStr (x,y,c)) = {[<*x,y*>,'&'],[<*y,c*>,'&'],[<*c,x*>,'&']} \/ {(MajorityOutput (x,y,c))}
theorem Th21:
for
x,
y,
c being
set st
x <> [<*y,c*>,'&'] &
y <> [<*c,x*>,'&'] &
c <> [<*x,y*>,'&'] holds
InputVertices (MajorityStr (x,y,c)) = {x,y,c}
theorem Th22:
theorem Th23:
for
x,
y,
c being
set st
x <> [<*y,c*>,'&'] &
y <> [<*c,x*>,'&'] &
c <> [<*x,y*>,'&'] &
c <> [<*x,y*>,'xor'] holds
InputVertices (BitAdderWithOverflowStr (x,y,c)) = {x,y,c}
theorem Th24:
for
x,
y,
c being
set holds
InnerVertices (BitAdderWithOverflowStr (x,y,c)) = ({[<*x,y*>,'xor'],(2GatesCircOutput (x,y,c,'xor'))} \/ {[<*x,y*>,'&'],[<*y,c*>,'&'],[<*c,x*>,'&']}) \/ {(MajorityOutput (x,y,c))}
theorem Th25:
theorem Th26:
theorem Th27:
theorem
Lm3:
for x, y, c being set
for s being State of (MajorityCirc (x,y,c))
for a1, a2, a3 being Element of BOOLEAN st a1 = s . x & a2 = s . y & a3 = s . c holds
( (Following s) . [<*x,y*>,'&'] = a1 '&' a2 & (Following s) . [<*y,c*>,'&'] = a2 '&' a3 & (Following s) . [<*c,x*>,'&'] = a3 '&' a1 )
theorem Th29:
for
x,
y,
c being
set for
s being
State of
(MajorityCirc (x,y,c)) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . [<*x,y*>,'&'] &
a2 = s . [<*y,c*>,'&'] &
a3 = s . [<*c,x*>,'&'] holds
(Following s) . (MajorityOutput (x,y,c)) = (a1 'or' a2) 'or' a3
Lm4:
for x, y, c being set st x <> [<*y,c*>,'&'] & y <> [<*c,x*>,'&'] & c <> [<*x,y*>,'&'] holds
for s being State of (MajorityCirc (x,y,c))
for a1, a2, a3 being Element of BOOLEAN st a1 = s . x & a2 = s . y & a3 = s . c holds
( (Following (s,2)) . (MajorityOutput (x,y,c)) = ((a1 '&' a2) 'or' (a2 '&' a3)) 'or' (a3 '&' a1) & (Following (s,2)) . [<*x,y*>,'&'] = a1 '&' a2 & (Following (s,2)) . [<*y,c*>,'&'] = a2 '&' a3 & (Following (s,2)) . [<*c,x*>,'&'] = a3 '&' a1 )
theorem Th30:
for
x,
y,
c being
set st
x <> [<*y,c*>,'&'] &
y <> [<*c,x*>,'&'] &
c <> [<*x,y*>,'&'] holds
for
s being
State of
(MajorityCirc (x,y,c)) holds
Following (
s,2) is
stable
theorem
for
x,
y,
c being
set st
x <> [<*y,c*>,'&'] &
y <> [<*c,x*>,'&'] &
c <> [<*x,y*>,'&'] &
c <> [<*x,y*>,'xor'] holds
for
s being
State of
(BitAdderWithOverflowCirc (x,y,c)) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . x &
a2 = s . y &
a3 = s . c holds
(
(Following (s,2)) . (BitAdderOutput (x,y,c)) = (a1 'xor' a2) 'xor' a3 &
(Following (s,2)) . (MajorityOutput (x,y,c)) = ((a1 '&' a2) 'or' (a2 '&' a3)) 'or' (a3 '&' a1) )
theorem Th32:
for
x,
y,
c being
set st
x <> [<*y,c*>,'&'] &
y <> [<*c,x*>,'&'] &
c <> [<*x,y*>,'&'] &
c <> [<*x,y*>,'xor'] holds
for
s being
State of
(BitAdderWithOverflowCirc (x,y,c)) holds
Following (
s,2) is
stable
Lm5:
for n being Element of NAT ex N being Function of NAT,NAT st
( N . 0 = 1 & N . 1 = 2 & N . 2 = n )
theorem
theorem
theorem
canceled;
theorem
theorem