begin
Lm1:
for n being Nat
for p, q being Tuple of n, BOOLEAN st len p = n & len q = n & ( for i being Nat st i in Seg n holds
p /. i = q /. i ) holds
p = q
definition
let n be
Nat;
func Bin1 n -> Tuple of
n,
BOOLEAN means :
Def1:
for
i being
Nat st
i in Seg n holds
it /. i = IFEQ (
i,1,
TRUE,
FALSE);
existence
ex b1 being Tuple of n, BOOLEAN st
for i being Nat st i in Seg n holds
b1 /. i = IFEQ (i,1,TRUE,FALSE)
uniqueness
for b1, b2 being Tuple of n, BOOLEAN st ( for i being Nat st i in Seg n holds
b1 /. i = IFEQ (i,1,TRUE,FALSE) ) & ( for i being Nat st i in Seg n holds
b2 /. i = IFEQ (i,1,TRUE,FALSE) ) holds
b1 = b2
end;
:: deftheorem Def1 defines Bin1 BINARI_2:def 1 :
for n being Nat
for b2 being Tuple of n, BOOLEAN holds
( b2 = Bin1 n iff for i being Nat st i in Seg n holds
b2 /. i = IFEQ (i,1,TRUE,FALSE) );
:: deftheorem defines Neg2 BINARI_2:def 2 :
for n being non empty Nat
for x being Tuple of n, BOOLEAN holds Neg2 x = ('not' x) + (Bin1 n);
:: deftheorem Def3 defines Intval BINARI_2:def 3 :
for n being Nat
for x being Tuple of n, BOOLEAN holds
( ( x /. n = FALSE implies Intval x = Absval x ) & ( not x /. n = FALSE implies Intval x = (Absval x) - (2 to_power n) ) );
:: deftheorem defines Int_add_ovfl BINARI_2:def 4 :
for n being non empty Nat
for z1, z2 being Tuple of n, BOOLEAN holds Int_add_ovfl (z1,z2) = (('not' (z1 /. n)) '&' ('not' (z2 /. n))) '&' ((carry (z1,z2)) /. n);
:: deftheorem defines Int_add_udfl BINARI_2:def 5 :
for n being non empty Nat
for z1, z2 being Tuple of n, BOOLEAN holds Int_add_udfl (z1,z2) = ((z1 /. n) '&' (z2 /. n)) '&' ('not' ((carry (z1,z2)) /. n));
theorem
canceled;
theorem
canceled;
theorem
theorem Th4:
theorem
theorem
theorem Th7:
theorem Th8:
theorem Th9:
theorem Th10:
theorem Th11:
theorem Th12:
theorem Th13:
for
m being non
empty Nat for
z1,
z2 being
Tuple of
m,
BOOLEAN for
d1,
d2 being
Element of
BOOLEAN holds
((Intval ((z1 ^ <*d1*>) + (z2 ^ <*d2*>))) + (IFEQ ((Int_add_ovfl ((z1 ^ <*d1*>),(z2 ^ <*d2*>))),FALSE,0,(2 to_power (m + 1))))) - (IFEQ ((Int_add_udfl ((z1 ^ <*d1*>),(z2 ^ <*d2*>))),FALSE,0,(2 to_power (m + 1)))) = (Intval (z1 ^ <*d1*>)) + (Intval (z2 ^ <*d2*>))
theorem Th14:
for
m being non
empty Nat for
z1,
z2 being
Tuple of
m,
BOOLEAN for
d1,
d2 being
Element of
BOOLEAN holds
Intval ((z1 ^ <*d1*>) + (z2 ^ <*d2*>)) = (((Intval (z1 ^ <*d1*>)) + (Intval (z2 ^ <*d2*>))) - (IFEQ ((Int_add_ovfl ((z1 ^ <*d1*>),(z2 ^ <*d2*>))),FALSE,0,(2 to_power (m + 1))))) + (IFEQ ((Int_add_udfl ((z1 ^ <*d1*>),(z2 ^ <*d2*>))),FALSE,0,(2 to_power (m + 1))))
theorem Th15:
theorem Th16:
theorem Th17:
theorem
definition
let n be non
empty Nat;
let x,
y be
Tuple of
n,
BOOLEAN ;
func x - y -> Tuple of
n,
BOOLEAN means :
Def6:
for
i being
Nat st
i in Seg n holds
it /. i = ((x /. i) 'xor' ((Neg2 y) /. i)) 'xor' ((carry (x,(Neg2 y))) /. i);
existence
ex b1 being Tuple of n, BOOLEAN st
for i being Nat st i in Seg n holds
b1 /. i = ((x /. i) 'xor' ((Neg2 y) /. i)) 'xor' ((carry (x,(Neg2 y))) /. i)
uniqueness
for b1, b2 being Tuple of n, BOOLEAN st ( for i being Nat st i in Seg n holds
b1 /. i = ((x /. i) 'xor' ((Neg2 y) /. i)) 'xor' ((carry (x,(Neg2 y))) /. i) ) & ( for i being Nat st i in Seg n holds
b2 /. i = ((x /. i) 'xor' ((Neg2 y) /. i)) 'xor' ((carry (x,(Neg2 y))) /. i) ) holds
b1 = b2
end;
:: deftheorem Def6 defines - BINARI_2:def 6 :
for n being non empty Nat
for x, y, b4 being Tuple of n, BOOLEAN holds
( b4 = x - y iff for i being Nat st i in Seg n holds
b4 /. i = ((x /. i) 'xor' ((Neg2 y) /. i)) 'xor' ((carry (x,(Neg2 y))) /. i) );
theorem Th19:
theorem
theorem
for
m being non
empty Nat for
z1,
z2 being
Tuple of
m,
BOOLEAN for
d1,
d2 being
Element of
BOOLEAN holds
(((Intval ((z1 ^ <*d1*>) - (z2 ^ <*d2*>))) + (IFEQ ((Int_add_ovfl ((z1 ^ <*d1*>),(Neg2 (z2 ^ <*d2*>)))),FALSE,0,(2 to_power (m + 1))))) - (IFEQ ((Int_add_udfl ((z1 ^ <*d1*>),(Neg2 (z2 ^ <*d2*>)))),FALSE,0,(2 to_power (m + 1))))) + (IFEQ ((Int_add_ovfl (('not' (z2 ^ <*d2*>)),(Bin1 (m + 1)))),FALSE,0,(2 to_power (m + 1)))) = (Intval (z1 ^ <*d1*>)) - (Intval (z2 ^ <*d2*>))