definition
let A,
B be non
empty preBoolean set ;
let x,
y be
Element of
[:A,B:];
reflexivity
for x being Element of [:A,B:] holds
( x `1 c= x `1 & x `2 c= x `2 )
;
correctness
coherence
[((x `1) \/ (y `1)),((x `2) \/ (y `2))] is Element of [:A,B:];
;
commutativity
for b1, x, y being Element of [:A,B:] st b1 = [((x `1) \/ (y `1)),((x `2) \/ (y `2))] holds
b1 = [((y `1) \/ (x `1)),((y `2) \/ (x `2))]
;
idempotence
for x being Element of [:A,B:] holds x = [((x `1) \/ (x `1)),((x `2) \/ (x `2))]
by MCART_1:21;
correctness
coherence
[((x `1) /\ (y `1)),((x `2) /\ (y `2))] is Element of [:A,B:];
;
commutativity
for b1, x, y being Element of [:A,B:] st b1 = [((x `1) /\ (y `1)),((x `2) /\ (y `2))] holds
b1 = [((y `1) /\ (x `1)),((y `2) /\ (x `2))]
;
idempotence
for x being Element of [:A,B:] holds x = [((x `1) /\ (x `1)),((x `2) /\ (x `2))]
by MCART_1:21;
correctness
coherence
[((x `1) \ (y `1)),((x `2) \ (y `2))] is Element of [:A,B:];
;
correctness
coherence
[((x `1) \+\ (y `1)),((x `2) \+\ (y `2))] is Element of [:A,B:];
;
commutativity
for b1, x, y being Element of [:A,B:] st b1 = [((x `1) \+\ (y `1)),((x `2) \+\ (y `2))] holds
b1 = [((y `1) \+\ (x `1)),((y `2) \+\ (x `2))]
;
end;
definition
let A be
set ;
existence
ex b1 being BinOp of [:(Fin A),(Fin A):] st
for x, y being Element of [:(Fin A),(Fin A):] holds b1 . (x,y) = x \/ y
uniqueness
for b1, b2 being BinOp of [:(Fin A),(Fin A):] st ( for x, y being Element of [:(Fin A),(Fin A):] holds b1 . (x,y) = x \/ y ) & ( for x, y being Element of [:(Fin A),(Fin A):] holds b2 . (x,y) = x \/ y ) holds
b1 = b2
end;
Lm1:
for A being set holds FinPairUnion A is idempotent
Lm2:
for A being set holds FinPairUnion A is commutative
Lm3:
for A being set holds FinPairUnion A is associative
Lm4:
for A being set holds {} in { B where B is Element of Fin (DISJOINT_PAIRS A) : for a, b being Element of DISJOINT_PAIRS A st a in B & b in B & a c= b holds
a = b }
Lm5:
for A being set
for B, C being Element of Fin (DISJOINT_PAIRS A) st ( for a being Element of DISJOINT_PAIRS A st a in B holds
a in C ) holds
B c= C
Lm6:
for A being set
for a being Element of DISJOINT_PAIRS A
for B, C being Element of Fin (DISJOINT_PAIRS A) st a in B ^ C holds
ex b being Element of DISJOINT_PAIRS A st
( b c= a & b in (mi B) ^ C )
Lm7:
for A being set
for a being Element of DISJOINT_PAIRS A
for B, C being Element of Fin (DISJOINT_PAIRS A) st a in B ^ C holds
ex c being Element of DISJOINT_PAIRS A st
( c in C & c c= a )
Lm8:
for A being set
for K, L being Element of Normal_forms_on A holds mi ((K ^ L) \/ L) = mi L
definition
let A be
set ;
existence
ex b1 being strict LattStr st
( the carrier of b1 = Normal_forms_on A & ( for B, C being Element of Normal_forms_on A holds
( the L_join of b1 . (B,C) = mi (B \/ C) & the L_meet of b1 . (B,C) = mi (B ^ C) ) ) )
uniqueness
for b1, b2 being strict LattStr st the carrier of b1 = Normal_forms_on A & ( for B, C being Element of Normal_forms_on A holds
( the L_join of b1 . (B,C) = mi (B \/ C) & the L_meet of b1 . (B,C) = mi (B ^ C) ) ) & the carrier of b2 = Normal_forms_on A & ( for B, C being Element of Normal_forms_on A holds
( the L_join of b2 . (B,C) = mi (B \/ C) & the L_meet of b2 . (B,C) = mi (B ^ C) ) ) holds
b1 = b2
end;
Lm9:
for A being set
for a, b being Element of (NormForm A) holds a "\/" b = b "\/" a
Lm10:
for A being set
for a, b, c being Element of (NormForm A) holds a "\/" (b "\/" c) = (a "\/" b) "\/" c
Lm11:
for A being set
for K, L being Element of Normal_forms_on A holds the L_join of (NormForm A) . (( the L_meet of (NormForm A) . (K,L)),L) = L
Lm12:
for A being set
for a, b being Element of (NormForm A) holds (a "/\" b) "\/" b = b
Lm13:
for A being set
for a, b being Element of (NormForm A) holds a "/\" b = b "/\" a
Lm14:
for A being set
for a, b, c being Element of (NormForm A) holds a "/\" (b "/\" c) = (a "/\" b) "/\" c
Lm15:
for A being set
for K, L, M being Element of Normal_forms_on A holds the L_meet of (NormForm A) . (K,( the L_join of (NormForm A) . (L,M))) = the L_join of (NormForm A) . (( the L_meet of (NormForm A) . (K,L)),( the L_meet of (NormForm A) . (K,M)))
Lm16:
for A being set
for a, b being Element of (NormForm A) holds a "/\" (a "\/" b) = a