:: GFACIRC1 semantic presentation
:: deftheorem Def1 defines inv1 GFACIRC1:def 1 :
theorem Th1: :: GFACIRC1:1
:: deftheorem Def2 defines buf1 GFACIRC1:def 2 :
theorem Th2: :: GFACIRC1:2
definition
func and2c -> Function of 2
-tuples_on BOOLEAN ,
BOOLEAN means :
Def3:
:: GFACIRC1:def 3
for
x,
y being
Element of
BOOLEAN holds
it . <*x,y*> = x '&' ('not' y);
existence
ex b1 being Function of 2 -tuples_on BOOLEAN , BOOLEAN st
for x, y being Element of BOOLEAN holds b1 . <*x,y*> = x '&' ('not' y)
uniqueness
for b1, b2 being Function of 2 -tuples_on BOOLEAN , BOOLEAN st ( for x, y being Element of BOOLEAN holds b1 . <*x,y*> = x '&' ('not' y) ) & ( for x, y being Element of BOOLEAN holds b2 . <*x,y*> = x '&' ('not' y) ) holds
b1 = b2
end;
:: deftheorem Def3 defines and2c GFACIRC1:def 3 :
theorem Th3: :: GFACIRC1:3
for
x,
y being
Element of
BOOLEAN holds
(
and2c . <*x,y*> = x '&' ('not' y) &
and2c . <*x,y*> = and2a . <*y,x*> &
and2c . <*x,y*> = nor2a . <*x,y*> &
and2c . <*0,0*> = 0 &
and2c . <*0,1*> = 0 &
and2c . <*1,0*> = 1 &
and2c . <*1,1*> = 0 )
definition
func xor2c -> Function of 2
-tuples_on BOOLEAN ,
BOOLEAN means :
Def4:
:: GFACIRC1:def 4
for
x,
y being
Element of
BOOLEAN holds
it . <*x,y*> = x 'xor' ('not' y);
existence
ex b1 being Function of 2 -tuples_on BOOLEAN , BOOLEAN st
for x, y being Element of BOOLEAN holds b1 . <*x,y*> = x 'xor' ('not' y)
uniqueness
for b1, b2 being Function of 2 -tuples_on BOOLEAN , BOOLEAN st ( for x, y being Element of BOOLEAN holds b1 . <*x,y*> = x 'xor' ('not' y) ) & ( for x, y being Element of BOOLEAN holds b2 . <*x,y*> = x 'xor' ('not' y) ) holds
b1 = b2
end;
:: deftheorem Def4 defines xor2c GFACIRC1:def 4 :
theorem Th4: :: GFACIRC1:4
for
x,
y being
Element of
BOOLEAN holds
(
xor2c . <*x,y*> = x 'xor' ('not' y) &
xor2c . <*x,y*> = xor2a . <*x,y*> &
xor2c . <*x,y*> = or2 . <*(and2b . <*x,y*>),(and2 . <*x,y*>)*> &
xor2c . <*0,0*> = 1 &
xor2c . <*0,1*> = 0 &
xor2c . <*1,0*> = 0 &
xor2c . <*1,1*> = 1 )
theorem Th5: :: GFACIRC1:5
theorem Th6: :: GFACIRC1:6
theorem Th7: :: GFACIRC1:7
theorem Th8: :: GFACIRC1:8
theorem Th9: :: GFACIRC1:9
theorem Th10: :: GFACIRC1:10
theorem Th11: :: GFACIRC1:11
theorem Th12: :: GFACIRC1:12
Lemma49:
for f1, f2, f3 being Function of 2 -tuples_on BOOLEAN , BOOLEAN
for x, y, z being set st x <> [<*y,z*>,f2] & y <> [<*z,x*>,f3] & z <> [<*x,y*>,f1] holds
( not [<*x,y*>,f1] in {y,z} & not z in {[<*x,y*>,f1],[<*y,z*>,f2]} & not x in {[<*x,y*>,f1],[<*y,z*>,f2]} & not [<*z,x*>,f3] in {x,y,z} )
Lemma70:
for f1, f2, f3 being Function of 2 -tuples_on BOOLEAN , BOOLEAN
for f4 being Function of 3 -tuples_on BOOLEAN , BOOLEAN
for x, y, z being set holds {x,y,z} \ {[<*[<*x,y*>,f1],[<*y,z*>,f2],[<*z,x*>,f3]*>,f4]} = {x,y,z}
Lemma73:
for f being Function of 2 -tuples_on BOOLEAN , BOOLEAN
for x, y, c being set st c <> [<*x,y*>,f] holds
for s being State of (2GatesCircuit x,y,c,f) holds
( (Following s) . (2GatesCircOutput x,y,c,f) = f . <*(s . [<*x,y*>,f]),(s . c)*> & (Following s) . [<*x,y*>,f] = f . <*(s . x),(s . y)*> & (Following s) . x = s . x & (Following s) . y = s . y & (Following s) . c = s . c )
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA0CarryIStr c1,
c2,
c3 -> non
empty strict non
void unsplit gate`1=arity gate`2isBoolean ManySortedSign equals :: GFACIRC1:def 5
((1GateCircStr <*x,y*>,and2 ) +* (1GateCircStr <*y,z*>,and2 )) +* (1GateCircStr <*z,x*>,and2 );
coherence
((1GateCircStr <*x,y*>,and2 ) +* (1GateCircStr <*y,z*>,and2 )) +* (1GateCircStr <*z,x*>,and2 ) is non empty strict non void unsplit gate`1=arity gate`2isBoolean ManySortedSign
;
end;
:: deftheorem Def5 defines GFA0CarryIStr GFACIRC1:def 5 :
for
x,
y,
z being
set holds
GFA0CarryIStr x,
y,
z = ((1GateCircStr <*x,y*>,and2 ) +* (1GateCircStr <*y,z*>,and2 )) +* (1GateCircStr <*z,x*>,and2 );
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA0CarryICirc c1,
c2,
c3 -> strict gate`2=den Boolean Circuit of
GFA0CarryIStr a1,
a2,
a3 equals :: GFACIRC1:def 6
((1GateCircuit x,y,and2 ) +* (1GateCircuit y,z,and2 )) +* (1GateCircuit z,x,and2 );
coherence
((1GateCircuit x,y,and2 ) +* (1GateCircuit y,z,and2 )) +* (1GateCircuit z,x,and2 ) is strict gate`2=den Boolean Circuit of GFA0CarryIStr x,y,z
;
end;
:: deftheorem Def6 defines GFA0CarryICirc GFACIRC1:def 6 :
for
x,
y,
z being
set holds
GFA0CarryICirc x,
y,
z = ((1GateCircuit x,y,and2 ) +* (1GateCircuit y,z,and2 )) +* (1GateCircuit z,x,and2 );
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA0CarryStr c1,
c2,
c3 -> non
empty strict non
void unsplit gate`1=arity gate`2isBoolean ManySortedSign equals :: GFACIRC1:def 7
(GFA0CarryIStr x,y,z) +* (1GateCircStr <*[<*x,y*>,and2 ],[<*y,z*>,and2 ],[<*z,x*>,and2 ]*>,or3 );
coherence
(GFA0CarryIStr x,y,z) +* (1GateCircStr <*[<*x,y*>,and2 ],[<*y,z*>,and2 ],[<*z,x*>,and2 ]*>,or3 ) is non empty strict non void unsplit gate`1=arity gate`2isBoolean ManySortedSign
;
end;
:: deftheorem Def7 defines GFA0CarryStr GFACIRC1:def 7 :
for
x,
y,
z being
set holds
GFA0CarryStr x,
y,
z = (GFA0CarryIStr x,y,z) +* (1GateCircStr <*[<*x,y*>,and2 ],[<*y,z*>,and2 ],[<*z,x*>,and2 ]*>,or3 );
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA0CarryCirc c1,
c2,
c3 -> strict gate`2=den Boolean Circuit of
GFA0CarryStr a1,
a2,
a3 equals :: GFACIRC1:def 8
(GFA0CarryICirc x,y,z) +* (1GateCircuit [<*x,y*>,and2 ],[<*y,z*>,and2 ],[<*z,x*>,and2 ],or3 );
coherence
(GFA0CarryICirc x,y,z) +* (1GateCircuit [<*x,y*>,and2 ],[<*y,z*>,and2 ],[<*z,x*>,and2 ],or3 ) is strict gate`2=den Boolean Circuit of GFA0CarryStr x,y,z
;
end;
:: deftheorem Def8 defines GFA0CarryCirc GFACIRC1:def 8 :
for
x,
y,
z being
set holds
GFA0CarryCirc x,
y,
z = (GFA0CarryICirc x,y,z) +* (1GateCircuit [<*x,y*>,and2 ],[<*y,z*>,and2 ],[<*z,x*>,and2 ],or3 );
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA0CarryOutput c1,
c2,
c3 -> Element of
InnerVertices (GFA0CarryStr a1,a2,a3) equals :: GFACIRC1:def 9
[<*[<*x,y*>,and2 ],[<*y,z*>,and2 ],[<*z,x*>,and2 ]*>,or3 ];
coherence
[<*[<*x,y*>,and2 ],[<*y,z*>,and2 ],[<*z,x*>,and2 ]*>,or3 ] is Element of InnerVertices (GFA0CarryStr x,y,z)
end;
:: deftheorem Def9 defines GFA0CarryOutput GFACIRC1:def 9 :
for
x,
y,
z being
set holds
GFA0CarryOutput x,
y,
z = [<*[<*x,y*>,and2 ],[<*y,z*>,and2 ],[<*z,x*>,and2 ]*>,or3 ];
theorem Th13: :: GFACIRC1:13
for
x,
y,
z being
set holds
InnerVertices (GFA0CarryIStr x,y,z) = {[<*x,y*>,and2 ],[<*y,z*>,and2 ],[<*z,x*>,and2 ]}
theorem Th14: :: GFACIRC1:14
for
x,
y,
z being
set holds
InnerVertices (GFA0CarryStr x,y,z) = {[<*x,y*>,and2 ],[<*y,z*>,and2 ],[<*z,x*>,and2 ]} \/ {(GFA0CarryOutput x,y,z)}
theorem Th15: :: GFACIRC1:15
theorem Th16: :: GFACIRC1:16
for
x,
y,
z being
set st
x <> [<*y,z*>,and2 ] &
y <> [<*z,x*>,and2 ] &
z <> [<*x,y*>,and2 ] holds
InputVertices (GFA0CarryIStr x,y,z) = {x,y,z}
theorem Th17: :: GFACIRC1:17
for
x,
y,
z being
set st
x <> [<*y,z*>,and2 ] &
y <> [<*z,x*>,and2 ] &
z <> [<*x,y*>,and2 ] holds
InputVertices (GFA0CarryStr x,y,z) = {x,y,z}
theorem Th18: :: GFACIRC1:18
theorem Th19: :: GFACIRC1:19
for
x,
y,
z being
set holds
(
x in the
carrier of
(GFA0CarryStr x,y,z) &
y in the
carrier of
(GFA0CarryStr x,y,z) &
z in the
carrier of
(GFA0CarryStr x,y,z) &
[<*x,y*>,and2 ] in the
carrier of
(GFA0CarryStr x,y,z) &
[<*y,z*>,and2 ] in the
carrier of
(GFA0CarryStr x,y,z) &
[<*z,x*>,and2 ] in the
carrier of
(GFA0CarryStr x,y,z) &
[<*[<*x,y*>,and2 ],[<*y,z*>,and2 ],[<*z,x*>,and2 ]*>,or3 ] in the
carrier of
(GFA0CarryStr x,y,z) )
theorem Th20: :: GFACIRC1:20
for
x,
y,
z being
set holds
(
[<*x,y*>,and2 ] in InnerVertices (GFA0CarryStr x,y,z) &
[<*y,z*>,and2 ] in InnerVertices (GFA0CarryStr x,y,z) &
[<*z,x*>,and2 ] in InnerVertices (GFA0CarryStr x,y,z) &
GFA0CarryOutput x,
y,
z in InnerVertices (GFA0CarryStr x,y,z) )
theorem Th21: :: GFACIRC1:21
for
x,
y,
z being
set st
x <> [<*y,z*>,and2 ] &
y <> [<*z,x*>,and2 ] &
z <> [<*x,y*>,and2 ] holds
(
x in InputVertices (GFA0CarryStr x,y,z) &
y in InputVertices (GFA0CarryStr x,y,z) &
z in InputVertices (GFA0CarryStr x,y,z) )
theorem Th22: :: GFACIRC1:22
theorem Th23: :: GFACIRC1:23
for
x,
y,
z being
set for
s being
State of
(GFA0CarryCirc x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(
(Following s) . [<*x,y*>,and2 ] = a1 '&' a2 &
(Following s) . [<*y,z*>,and2 ] = a2 '&' a3 &
(Following s) . [<*z,x*>,and2 ] = a3 '&' a1 )
theorem Th24: :: GFACIRC1:24
for
x,
y,
z being
set for
s being
State of
(GFA0CarryCirc x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . [<*x,y*>,and2 ] &
a2 = s . [<*y,z*>,and2 ] &
a3 = s . [<*z,x*>,and2 ] holds
(Following s) . (GFA0CarryOutput x,y,z) = (a1 'or' a2) 'or' a3
theorem Th25: :: GFACIRC1:25
for
x,
y,
z being
set st
x <> [<*y,z*>,and2 ] &
y <> [<*z,x*>,and2 ] &
z <> [<*x,y*>,and2 ] holds
for
s being
State of
(GFA0CarryCirc x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(
(Following s,2) . (GFA0CarryOutput x,y,z) = ((a1 '&' a2) 'or' (a2 '&' a3)) 'or' (a3 '&' a1) &
(Following s,2) . [<*x,y*>,and2 ] = a1 '&' a2 &
(Following s,2) . [<*y,z*>,and2 ] = a2 '&' a3 &
(Following s,2) . [<*z,x*>,and2 ] = a3 '&' a1 )
theorem Th26: :: GFACIRC1:26
for
x,
y,
z being
set st
x <> [<*y,z*>,and2 ] &
y <> [<*z,x*>,and2 ] &
z <> [<*x,y*>,and2 ] holds
for
s being
State of
(GFA0CarryCirc x,y,z) holds
Following s,2 is
stable
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA0AdderStr c1,
c2,
c3 -> non
empty strict non
void unsplit gate`1=arity gate`2isBoolean ManySortedSign equals :: GFACIRC1:def 10
2GatesCircStr x,
y,
z,
xor2 ;
coherence
2GatesCircStr x,y,z,xor2 is non empty strict non void unsplit gate`1=arity gate`2isBoolean ManySortedSign
;
end;
:: deftheorem Def10 defines GFA0AdderStr GFACIRC1:def 10 :
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA0AdderCirc c1,
c2,
c3 -> strict gate`2=den Boolean Circuit of
GFA0AdderStr a1,
a2,
a3 equals :: GFACIRC1:def 11
2GatesCircuit x,
y,
z,
xor2 ;
coherence
2GatesCircuit x,y,z,xor2 is strict gate`2=den Boolean Circuit of GFA0AdderStr x,y,z
;
end;
:: deftheorem Def11 defines GFA0AdderCirc GFACIRC1:def 11 :
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA0AdderOutput c1,
c2,
c3 -> Element of
InnerVertices (GFA0AdderStr a1,a2,a3) equals :: GFACIRC1:def 12
2GatesCircOutput x,
y,
z,
xor2 ;
coherence
2GatesCircOutput x,y,z,xor2 is Element of InnerVertices (GFA0AdderStr x,y,z)
;
end;
:: deftheorem Def12 defines GFA0AdderOutput GFACIRC1:def 12 :
theorem Th27: :: GFACIRC1:27
theorem Th28: :: GFACIRC1:28
theorem Th29: :: GFACIRC1:29
theorem Th30: :: GFACIRC1:30
theorem Th31: :: GFACIRC1:31
for
x,
y,
z being
set holds
(
x in the
carrier of
(GFA0AdderStr x,y,z) &
y in the
carrier of
(GFA0AdderStr x,y,z) &
z in the
carrier of
(GFA0AdderStr x,y,z) &
[<*x,y*>,xor2 ] in the
carrier of
(GFA0AdderStr x,y,z) &
[<*[<*x,y*>,xor2 ],z*>,xor2 ] in the
carrier of
(GFA0AdderStr x,y,z) )
by FACIRC_1:60, FACIRC_1:61;
theorem Th32: :: GFACIRC1:32
for
x,
y,
z being
set holds
(
[<*x,y*>,xor2 ] in InnerVertices (GFA0AdderStr x,y,z) &
GFA0AdderOutput x,
y,
z in InnerVertices (GFA0AdderStr x,y,z) )
theorem Th33: :: GFACIRC1:33
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2 ] holds
(
x in InputVertices (GFA0AdderStr x,y,z) &
y in InputVertices (GFA0AdderStr x,y,z) &
z in InputVertices (GFA0AdderStr x,y,z) )
theorem Th34: :: GFACIRC1:34
theorem Th35: :: GFACIRC1:35
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2 ] holds
for
s being
State of
(GFA0AdderCirc x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(
(Following s) . [<*x,y*>,xor2 ] = a1 'xor' a2 &
(Following s) . x = a1 &
(Following s) . y = a2 &
(Following s) . z = a3 )
theorem Th36: :: GFACIRC1:36
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2 ] holds
for
s being
State of
(GFA0AdderCirc x,y,z) for
a1a2,
a1,
a2,
a3 being
Element of
BOOLEAN st
a1a2 = s . [<*x,y*>,xor2 ] &
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(Following s) . (GFA0AdderOutput x,y,z) = a1a2 'xor' a3
theorem Th37: :: GFACIRC1:37
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2 ] holds
for
s being
State of
(GFA0AdderCirc x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(
(Following s,2) . (GFA0AdderOutput x,y,z) = (a1 'xor' a2) 'xor' a3 &
(Following s,2) . [<*x,y*>,xor2 ] = a1 'xor' a2 &
(Following s,2) . x = a1 &
(Following s,2) . y = a2 &
(Following s,2) . z = a3 )
theorem Th38: :: GFACIRC1:38
definition
let x be
set ,
y be
set ,
z be
set ;
func BitGFA0Str c1,
c2,
c3 -> non
empty strict non
void unsplit gate`1=arity gate`2isBoolean ManySortedSign equals :: GFACIRC1:def 13
(GFA0AdderStr x,y,z) +* (GFA0CarryStr x,y,z);
coherence
(GFA0AdderStr x,y,z) +* (GFA0CarryStr x,y,z) is non empty strict non void unsplit gate`1=arity gate`2isBoolean ManySortedSign
;
end;
:: deftheorem Def13 defines BitGFA0Str GFACIRC1:def 13 :
definition
let x be
set ,
y be
set ,
z be
set ;
func BitGFA0Circ c1,
c2,
c3 -> strict gate`2=den Boolean Circuit of
BitGFA0Str a1,
a2,
a3 equals :: GFACIRC1:def 14
(GFA0AdderCirc x,y,z) +* (GFA0CarryCirc x,y,z);
coherence
(GFA0AdderCirc x,y,z) +* (GFA0CarryCirc x,y,z) is strict gate`2=den Boolean Circuit of BitGFA0Str x,y,z
;
end;
:: deftheorem Def14 defines BitGFA0Circ GFACIRC1:def 14 :
theorem Th39: :: GFACIRC1:39
for
x,
y,
z being
set holds
InnerVertices (BitGFA0Str x,y,z) = (({[<*x,y*>,xor2 ]} \/ {(GFA0AdderOutput x,y,z)}) \/ {[<*x,y*>,and2 ],[<*y,z*>,and2 ],[<*z,x*>,and2 ]}) \/ {(GFA0CarryOutput x,y,z)}
theorem Th40: :: GFACIRC1:40
theorem Th41: :: GFACIRC1:41
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2 ] &
x <> [<*y,z*>,and2 ] &
y <> [<*z,x*>,and2 ] &
z <> [<*x,y*>,and2 ] holds
InputVertices (BitGFA0Str x,y,z) = {x,y,z}
theorem Th42: :: GFACIRC1:42
theorem Th43: :: GFACIRC1:43
theorem Th44: :: GFACIRC1:44
for
x,
y,
z being
set holds
(
x in the
carrier of
(BitGFA0Str x,y,z) &
y in the
carrier of
(BitGFA0Str x,y,z) &
z in the
carrier of
(BitGFA0Str x,y,z) &
[<*x,y*>,xor2 ] in the
carrier of
(BitGFA0Str x,y,z) &
[<*[<*x,y*>,xor2 ],z*>,xor2 ] in the
carrier of
(BitGFA0Str x,y,z) &
[<*x,y*>,and2 ] in the
carrier of
(BitGFA0Str x,y,z) &
[<*y,z*>,and2 ] in the
carrier of
(BitGFA0Str x,y,z) &
[<*z,x*>,and2 ] in the
carrier of
(BitGFA0Str x,y,z) &
[<*[<*x,y*>,and2 ],[<*y,z*>,and2 ],[<*z,x*>,and2 ]*>,or3 ] in the
carrier of
(BitGFA0Str x,y,z) )
theorem Th45: :: GFACIRC1:45
for
x,
y,
z being
set holds
(
[<*x,y*>,xor2 ] in InnerVertices (BitGFA0Str x,y,z) &
GFA0AdderOutput x,
y,
z in InnerVertices (BitGFA0Str x,y,z) &
[<*x,y*>,and2 ] in InnerVertices (BitGFA0Str x,y,z) &
[<*y,z*>,and2 ] in InnerVertices (BitGFA0Str x,y,z) &
[<*z,x*>,and2 ] in InnerVertices (BitGFA0Str x,y,z) &
GFA0CarryOutput x,
y,
z in InnerVertices (BitGFA0Str x,y,z) )
theorem Th46: :: GFACIRC1:46
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2 ] &
x <> [<*y,z*>,and2 ] &
y <> [<*z,x*>,and2 ] &
z <> [<*x,y*>,and2 ] holds
(
x in InputVertices (BitGFA0Str x,y,z) &
y in InputVertices (BitGFA0Str x,y,z) &
z in InputVertices (BitGFA0Str x,y,z) )
definition
let x be
set ,
y be
set ,
z be
set ;
func BitGFA0CarryOutput c1,
c2,
c3 -> Element of
InnerVertices (BitGFA0Str a1,a2,a3) equals :: GFACIRC1:def 15
[<*[<*x,y*>,and2 ],[<*y,z*>,and2 ],[<*z,x*>,and2 ]*>,or3 ];
coherence
[<*[<*x,y*>,and2 ],[<*y,z*>,and2 ],[<*z,x*>,and2 ]*>,or3 ] is Element of InnerVertices (BitGFA0Str x,y,z)
end;
:: deftheorem Def15 defines BitGFA0CarryOutput GFACIRC1:def 15 :
for
x,
y,
z being
set holds
BitGFA0CarryOutput x,
y,
z = [<*[<*x,y*>,and2 ],[<*y,z*>,and2 ],[<*z,x*>,and2 ]*>,or3 ];
definition
let x be
set ,
y be
set ,
z be
set ;
func BitGFA0AdderOutput c1,
c2,
c3 -> Element of
InnerVertices (BitGFA0Str a1,a2,a3) equals :: GFACIRC1:def 16
2GatesCircOutput x,
y,
z,
xor2 ;
coherence
2GatesCircOutput x,y,z,xor2 is Element of InnerVertices (BitGFA0Str x,y,z)
end;
:: deftheorem Def16 defines BitGFA0AdderOutput GFACIRC1:def 16 :
theorem Th47: :: GFACIRC1:47
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2 ] &
x <> [<*y,z*>,and2 ] &
y <> [<*z,x*>,and2 ] &
z <> [<*x,y*>,and2 ] holds
for
s being
State of
(BitGFA0Circ x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(
(Following s,2) . (GFA0AdderOutput x,y,z) = (a1 'xor' a2) 'xor' a3 &
(Following s,2) . (GFA0CarryOutput x,y,z) = ((a1 '&' a2) 'or' (a2 '&' a3)) 'or' (a3 '&' a1) )
theorem Th48: :: GFACIRC1:48
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2 ] &
x <> [<*y,z*>,and2 ] &
y <> [<*z,x*>,and2 ] &
z <> [<*x,y*>,and2 ] holds
for
s being
State of
(BitGFA0Circ x,y,z) holds
Following s,2 is
stable
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA1CarryIStr c1,
c2,
c3 -> non
empty strict non
void unsplit gate`1=arity gate`2isBoolean ManySortedSign equals :: GFACIRC1:def 17
((1GateCircStr <*x,y*>,and2c ) +* (1GateCircStr <*y,z*>,and2a )) +* (1GateCircStr <*z,x*>,and2 );
coherence
((1GateCircStr <*x,y*>,and2c ) +* (1GateCircStr <*y,z*>,and2a )) +* (1GateCircStr <*z,x*>,and2 ) is non empty strict non void unsplit gate`1=arity gate`2isBoolean ManySortedSign
;
end;
:: deftheorem Def17 defines GFA1CarryIStr GFACIRC1:def 17 :
for
x,
y,
z being
set holds
GFA1CarryIStr x,
y,
z = ((1GateCircStr <*x,y*>,and2c ) +* (1GateCircStr <*y,z*>,and2a )) +* (1GateCircStr <*z,x*>,and2 );
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA1CarryICirc c1,
c2,
c3 -> strict gate`2=den Boolean Circuit of
GFA1CarryIStr a1,
a2,
a3 equals :: GFACIRC1:def 18
((1GateCircuit x,y,and2c ) +* (1GateCircuit y,z,and2a )) +* (1GateCircuit z,x,and2 );
coherence
((1GateCircuit x,y,and2c ) +* (1GateCircuit y,z,and2a )) +* (1GateCircuit z,x,and2 ) is strict gate`2=den Boolean Circuit of GFA1CarryIStr x,y,z
;
end;
:: deftheorem Def18 defines GFA1CarryICirc GFACIRC1:def 18 :
for
x,
y,
z being
set holds
GFA1CarryICirc x,
y,
z = ((1GateCircuit x,y,and2c ) +* (1GateCircuit y,z,and2a )) +* (1GateCircuit z,x,and2 );
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA1CarryStr c1,
c2,
c3 -> non
empty strict non
void unsplit gate`1=arity gate`2isBoolean ManySortedSign equals :: GFACIRC1:def 19
(GFA1CarryIStr x,y,z) +* (1GateCircStr <*[<*x,y*>,and2c ],[<*y,z*>,and2a ],[<*z,x*>,and2 ]*>,or3 );
coherence
(GFA1CarryIStr x,y,z) +* (1GateCircStr <*[<*x,y*>,and2c ],[<*y,z*>,and2a ],[<*z,x*>,and2 ]*>,or3 ) is non empty strict non void unsplit gate`1=arity gate`2isBoolean ManySortedSign
;
end;
:: deftheorem Def19 defines GFA1CarryStr GFACIRC1:def 19 :
for
x,
y,
z being
set holds
GFA1CarryStr x,
y,
z = (GFA1CarryIStr x,y,z) +* (1GateCircStr <*[<*x,y*>,and2c ],[<*y,z*>,and2a ],[<*z,x*>,and2 ]*>,or3 );
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA1CarryCirc c1,
c2,
c3 -> strict gate`2=den Boolean Circuit of
GFA1CarryStr a1,
a2,
a3 equals :: GFACIRC1:def 20
(GFA1CarryICirc x,y,z) +* (1GateCircuit [<*x,y*>,and2c ],[<*y,z*>,and2a ],[<*z,x*>,and2 ],or3 );
coherence
(GFA1CarryICirc x,y,z) +* (1GateCircuit [<*x,y*>,and2c ],[<*y,z*>,and2a ],[<*z,x*>,and2 ],or3 ) is strict gate`2=den Boolean Circuit of GFA1CarryStr x,y,z
;
end;
:: deftheorem Def20 defines GFA1CarryCirc GFACIRC1:def 20 :
for
x,
y,
z being
set holds
GFA1CarryCirc x,
y,
z = (GFA1CarryICirc x,y,z) +* (1GateCircuit [<*x,y*>,and2c ],[<*y,z*>,and2a ],[<*z,x*>,and2 ],or3 );
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA1CarryOutput c1,
c2,
c3 -> Element of
InnerVertices (GFA1CarryStr a1,a2,a3) equals :: GFACIRC1:def 21
[<*[<*x,y*>,and2c ],[<*y,z*>,and2a ],[<*z,x*>,and2 ]*>,or3 ];
coherence
[<*[<*x,y*>,and2c ],[<*y,z*>,and2a ],[<*z,x*>,and2 ]*>,or3 ] is Element of InnerVertices (GFA1CarryStr x,y,z)
end;
:: deftheorem Def21 defines GFA1CarryOutput GFACIRC1:def 21 :
for
x,
y,
z being
set holds
GFA1CarryOutput x,
y,
z = [<*[<*x,y*>,and2c ],[<*y,z*>,and2a ],[<*z,x*>,and2 ]*>,or3 ];
theorem Th49: :: GFACIRC1:49
for
x,
y,
z being
set holds
InnerVertices (GFA1CarryIStr x,y,z) = {[<*x,y*>,and2c ],[<*y,z*>,and2a ],[<*z,x*>,and2 ]}
theorem Th50: :: GFACIRC1:50
for
x,
y,
z being
set holds
InnerVertices (GFA1CarryStr x,y,z) = {[<*x,y*>,and2c ],[<*y,z*>,and2a ],[<*z,x*>,and2 ]} \/ {(GFA1CarryOutput x,y,z)}
theorem Th51: :: GFACIRC1:51
theorem Th52: :: GFACIRC1:52
for
x,
y,
z being
set st
x <> [<*y,z*>,and2a ] &
y <> [<*z,x*>,and2 ] &
z <> [<*x,y*>,and2c ] holds
InputVertices (GFA1CarryIStr x,y,z) = {x,y,z}
theorem Th53: :: GFACIRC1:53
for
x,
y,
z being
set st
x <> [<*y,z*>,and2a ] &
y <> [<*z,x*>,and2 ] &
z <> [<*x,y*>,and2c ] holds
InputVertices (GFA1CarryStr x,y,z) = {x,y,z}
theorem Th54: :: GFACIRC1:54
theorem Th55: :: GFACIRC1:55
for
x,
y,
z being
set holds
(
x in the
carrier of
(GFA1CarryStr x,y,z) &
y in the
carrier of
(GFA1CarryStr x,y,z) &
z in the
carrier of
(GFA1CarryStr x,y,z) &
[<*x,y*>,and2c ] in the
carrier of
(GFA1CarryStr x,y,z) &
[<*y,z*>,and2a ] in the
carrier of
(GFA1CarryStr x,y,z) &
[<*z,x*>,and2 ] in the
carrier of
(GFA1CarryStr x,y,z) &
[<*[<*x,y*>,and2c ],[<*y,z*>,and2a ],[<*z,x*>,and2 ]*>,or3 ] in the
carrier of
(GFA1CarryStr x,y,z) )
theorem Th56: :: GFACIRC1:56
for
x,
y,
z being
set holds
(
[<*x,y*>,and2c ] in InnerVertices (GFA1CarryStr x,y,z) &
[<*y,z*>,and2a ] in InnerVertices (GFA1CarryStr x,y,z) &
[<*z,x*>,and2 ] in InnerVertices (GFA1CarryStr x,y,z) &
GFA1CarryOutput x,
y,
z in InnerVertices (GFA1CarryStr x,y,z) )
theorem Th57: :: GFACIRC1:57
for
x,
y,
z being
set st
x <> [<*y,z*>,and2a ] &
y <> [<*z,x*>,and2 ] &
z <> [<*x,y*>,and2c ] holds
(
x in InputVertices (GFA1CarryStr x,y,z) &
y in InputVertices (GFA1CarryStr x,y,z) &
z in InputVertices (GFA1CarryStr x,y,z) )
theorem Th58: :: GFACIRC1:58
theorem Th59: :: GFACIRC1:59
for
x,
y,
z being
set for
s being
State of
(GFA1CarryCirc x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(
(Following s) . [<*x,y*>,and2c ] = a1 '&' ('not' a2) &
(Following s) . [<*y,z*>,and2a ] = ('not' a2) '&' a3 &
(Following s) . [<*z,x*>,and2 ] = a3 '&' a1 )
theorem Th60: :: GFACIRC1:60
for
x,
y,
z being
set for
s being
State of
(GFA1CarryCirc x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . [<*x,y*>,and2c ] &
a2 = s . [<*y,z*>,and2a ] &
a3 = s . [<*z,x*>,and2 ] holds
(Following s) . (GFA1CarryOutput x,y,z) = (a1 'or' a2) 'or' a3
theorem Th61: :: GFACIRC1:61
for
x,
y,
z being
set st
x <> [<*y,z*>,and2a ] &
y <> [<*z,x*>,and2 ] &
z <> [<*x,y*>,and2c ] holds
for
s being
State of
(GFA1CarryCirc x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(
(Following s,2) . (GFA1CarryOutput x,y,z) = ((a1 '&' ('not' a2)) 'or' (('not' a2) '&' a3)) 'or' (a3 '&' a1) &
(Following s,2) . [<*x,y*>,and2c ] = a1 '&' ('not' a2) &
(Following s,2) . [<*y,z*>,and2a ] = ('not' a2) '&' a3 &
(Following s,2) . [<*z,x*>,and2 ] = a3 '&' a1 )
theorem Th62: :: GFACIRC1:62
for
x,
y,
z being
set st
x <> [<*y,z*>,and2a ] &
y <> [<*z,x*>,and2 ] &
z <> [<*x,y*>,and2c ] holds
for
s being
State of
(GFA1CarryCirc x,y,z) holds
Following s,2 is
stable
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA1AdderStr c1,
c2,
c3 -> non
empty strict non
void unsplit gate`1=arity gate`2isBoolean ManySortedSign equals :: GFACIRC1:def 22
2GatesCircStr x,
y,
z,
xor2c ;
coherence
2GatesCircStr x,y,z,xor2c is non empty strict non void unsplit gate`1=arity gate`2isBoolean ManySortedSign
;
end;
:: deftheorem Def22 defines GFA1AdderStr GFACIRC1:def 22 :
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA1AdderCirc c1,
c2,
c3 -> strict gate`2=den Boolean Circuit of
GFA1AdderStr a1,
a2,
a3 equals :: GFACIRC1:def 23
2GatesCircuit x,
y,
z,
xor2c ;
coherence
2GatesCircuit x,y,z,xor2c is strict gate`2=den Boolean Circuit of GFA1AdderStr x,y,z
;
end;
:: deftheorem Def23 defines GFA1AdderCirc GFACIRC1:def 23 :
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA1AdderOutput c1,
c2,
c3 -> Element of
InnerVertices (GFA1AdderStr a1,a2,a3) equals :: GFACIRC1:def 24
2GatesCircOutput x,
y,
z,
xor2c ;
coherence
2GatesCircOutput x,y,z,xor2c is Element of InnerVertices (GFA1AdderStr x,y,z)
;
end;
:: deftheorem Def24 defines GFA1AdderOutput GFACIRC1:def 24 :
theorem Th63: :: GFACIRC1:63
theorem Th64: :: GFACIRC1:64
theorem Th65: :: GFACIRC1:65
theorem Th66: :: GFACIRC1:66
theorem Th67: :: GFACIRC1:67
for
x,
y,
z being
set holds
(
x in the
carrier of
(GFA1AdderStr x,y,z) &
y in the
carrier of
(GFA1AdderStr x,y,z) &
z in the
carrier of
(GFA1AdderStr x,y,z) &
[<*x,y*>,xor2c ] in the
carrier of
(GFA1AdderStr x,y,z) &
[<*[<*x,y*>,xor2c ],z*>,xor2c ] in the
carrier of
(GFA1AdderStr x,y,z) )
by FACIRC_1:60, FACIRC_1:61;
theorem Th68: :: GFACIRC1:68
for
x,
y,
z being
set holds
(
[<*x,y*>,xor2c ] in InnerVertices (GFA1AdderStr x,y,z) &
GFA1AdderOutput x,
y,
z in InnerVertices (GFA1AdderStr x,y,z) )
theorem Th69: :: GFACIRC1:69
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2c ] holds
(
x in InputVertices (GFA1AdderStr x,y,z) &
y in InputVertices (GFA1AdderStr x,y,z) &
z in InputVertices (GFA1AdderStr x,y,z) )
theorem Th70: :: GFACIRC1:70
theorem Th71: :: GFACIRC1:71
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2c ] holds
for
s being
State of
(GFA1AdderCirc x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(
(Following s) . [<*x,y*>,xor2c ] = a1 'xor' ('not' a2) &
(Following s) . x = a1 &
(Following s) . y = a2 &
(Following s) . z = a3 )
theorem Th72: :: GFACIRC1:72
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2c ] holds
for
s being
State of
(GFA1AdderCirc x,y,z) for
a1a2,
a1,
a2,
a3 being
Element of
BOOLEAN st
a1a2 = s . [<*x,y*>,xor2c ] &
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(Following s) . (GFA1AdderOutput x,y,z) = a1a2 'xor' ('not' a3)
theorem Th73: :: GFACIRC1:73
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2c ] holds
for
s being
State of
(GFA1AdderCirc x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(
(Following s,2) . (GFA1AdderOutput x,y,z) = (a1 'xor' ('not' a2)) 'xor' ('not' a3) &
(Following s,2) . [<*x,y*>,xor2c ] = a1 'xor' ('not' a2) &
(Following s,2) . x = a1 &
(Following s,2) . y = a2 &
(Following s,2) . z = a3 )
theorem Th74: :: GFACIRC1:74
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2c ] holds
for
s being
State of
(GFA1AdderCirc x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(Following s,2) . (GFA1AdderOutput x,y,z) = 'not' ((a1 'xor' ('not' a2)) 'xor' a3)
theorem Th75: :: GFACIRC1:75
definition
let x be
set ,
y be
set ,
z be
set ;
func BitGFA1Str c1,
c2,
c3 -> non
empty strict non
void unsplit gate`1=arity gate`2isBoolean ManySortedSign equals :: GFACIRC1:def 25
(GFA1AdderStr x,y,z) +* (GFA1CarryStr x,y,z);
coherence
(GFA1AdderStr x,y,z) +* (GFA1CarryStr x,y,z) is non empty strict non void unsplit gate`1=arity gate`2isBoolean ManySortedSign
;
end;
:: deftheorem Def25 defines BitGFA1Str GFACIRC1:def 25 :
definition
let x be
set ,
y be
set ,
z be
set ;
func BitGFA1Circ c1,
c2,
c3 -> strict gate`2=den Boolean Circuit of
BitGFA1Str a1,
a2,
a3 equals :: GFACIRC1:def 26
(GFA1AdderCirc x,y,z) +* (GFA1CarryCirc x,y,z);
coherence
(GFA1AdderCirc x,y,z) +* (GFA1CarryCirc x,y,z) is strict gate`2=den Boolean Circuit of BitGFA1Str x,y,z
;
end;
:: deftheorem Def26 defines BitGFA1Circ GFACIRC1:def 26 :
theorem Th76: :: GFACIRC1:76
for
x,
y,
z being
set holds
InnerVertices (BitGFA1Str x,y,z) = (({[<*x,y*>,xor2c ]} \/ {(GFA1AdderOutput x,y,z)}) \/ {[<*x,y*>,and2c ],[<*y,z*>,and2a ],[<*z,x*>,and2 ]}) \/ {(GFA1CarryOutput x,y,z)}
theorem Th77: :: GFACIRC1:77
theorem Th78: :: GFACIRC1:78
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2c ] &
x <> [<*y,z*>,and2a ] &
y <> [<*z,x*>,and2 ] &
z <> [<*x,y*>,and2c ] holds
InputVertices (BitGFA1Str x,y,z) = {x,y,z}
theorem Th79: :: GFACIRC1:79
theorem Th80: :: GFACIRC1:80
theorem Th81: :: GFACIRC1:81
for
x,
y,
z being
set holds
(
x in the
carrier of
(BitGFA1Str x,y,z) &
y in the
carrier of
(BitGFA1Str x,y,z) &
z in the
carrier of
(BitGFA1Str x,y,z) &
[<*x,y*>,xor2c ] in the
carrier of
(BitGFA1Str x,y,z) &
[<*[<*x,y*>,xor2c ],z*>,xor2c ] in the
carrier of
(BitGFA1Str x,y,z) &
[<*x,y*>,and2c ] in the
carrier of
(BitGFA1Str x,y,z) &
[<*y,z*>,and2a ] in the
carrier of
(BitGFA1Str x,y,z) &
[<*z,x*>,and2 ] in the
carrier of
(BitGFA1Str x,y,z) &
[<*[<*x,y*>,and2c ],[<*y,z*>,and2a ],[<*z,x*>,and2 ]*>,or3 ] in the
carrier of
(BitGFA1Str x,y,z) )
theorem Th82: :: GFACIRC1:82
for
x,
y,
z being
set holds
(
[<*x,y*>,xor2c ] in InnerVertices (BitGFA1Str x,y,z) &
GFA1AdderOutput x,
y,
z in InnerVertices (BitGFA1Str x,y,z) &
[<*x,y*>,and2c ] in InnerVertices (BitGFA1Str x,y,z) &
[<*y,z*>,and2a ] in InnerVertices (BitGFA1Str x,y,z) &
[<*z,x*>,and2 ] in InnerVertices (BitGFA1Str x,y,z) &
GFA1CarryOutput x,
y,
z in InnerVertices (BitGFA1Str x,y,z) )
theorem Th83: :: GFACIRC1:83
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2c ] &
x <> [<*y,z*>,and2a ] &
y <> [<*z,x*>,and2 ] &
z <> [<*x,y*>,and2c ] holds
(
x in InputVertices (BitGFA1Str x,y,z) &
y in InputVertices (BitGFA1Str x,y,z) &
z in InputVertices (BitGFA1Str x,y,z) )
definition
let x be
set ,
y be
set ,
z be
set ;
func BitGFA1CarryOutput c1,
c2,
c3 -> Element of
InnerVertices (BitGFA1Str a1,a2,a3) equals :: GFACIRC1:def 27
[<*[<*x,y*>,and2c ],[<*y,z*>,and2a ],[<*z,x*>,and2 ]*>,or3 ];
coherence
[<*[<*x,y*>,and2c ],[<*y,z*>,and2a ],[<*z,x*>,and2 ]*>,or3 ] is Element of InnerVertices (BitGFA1Str x,y,z)
end;
:: deftheorem Def27 defines BitGFA1CarryOutput GFACIRC1:def 27 :
for
x,
y,
z being
set holds
BitGFA1CarryOutput x,
y,
z = [<*[<*x,y*>,and2c ],[<*y,z*>,and2a ],[<*z,x*>,and2 ]*>,or3 ];
definition
let x be
set ,
y be
set ,
z be
set ;
func BitGFA1AdderOutput c1,
c2,
c3 -> Element of
InnerVertices (BitGFA1Str a1,a2,a3) equals :: GFACIRC1:def 28
2GatesCircOutput x,
y,
z,
xor2c ;
coherence
2GatesCircOutput x,y,z,xor2c is Element of InnerVertices (BitGFA1Str x,y,z)
end;
:: deftheorem Def28 defines BitGFA1AdderOutput GFACIRC1:def 28 :
theorem Th84: :: GFACIRC1:84
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2c ] &
x <> [<*y,z*>,and2a ] &
y <> [<*z,x*>,and2 ] &
z <> [<*x,y*>,and2c ] holds
for
s being
State of
(BitGFA1Circ x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(
(Following s,2) . (GFA1AdderOutput x,y,z) = 'not' ((a1 'xor' ('not' a2)) 'xor' a3) &
(Following s,2) . (GFA1CarryOutput x,y,z) = ((a1 '&' ('not' a2)) 'or' (('not' a2) '&' a3)) 'or' (a3 '&' a1) )
theorem Th85: :: GFACIRC1:85
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2c ] &
x <> [<*y,z*>,and2a ] &
y <> [<*z,x*>,and2 ] &
z <> [<*x,y*>,and2c ] holds
for
s being
State of
(BitGFA1Circ x,y,z) holds
Following s,2 is
stable
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA2CarryIStr c1,
c2,
c3 -> non
empty strict non
void unsplit gate`1=arity gate`2isBoolean ManySortedSign equals :: GFACIRC1:def 29
((1GateCircStr <*x,y*>,and2a ) +* (1GateCircStr <*y,z*>,and2c )) +* (1GateCircStr <*z,x*>,and2b );
coherence
((1GateCircStr <*x,y*>,and2a ) +* (1GateCircStr <*y,z*>,and2c )) +* (1GateCircStr <*z,x*>,and2b ) is non empty strict non void unsplit gate`1=arity gate`2isBoolean ManySortedSign
;
end;
:: deftheorem Def29 defines GFA2CarryIStr GFACIRC1:def 29 :
for
x,
y,
z being
set holds
GFA2CarryIStr x,
y,
z = ((1GateCircStr <*x,y*>,and2a ) +* (1GateCircStr <*y,z*>,and2c )) +* (1GateCircStr <*z,x*>,and2b );
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA2CarryICirc c1,
c2,
c3 -> strict gate`2=den Boolean Circuit of
GFA2CarryIStr a1,
a2,
a3 equals :: GFACIRC1:def 30
((1GateCircuit x,y,and2a ) +* (1GateCircuit y,z,and2c )) +* (1GateCircuit z,x,and2b );
coherence
((1GateCircuit x,y,and2a ) +* (1GateCircuit y,z,and2c )) +* (1GateCircuit z,x,and2b ) is strict gate`2=den Boolean Circuit of GFA2CarryIStr x,y,z
;
end;
:: deftheorem Def30 defines GFA2CarryICirc GFACIRC1:def 30 :
for
x,
y,
z being
set holds
GFA2CarryICirc x,
y,
z = ((1GateCircuit x,y,and2a ) +* (1GateCircuit y,z,and2c )) +* (1GateCircuit z,x,and2b );
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA2CarryStr c1,
c2,
c3 -> non
empty strict non
void unsplit gate`1=arity gate`2isBoolean ManySortedSign equals :: GFACIRC1:def 31
(GFA2CarryIStr x,y,z) +* (1GateCircStr <*[<*x,y*>,and2a ],[<*y,z*>,and2c ],[<*z,x*>,and2b ]*>,nor3 );
coherence
(GFA2CarryIStr x,y,z) +* (1GateCircStr <*[<*x,y*>,and2a ],[<*y,z*>,and2c ],[<*z,x*>,and2b ]*>,nor3 ) is non empty strict non void unsplit gate`1=arity gate`2isBoolean ManySortedSign
;
end;
:: deftheorem Def31 defines GFA2CarryStr GFACIRC1:def 31 :
for
x,
y,
z being
set holds
GFA2CarryStr x,
y,
z = (GFA2CarryIStr x,y,z) +* (1GateCircStr <*[<*x,y*>,and2a ],[<*y,z*>,and2c ],[<*z,x*>,and2b ]*>,nor3 );
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA2CarryCirc c1,
c2,
c3 -> strict gate`2=den Boolean Circuit of
GFA2CarryStr a1,
a2,
a3 equals :: GFACIRC1:def 32
(GFA2CarryICirc x,y,z) +* (1GateCircuit [<*x,y*>,and2a ],[<*y,z*>,and2c ],[<*z,x*>,and2b ],nor3 );
coherence
(GFA2CarryICirc x,y,z) +* (1GateCircuit [<*x,y*>,and2a ],[<*y,z*>,and2c ],[<*z,x*>,and2b ],nor3 ) is strict gate`2=den Boolean Circuit of GFA2CarryStr x,y,z
;
end;
:: deftheorem Def32 defines GFA2CarryCirc GFACIRC1:def 32 :
for
x,
y,
z being
set holds
GFA2CarryCirc x,
y,
z = (GFA2CarryICirc x,y,z) +* (1GateCircuit [<*x,y*>,and2a ],[<*y,z*>,and2c ],[<*z,x*>,and2b ],nor3 );
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA2CarryOutput c1,
c2,
c3 -> Element of
InnerVertices (GFA2CarryStr a1,a2,a3) equals :: GFACIRC1:def 33
[<*[<*x,y*>,and2a ],[<*y,z*>,and2c ],[<*z,x*>,and2b ]*>,nor3 ];
coherence
[<*[<*x,y*>,and2a ],[<*y,z*>,and2c ],[<*z,x*>,and2b ]*>,nor3 ] is Element of InnerVertices (GFA2CarryStr x,y,z)
end;
:: deftheorem Def33 defines GFA2CarryOutput GFACIRC1:def 33 :
for
x,
y,
z being
set holds
GFA2CarryOutput x,
y,
z = [<*[<*x,y*>,and2a ],[<*y,z*>,and2c ],[<*z,x*>,and2b ]*>,nor3 ];
theorem Th86: :: GFACIRC1:86
for
x,
y,
z being
set holds
InnerVertices (GFA2CarryIStr x,y,z) = {[<*x,y*>,and2a ],[<*y,z*>,and2c ],[<*z,x*>,and2b ]}
theorem Th87: :: GFACIRC1:87
for
x,
y,
z being
set holds
InnerVertices (GFA2CarryStr x,y,z) = {[<*x,y*>,and2a ],[<*y,z*>,and2c ],[<*z,x*>,and2b ]} \/ {(GFA2CarryOutput x,y,z)}
theorem Th88: :: GFACIRC1:88
theorem Th89: :: GFACIRC1:89
for
x,
y,
z being
set st
x <> [<*y,z*>,and2c ] &
y <> [<*z,x*>,and2b ] &
z <> [<*x,y*>,and2a ] holds
InputVertices (GFA2CarryIStr x,y,z) = {x,y,z}
theorem Th90: :: GFACIRC1:90
for
x,
y,
z being
set st
x <> [<*y,z*>,and2c ] &
y <> [<*z,x*>,and2b ] &
z <> [<*x,y*>,and2a ] holds
InputVertices (GFA2CarryStr x,y,z) = {x,y,z}
theorem Th91: :: GFACIRC1:91
theorem Th92: :: GFACIRC1:92
for
x,
y,
z being
set holds
(
x in the
carrier of
(GFA2CarryStr x,y,z) &
y in the
carrier of
(GFA2CarryStr x,y,z) &
z in the
carrier of
(GFA2CarryStr x,y,z) &
[<*x,y*>,and2a ] in the
carrier of
(GFA2CarryStr x,y,z) &
[<*y,z*>,and2c ] in the
carrier of
(GFA2CarryStr x,y,z) &
[<*z,x*>,and2b ] in the
carrier of
(GFA2CarryStr x,y,z) &
[<*[<*x,y*>,and2a ],[<*y,z*>,and2c ],[<*z,x*>,and2b ]*>,nor3 ] in the
carrier of
(GFA2CarryStr x,y,z) )
theorem Th93: :: GFACIRC1:93
for
x,
y,
z being
set holds
(
[<*x,y*>,and2a ] in InnerVertices (GFA2CarryStr x,y,z) &
[<*y,z*>,and2c ] in InnerVertices (GFA2CarryStr x,y,z) &
[<*z,x*>,and2b ] in InnerVertices (GFA2CarryStr x,y,z) &
GFA2CarryOutput x,
y,
z in InnerVertices (GFA2CarryStr x,y,z) )
theorem Th94: :: GFACIRC1:94
for
x,
y,
z being
set st
x <> [<*y,z*>,and2c ] &
y <> [<*z,x*>,and2b ] &
z <> [<*x,y*>,and2a ] holds
(
x in InputVertices (GFA2CarryStr x,y,z) &
y in InputVertices (GFA2CarryStr x,y,z) &
z in InputVertices (GFA2CarryStr x,y,z) )
theorem Th95: :: GFACIRC1:95
theorem Th96: :: GFACIRC1:96
for
x,
y,
z being
set for
s being
State of
(GFA2CarryCirc x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(
(Following s) . [<*x,y*>,and2a ] = ('not' a1) '&' a2 &
(Following s) . [<*y,z*>,and2c ] = a2 '&' ('not' a3) &
(Following s) . [<*z,x*>,and2b ] = ('not' a3) '&' ('not' a1) )
theorem Th97: :: GFACIRC1:97
for
x,
y,
z being
set for
s being
State of
(GFA2CarryCirc x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . [<*x,y*>,and2a ] &
a2 = s . [<*y,z*>,and2c ] &
a3 = s . [<*z,x*>,and2b ] holds
(Following s) . (GFA2CarryOutput x,y,z) = 'not' ((a1 'or' a2) 'or' a3)
theorem Th98: :: GFACIRC1:98
for
x,
y,
z being
set st
x <> [<*y,z*>,and2c ] &
y <> [<*z,x*>,and2b ] &
z <> [<*x,y*>,and2a ] holds
for
s being
State of
(GFA2CarryCirc x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(
(Following s,2) . (GFA2CarryOutput x,y,z) = 'not' (((('not' a1) '&' a2) 'or' (a2 '&' ('not' a3))) 'or' (('not' a3) '&' ('not' a1))) &
(Following s,2) . [<*x,y*>,and2a ] = ('not' a1) '&' a2 &
(Following s,2) . [<*y,z*>,and2c ] = a2 '&' ('not' a3) &
(Following s,2) . [<*z,x*>,and2b ] = ('not' a3) '&' ('not' a1) )
theorem Th99: :: GFACIRC1:99
for
x,
y,
z being
set st
x <> [<*y,z*>,and2c ] &
y <> [<*z,x*>,and2b ] &
z <> [<*x,y*>,and2a ] holds
for
s being
State of
(GFA2CarryCirc x,y,z) holds
Following s,2 is
stable
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA2AdderStr c1,
c2,
c3 -> non
empty strict non
void unsplit gate`1=arity gate`2isBoolean ManySortedSign equals :: GFACIRC1:def 34
2GatesCircStr x,
y,
z,
xor2c ;
coherence
2GatesCircStr x,y,z,xor2c is non empty strict non void unsplit gate`1=arity gate`2isBoolean ManySortedSign
;
end;
:: deftheorem Def34 defines GFA2AdderStr GFACIRC1:def 34 :
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA2AdderCirc c1,
c2,
c3 -> strict gate`2=den Boolean Circuit of
GFA2AdderStr a1,
a2,
a3 equals :: GFACIRC1:def 35
2GatesCircuit x,
y,
z,
xor2c ;
coherence
2GatesCircuit x,y,z,xor2c is strict gate`2=den Boolean Circuit of GFA2AdderStr x,y,z
;
end;
:: deftheorem Def35 defines GFA2AdderCirc GFACIRC1:def 35 :
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA2AdderOutput c1,
c2,
c3 -> Element of
InnerVertices (GFA2AdderStr a1,a2,a3) equals :: GFACIRC1:def 36
2GatesCircOutput x,
y,
z,
xor2c ;
coherence
2GatesCircOutput x,y,z,xor2c is Element of InnerVertices (GFA2AdderStr x,y,z)
;
end;
:: deftheorem Def36 defines GFA2AdderOutput GFACIRC1:def 36 :
theorem Th100: :: GFACIRC1:100
theorem Th101: :: GFACIRC1:101
theorem Th102: :: GFACIRC1:102
theorem Th103: :: GFACIRC1:103
theorem Th104: :: GFACIRC1:104
for
x,
y,
z being
set holds
(
x in the
carrier of
(GFA2AdderStr x,y,z) &
y in the
carrier of
(GFA2AdderStr x,y,z) &
z in the
carrier of
(GFA2AdderStr x,y,z) &
[<*x,y*>,xor2c ] in the
carrier of
(GFA2AdderStr x,y,z) &
[<*[<*x,y*>,xor2c ],z*>,xor2c ] in the
carrier of
(GFA2AdderStr x,y,z) )
by FACIRC_1:60, FACIRC_1:61;
theorem Th105: :: GFACIRC1:105
for
x,
y,
z being
set holds
(
[<*x,y*>,xor2c ] in InnerVertices (GFA2AdderStr x,y,z) &
GFA2AdderOutput x,
y,
z in InnerVertices (GFA2AdderStr x,y,z) )
theorem Th106: :: GFACIRC1:106
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2c ] holds
(
x in InputVertices (GFA2AdderStr x,y,z) &
y in InputVertices (GFA2AdderStr x,y,z) &
z in InputVertices (GFA2AdderStr x,y,z) )
theorem Th107: :: GFACIRC1:107
theorem Th108: :: GFACIRC1:108
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2c ] holds
for
s being
State of
(GFA2AdderCirc x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(
(Following s) . [<*x,y*>,xor2c ] = a1 'xor' ('not' a2) &
(Following s) . x = a1 &
(Following s) . y = a2 &
(Following s) . z = a3 )
theorem Th109: :: GFACIRC1:109
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2c ] holds
for
s being
State of
(GFA2AdderCirc x,y,z) for
a1a2,
a1,
a2,
a3 being
Element of
BOOLEAN st
a1a2 = s . [<*x,y*>,xor2c ] &
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(Following s) . (GFA2AdderOutput x,y,z) = a1a2 'xor' ('not' a3)
theorem Th110: :: GFACIRC1:110
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2c ] holds
for
s being
State of
(GFA2AdderCirc x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(
(Following s,2) . (GFA2AdderOutput x,y,z) = (a1 'xor' ('not' a2)) 'xor' ('not' a3) &
(Following s,2) . [<*x,y*>,xor2c ] = a1 'xor' ('not' a2) &
(Following s,2) . x = a1 &
(Following s,2) . y = a2 &
(Following s,2) . z = a3 )
theorem Th111: :: GFACIRC1:111
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2c ] holds
for
s being
State of
(GFA2AdderCirc x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(Following s,2) . (GFA2AdderOutput x,y,z) = (('not' a1) 'xor' a2) 'xor' ('not' a3)
theorem Th112: :: GFACIRC1:112
definition
let x be
set ,
y be
set ,
z be
set ;
func BitGFA2Str c1,
c2,
c3 -> non
empty strict non
void unsplit gate`1=arity gate`2isBoolean ManySortedSign equals :: GFACIRC1:def 37
(GFA2AdderStr x,y,z) +* (GFA2CarryStr x,y,z);
coherence
(GFA2AdderStr x,y,z) +* (GFA2CarryStr x,y,z) is non empty strict non void unsplit gate`1=arity gate`2isBoolean ManySortedSign
;
end;
:: deftheorem Def37 defines BitGFA2Str GFACIRC1:def 37 :
definition
let x be
set ,
y be
set ,
z be
set ;
func BitGFA2Circ c1,
c2,
c3 -> strict gate`2=den Boolean Circuit of
BitGFA2Str a1,
a2,
a3 equals :: GFACIRC1:def 38
(GFA2AdderCirc x,y,z) +* (GFA2CarryCirc x,y,z);
coherence
(GFA2AdderCirc x,y,z) +* (GFA2CarryCirc x,y,z) is strict gate`2=den Boolean Circuit of BitGFA2Str x,y,z
;
end;
:: deftheorem Def38 defines BitGFA2Circ GFACIRC1:def 38 :
theorem Th113: :: GFACIRC1:113
for
x,
y,
z being
set holds
InnerVertices (BitGFA2Str x,y,z) = (({[<*x,y*>,xor2c ]} \/ {(GFA2AdderOutput x,y,z)}) \/ {[<*x,y*>,and2a ],[<*y,z*>,and2c ],[<*z,x*>,and2b ]}) \/ {(GFA2CarryOutput x,y,z)}
theorem Th114: :: GFACIRC1:114
theorem Th115: :: GFACIRC1:115
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2c ] &
x <> [<*y,z*>,and2c ] &
y <> [<*z,x*>,and2b ] &
z <> [<*x,y*>,and2a ] holds
InputVertices (BitGFA2Str x,y,z) = {x,y,z}
theorem Th116: :: GFACIRC1:116
theorem Th117: :: GFACIRC1:117
theorem Th118: :: GFACIRC1:118
for
x,
y,
z being
set holds
(
x in the
carrier of
(BitGFA2Str x,y,z) &
y in the
carrier of
(BitGFA2Str x,y,z) &
z in the
carrier of
(BitGFA2Str x,y,z) &
[<*x,y*>,xor2c ] in the
carrier of
(BitGFA2Str x,y,z) &
[<*[<*x,y*>,xor2c ],z*>,xor2c ] in the
carrier of
(BitGFA2Str x,y,z) &
[<*x,y*>,and2a ] in the
carrier of
(BitGFA2Str x,y,z) &
[<*y,z*>,and2c ] in the
carrier of
(BitGFA2Str x,y,z) &
[<*z,x*>,and2b ] in the
carrier of
(BitGFA2Str x,y,z) &
[<*[<*x,y*>,and2a ],[<*y,z*>,and2c ],[<*z,x*>,and2b ]*>,nor3 ] in the
carrier of
(BitGFA2Str x,y,z) )
theorem Th119: :: GFACIRC1:119
for
x,
y,
z being
set holds
(
[<*x,y*>,xor2c ] in InnerVertices (BitGFA2Str x,y,z) &
GFA2AdderOutput x,
y,
z in InnerVertices (BitGFA2Str x,y,z) &
[<*x,y*>,and2a ] in InnerVertices (BitGFA2Str x,y,z) &
[<*y,z*>,and2c ] in InnerVertices (BitGFA2Str x,y,z) &
[<*z,x*>,and2b ] in InnerVertices (BitGFA2Str x,y,z) &
GFA2CarryOutput x,
y,
z in InnerVertices (BitGFA2Str x,y,z) )
theorem Th120: :: GFACIRC1:120
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2c ] &
x <> [<*y,z*>,and2c ] &
y <> [<*z,x*>,and2b ] &
z <> [<*x,y*>,and2a ] holds
(
x in InputVertices (BitGFA2Str x,y,z) &
y in InputVertices (BitGFA2Str x,y,z) &
z in InputVertices (BitGFA2Str x,y,z) )
definition
let x be
set ,
y be
set ,
z be
set ;
func BitGFA2CarryOutput c1,
c2,
c3 -> Element of
InnerVertices (BitGFA2Str a1,a2,a3) equals :: GFACIRC1:def 39
[<*[<*x,y*>,and2a ],[<*y,z*>,and2c ],[<*z,x*>,and2b ]*>,nor3 ];
coherence
[<*[<*x,y*>,and2a ],[<*y,z*>,and2c ],[<*z,x*>,and2b ]*>,nor3 ] is Element of InnerVertices (BitGFA2Str x,y,z)
end;
:: deftheorem Def39 defines BitGFA2CarryOutput GFACIRC1:def 39 :
for
x,
y,
z being
set holds
BitGFA2CarryOutput x,
y,
z = [<*[<*x,y*>,and2a ],[<*y,z*>,and2c ],[<*z,x*>,and2b ]*>,nor3 ];
definition
let x be
set ,
y be
set ,
z be
set ;
func BitGFA2AdderOutput c1,
c2,
c3 -> Element of
InnerVertices (BitGFA2Str a1,a2,a3) equals :: GFACIRC1:def 40
2GatesCircOutput x,
y,
z,
xor2c ;
coherence
2GatesCircOutput x,y,z,xor2c is Element of InnerVertices (BitGFA2Str x,y,z)
end;
:: deftheorem Def40 defines BitGFA2AdderOutput GFACIRC1:def 40 :
theorem Th121: :: GFACIRC1:121
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2c ] &
x <> [<*y,z*>,and2c ] &
y <> [<*z,x*>,and2b ] &
z <> [<*x,y*>,and2a ] holds
for
s being
State of
(BitGFA2Circ x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(
(Following s,2) . (GFA2AdderOutput x,y,z) = (('not' a1) 'xor' a2) 'xor' ('not' a3) &
(Following s,2) . (GFA2CarryOutput x,y,z) = 'not' (((('not' a1) '&' a2) 'or' (a2 '&' ('not' a3))) 'or' (('not' a3) '&' ('not' a1))) )
theorem Th122: :: GFACIRC1:122
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2c ] &
x <> [<*y,z*>,and2c ] &
y <> [<*z,x*>,and2b ] &
z <> [<*x,y*>,and2a ] holds
for
s being
State of
(BitGFA2Circ x,y,z) holds
Following s,2 is
stable
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA3CarryIStr c1,
c2,
c3 -> non
empty strict non
void unsplit gate`1=arity gate`2isBoolean ManySortedSign equals :: GFACIRC1:def 41
((1GateCircStr <*x,y*>,and2b ) +* (1GateCircStr <*y,z*>,and2b )) +* (1GateCircStr <*z,x*>,and2b );
coherence
((1GateCircStr <*x,y*>,and2b ) +* (1GateCircStr <*y,z*>,and2b )) +* (1GateCircStr <*z,x*>,and2b ) is non empty strict non void unsplit gate`1=arity gate`2isBoolean ManySortedSign
;
end;
:: deftheorem Def41 defines GFA3CarryIStr GFACIRC1:def 41 :
for
x,
y,
z being
set holds
GFA3CarryIStr x,
y,
z = ((1GateCircStr <*x,y*>,and2b ) +* (1GateCircStr <*y,z*>,and2b )) +* (1GateCircStr <*z,x*>,and2b );
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA3CarryICirc c1,
c2,
c3 -> strict gate`2=den Boolean Circuit of
GFA3CarryIStr a1,
a2,
a3 equals :: GFACIRC1:def 42
((1GateCircuit x,y,and2b ) +* (1GateCircuit y,z,and2b )) +* (1GateCircuit z,x,and2b );
coherence
((1GateCircuit x,y,and2b ) +* (1GateCircuit y,z,and2b )) +* (1GateCircuit z,x,and2b ) is strict gate`2=den Boolean Circuit of GFA3CarryIStr x,y,z
;
end;
:: deftheorem Def42 defines GFA3CarryICirc GFACIRC1:def 42 :
for
x,
y,
z being
set holds
GFA3CarryICirc x,
y,
z = ((1GateCircuit x,y,and2b ) +* (1GateCircuit y,z,and2b )) +* (1GateCircuit z,x,and2b );
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA3CarryStr c1,
c2,
c3 -> non
empty strict non
void unsplit gate`1=arity gate`2isBoolean ManySortedSign equals :: GFACIRC1:def 43
(GFA3CarryIStr x,y,z) +* (1GateCircStr <*[<*x,y*>,and2b ],[<*y,z*>,and2b ],[<*z,x*>,and2b ]*>,nor3 );
coherence
(GFA3CarryIStr x,y,z) +* (1GateCircStr <*[<*x,y*>,and2b ],[<*y,z*>,and2b ],[<*z,x*>,and2b ]*>,nor3 ) is non empty strict non void unsplit gate`1=arity gate`2isBoolean ManySortedSign
;
end;
:: deftheorem Def43 defines GFA3CarryStr GFACIRC1:def 43 :
for
x,
y,
z being
set holds
GFA3CarryStr x,
y,
z = (GFA3CarryIStr x,y,z) +* (1GateCircStr <*[<*x,y*>,and2b ],[<*y,z*>,and2b ],[<*z,x*>,and2b ]*>,nor3 );
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA3CarryCirc c1,
c2,
c3 -> strict gate`2=den Boolean Circuit of
GFA3CarryStr a1,
a2,
a3 equals :: GFACIRC1:def 44
(GFA3CarryICirc x,y,z) +* (1GateCircuit [<*x,y*>,and2b ],[<*y,z*>,and2b ],[<*z,x*>,and2b ],nor3 );
coherence
(GFA3CarryICirc x,y,z) +* (1GateCircuit [<*x,y*>,and2b ],[<*y,z*>,and2b ],[<*z,x*>,and2b ],nor3 ) is strict gate`2=den Boolean Circuit of GFA3CarryStr x,y,z
;
end;
:: deftheorem Def44 defines GFA3CarryCirc GFACIRC1:def 44 :
for
x,
y,
z being
set holds
GFA3CarryCirc x,
y,
z = (GFA3CarryICirc x,y,z) +* (1GateCircuit [<*x,y*>,and2b ],[<*y,z*>,and2b ],[<*z,x*>,and2b ],nor3 );
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA3CarryOutput c1,
c2,
c3 -> Element of
InnerVertices (GFA3CarryStr a1,a2,a3) equals :: GFACIRC1:def 45
[<*[<*x,y*>,and2b ],[<*y,z*>,and2b ],[<*z,x*>,and2b ]*>,nor3 ];
coherence
[<*[<*x,y*>,and2b ],[<*y,z*>,and2b ],[<*z,x*>,and2b ]*>,nor3 ] is Element of InnerVertices (GFA3CarryStr x,y,z)
end;
:: deftheorem Def45 defines GFA3CarryOutput GFACIRC1:def 45 :
for
x,
y,
z being
set holds
GFA3CarryOutput x,
y,
z = [<*[<*x,y*>,and2b ],[<*y,z*>,and2b ],[<*z,x*>,and2b ]*>,nor3 ];
theorem Th123: :: GFACIRC1:123
for
x,
y,
z being
set holds
InnerVertices (GFA3CarryIStr x,y,z) = {[<*x,y*>,and2b ],[<*y,z*>,and2b ],[<*z,x*>,and2b ]}
theorem Th124: :: GFACIRC1:124
for
x,
y,
z being
set holds
InnerVertices (GFA3CarryStr x,y,z) = {[<*x,y*>,and2b ],[<*y,z*>,and2b ],[<*z,x*>,and2b ]} \/ {(GFA3CarryOutput x,y,z)}
theorem Th125: :: GFACIRC1:125
theorem Th126: :: GFACIRC1:126
for
x,
y,
z being
set st
x <> [<*y,z*>,and2b ] &
y <> [<*z,x*>,and2b ] &
z <> [<*x,y*>,and2b ] holds
InputVertices (GFA3CarryIStr x,y,z) = {x,y,z}
theorem Th127: :: GFACIRC1:127
for
x,
y,
z being
set st
x <> [<*y,z*>,and2b ] &
y <> [<*z,x*>,and2b ] &
z <> [<*x,y*>,and2b ] holds
InputVertices (GFA3CarryStr x,y,z) = {x,y,z}
theorem Th128: :: GFACIRC1:128
theorem Th129: :: GFACIRC1:129
for
x,
y,
z being
set holds
(
x in the
carrier of
(GFA3CarryStr x,y,z) &
y in the
carrier of
(GFA3CarryStr x,y,z) &
z in the
carrier of
(GFA3CarryStr x,y,z) &
[<*x,y*>,and2b ] in the
carrier of
(GFA3CarryStr x,y,z) &
[<*y,z*>,and2b ] in the
carrier of
(GFA3CarryStr x,y,z) &
[<*z,x*>,and2b ] in the
carrier of
(GFA3CarryStr x,y,z) &
[<*[<*x,y*>,and2b ],[<*y,z*>,and2b ],[<*z,x*>,and2b ]*>,nor3 ] in the
carrier of
(GFA3CarryStr x,y,z) )
theorem Th130: :: GFACIRC1:130
for
x,
y,
z being
set holds
(
[<*x,y*>,and2b ] in InnerVertices (GFA3CarryStr x,y,z) &
[<*y,z*>,and2b ] in InnerVertices (GFA3CarryStr x,y,z) &
[<*z,x*>,and2b ] in InnerVertices (GFA3CarryStr x,y,z) &
GFA3CarryOutput x,
y,
z in InnerVertices (GFA3CarryStr x,y,z) )
theorem Th131: :: GFACIRC1:131
for
x,
y,
z being
set st
x <> [<*y,z*>,and2b ] &
y <> [<*z,x*>,and2b ] &
z <> [<*x,y*>,and2b ] holds
(
x in InputVertices (GFA3CarryStr x,y,z) &
y in InputVertices (GFA3CarryStr x,y,z) &
z in InputVertices (GFA3CarryStr x,y,z) )
theorem Th132: :: GFACIRC1:132
theorem Th133: :: GFACIRC1:133
for
x,
y,
z being
set for
s being
State of
(GFA3CarryCirc x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(
(Following s) . [<*x,y*>,and2b ] = ('not' a1) '&' ('not' a2) &
(Following s) . [<*y,z*>,and2b ] = ('not' a2) '&' ('not' a3) &
(Following s) . [<*z,x*>,and2b ] = ('not' a3) '&' ('not' a1) )
theorem Th134: :: GFACIRC1:134
for
x,
y,
z being
set for
s being
State of
(GFA3CarryCirc x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . [<*x,y*>,and2b ] &
a2 = s . [<*y,z*>,and2b ] &
a3 = s . [<*z,x*>,and2b ] holds
(Following s) . (GFA3CarryOutput x,y,z) = 'not' ((a1 'or' a2) 'or' a3)
theorem Th135: :: GFACIRC1:135
for
x,
y,
z being
set st
x <> [<*y,z*>,and2b ] &
y <> [<*z,x*>,and2b ] &
z <> [<*x,y*>,and2b ] holds
for
s being
State of
(GFA3CarryCirc x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(
(Following s,2) . (GFA3CarryOutput x,y,z) = 'not' (((('not' a1) '&' ('not' a2)) 'or' (('not' a2) '&' ('not' a3))) 'or' (('not' a3) '&' ('not' a1))) &
(Following s,2) . [<*x,y*>,and2b ] = ('not' a1) '&' ('not' a2) &
(Following s,2) . [<*y,z*>,and2b ] = ('not' a2) '&' ('not' a3) &
(Following s,2) . [<*z,x*>,and2b ] = ('not' a3) '&' ('not' a1) )
theorem Th136: :: GFACIRC1:136
for
x,
y,
z being
set st
x <> [<*y,z*>,and2b ] &
y <> [<*z,x*>,and2b ] &
z <> [<*x,y*>,and2b ] holds
for
s being
State of
(GFA3CarryCirc x,y,z) holds
Following s,2 is
stable
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA3AdderStr c1,
c2,
c3 -> non
empty strict non
void unsplit gate`1=arity gate`2isBoolean ManySortedSign equals :: GFACIRC1:def 46
2GatesCircStr x,
y,
z,
xor2 ;
coherence
2GatesCircStr x,y,z,xor2 is non empty strict non void unsplit gate`1=arity gate`2isBoolean ManySortedSign
;
end;
:: deftheorem Def46 defines GFA3AdderStr GFACIRC1:def 46 :
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA3AdderCirc c1,
c2,
c3 -> strict gate`2=den Boolean Circuit of
GFA3AdderStr a1,
a2,
a3 equals :: GFACIRC1:def 47
2GatesCircuit x,
y,
z,
xor2 ;
coherence
2GatesCircuit x,y,z,xor2 is strict gate`2=den Boolean Circuit of GFA3AdderStr x,y,z
;
end;
:: deftheorem Def47 defines GFA3AdderCirc GFACIRC1:def 47 :
definition
let x be
set ,
y be
set ,
z be
set ;
func GFA3AdderOutput c1,
c2,
c3 -> Element of
InnerVertices (GFA3AdderStr a1,a2,a3) equals :: GFACIRC1:def 48
2GatesCircOutput x,
y,
z,
xor2 ;
coherence
2GatesCircOutput x,y,z,xor2 is Element of InnerVertices (GFA3AdderStr x,y,z)
;
end;
:: deftheorem Def48 defines GFA3AdderOutput GFACIRC1:def 48 :
theorem Th137: :: GFACIRC1:137
theorem Th138: :: GFACIRC1:138
theorem Th139: :: GFACIRC1:139
theorem Th140: :: GFACIRC1:140
theorem Th141: :: GFACIRC1:141
for
x,
y,
z being
set holds
(
x in the
carrier of
(GFA3AdderStr x,y,z) &
y in the
carrier of
(GFA3AdderStr x,y,z) &
z in the
carrier of
(GFA3AdderStr x,y,z) &
[<*x,y*>,xor2 ] in the
carrier of
(GFA3AdderStr x,y,z) &
[<*[<*x,y*>,xor2 ],z*>,xor2 ] in the
carrier of
(GFA3AdderStr x,y,z) )
by FACIRC_1:60, FACIRC_1:61;
theorem Th142: :: GFACIRC1:142
for
x,
y,
z being
set holds
(
[<*x,y*>,xor2 ] in InnerVertices (GFA3AdderStr x,y,z) &
GFA3AdderOutput x,
y,
z in InnerVertices (GFA3AdderStr x,y,z) )
theorem Th143: :: GFACIRC1:143
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2 ] holds
(
x in InputVertices (GFA3AdderStr x,y,z) &
y in InputVertices (GFA3AdderStr x,y,z) &
z in InputVertices (GFA3AdderStr x,y,z) )
theorem Th144: :: GFACIRC1:144
theorem Th145: :: GFACIRC1:145
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2 ] holds
for
s being
State of
(GFA3AdderCirc x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(
(Following s) . [<*x,y*>,xor2 ] = a1 'xor' a2 &
(Following s) . x = a1 &
(Following s) . y = a2 &
(Following s) . z = a3 )
theorem Th146: :: GFACIRC1:146
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2 ] holds
for
s being
State of
(GFA3AdderCirc x,y,z) for
a1a2,
a1,
a2,
a3 being
Element of
BOOLEAN st
a1a2 = s . [<*x,y*>,xor2 ] &
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(Following s) . (GFA3AdderOutput x,y,z) = a1a2 'xor' a3
theorem Th147: :: GFACIRC1:147
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2 ] holds
for
s being
State of
(GFA3AdderCirc x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(
(Following s,2) . (GFA3AdderOutput x,y,z) = (a1 'xor' a2) 'xor' a3 &
(Following s,2) . [<*x,y*>,xor2 ] = a1 'xor' a2 &
(Following s,2) . x = a1 &
(Following s,2) . y = a2 &
(Following s,2) . z = a3 )
theorem Th148: :: GFACIRC1:148
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2 ] holds
for
s being
State of
(GFA3AdderCirc x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(Following s,2) . (GFA3AdderOutput x,y,z) = 'not' ((('not' a1) 'xor' ('not' a2)) 'xor' ('not' a3))
theorem Th149: :: GFACIRC1:149
definition
let x be
set ,
y be
set ,
z be
set ;
func BitGFA3Str c1,
c2,
c3 -> non
empty strict non
void unsplit gate`1=arity gate`2isBoolean ManySortedSign equals :: GFACIRC1:def 49
(GFA3AdderStr x,y,z) +* (GFA3CarryStr x,y,z);
coherence
(GFA3AdderStr x,y,z) +* (GFA3CarryStr x,y,z) is non empty strict non void unsplit gate`1=arity gate`2isBoolean ManySortedSign
;
end;
:: deftheorem Def49 defines BitGFA3Str GFACIRC1:def 49 :
definition
let x be
set ,
y be
set ,
z be
set ;
func BitGFA3Circ c1,
c2,
c3 -> strict gate`2=den Boolean Circuit of
BitGFA3Str a1,
a2,
a3 equals :: GFACIRC1:def 50
(GFA3AdderCirc x,y,z) +* (GFA3CarryCirc x,y,z);
coherence
(GFA3AdderCirc x,y,z) +* (GFA3CarryCirc x,y,z) is strict gate`2=den Boolean Circuit of BitGFA3Str x,y,z
;
end;
:: deftheorem Def50 defines BitGFA3Circ GFACIRC1:def 50 :
theorem Th150: :: GFACIRC1:150
for
x,
y,
z being
set holds
InnerVertices (BitGFA3Str x,y,z) = (({[<*x,y*>,xor2 ]} \/ {(GFA3AdderOutput x,y,z)}) \/ {[<*x,y*>,and2b ],[<*y,z*>,and2b ],[<*z,x*>,and2b ]}) \/ {(GFA3CarryOutput x,y,z)}
theorem Th151: :: GFACIRC1:151
theorem Th152: :: GFACIRC1:152
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2 ] &
x <> [<*y,z*>,and2b ] &
y <> [<*z,x*>,and2b ] &
z <> [<*x,y*>,and2b ] holds
InputVertices (BitGFA3Str x,y,z) = {x,y,z}
theorem Th153: :: GFACIRC1:153
theorem Th154: :: GFACIRC1:154
theorem Th155: :: GFACIRC1:155
for
x,
y,
z being
set holds
(
x in the
carrier of
(BitGFA3Str x,y,z) &
y in the
carrier of
(BitGFA3Str x,y,z) &
z in the
carrier of
(BitGFA3Str x,y,z) &
[<*x,y*>,xor2 ] in the
carrier of
(BitGFA3Str x,y,z) &
[<*[<*x,y*>,xor2 ],z*>,xor2 ] in the
carrier of
(BitGFA3Str x,y,z) &
[<*x,y*>,and2b ] in the
carrier of
(BitGFA3Str x,y,z) &
[<*y,z*>,and2b ] in the
carrier of
(BitGFA3Str x,y,z) &
[<*z,x*>,and2b ] in the
carrier of
(BitGFA3Str x,y,z) &
[<*[<*x,y*>,and2b ],[<*y,z*>,and2b ],[<*z,x*>,and2b ]*>,nor3 ] in the
carrier of
(BitGFA3Str x,y,z) )
theorem Th156: :: GFACIRC1:156
for
x,
y,
z being
set holds
(
[<*x,y*>,xor2 ] in InnerVertices (BitGFA3Str x,y,z) &
GFA3AdderOutput x,
y,
z in InnerVertices (BitGFA3Str x,y,z) &
[<*x,y*>,and2b ] in InnerVertices (BitGFA3Str x,y,z) &
[<*y,z*>,and2b ] in InnerVertices (BitGFA3Str x,y,z) &
[<*z,x*>,and2b ] in InnerVertices (BitGFA3Str x,y,z) &
GFA3CarryOutput x,
y,
z in InnerVertices (BitGFA3Str x,y,z) )
theorem Th157: :: GFACIRC1:157
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2 ] &
x <> [<*y,z*>,and2b ] &
y <> [<*z,x*>,and2b ] &
z <> [<*x,y*>,and2b ] holds
(
x in InputVertices (BitGFA3Str x,y,z) &
y in InputVertices (BitGFA3Str x,y,z) &
z in InputVertices (BitGFA3Str x,y,z) )
definition
let x be
set ,
y be
set ,
z be
set ;
func BitGFA3CarryOutput c1,
c2,
c3 -> Element of
InnerVertices (BitGFA3Str a1,a2,a3) equals :: GFACIRC1:def 51
[<*[<*x,y*>,and2b ],[<*y,z*>,and2b ],[<*z,x*>,and2b ]*>,nor3 ];
coherence
[<*[<*x,y*>,and2b ],[<*y,z*>,and2b ],[<*z,x*>,and2b ]*>,nor3 ] is Element of InnerVertices (BitGFA3Str x,y,z)
end;
:: deftheorem Def51 defines BitGFA3CarryOutput GFACIRC1:def 51 :
for
x,
y,
z being
set holds
BitGFA3CarryOutput x,
y,
z = [<*[<*x,y*>,and2b ],[<*y,z*>,and2b ],[<*z,x*>,and2b ]*>,nor3 ];
definition
let x be
set ,
y be
set ,
z be
set ;
func BitGFA3AdderOutput c1,
c2,
c3 -> Element of
InnerVertices (BitGFA3Str a1,a2,a3) equals :: GFACIRC1:def 52
2GatesCircOutput x,
y,
z,
xor2 ;
coherence
2GatesCircOutput x,y,z,xor2 is Element of InnerVertices (BitGFA3Str x,y,z)
end;
:: deftheorem Def52 defines BitGFA3AdderOutput GFACIRC1:def 52 :
theorem Th158: :: GFACIRC1:158
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2 ] &
x <> [<*y,z*>,and2b ] &
y <> [<*z,x*>,and2b ] &
z <> [<*x,y*>,and2b ] holds
for
s being
State of
(BitGFA3Circ x,y,z) for
a1,
a2,
a3 being
Element of
BOOLEAN st
a1 = s . x &
a2 = s . y &
a3 = s . z holds
(
(Following s,2) . (GFA3AdderOutput x,y,z) = 'not' ((('not' a1) 'xor' ('not' a2)) 'xor' ('not' a3)) &
(Following s,2) . (GFA3CarryOutput x,y,z) = 'not' (((('not' a1) '&' ('not' a2)) 'or' (('not' a2) '&' ('not' a3))) 'or' (('not' a3) '&' ('not' a1))) )
theorem Th159: :: GFACIRC1:159
for
x,
y,
z being
set st
z <> [<*x,y*>,xor2 ] &
x <> [<*y,z*>,and2b ] &
y <> [<*z,x*>,and2b ] &
z <> [<*x,y*>,and2b ] holds
for
s being
State of
(BitGFA3Circ x,y,z) holds
Following s,2 is
stable