begin
:: deftheorem defines "\/" LATTICES:def 1 :
for G being non empty \/-SemiLattStr
for p, q being Element of G holds p "\/" q = the L_join of G . (p,q);
:: deftheorem defines "/\" LATTICES:def 2 :
for G being non empty /\-SemiLattStr
for p, q being Element of G holds p "/\" q = the L_meet of G . (p,q);
:: deftheorem Def3 defines [= LATTICES:def 3 :
for G being non empty \/-SemiLattStr
for p, q being Element of G holds
( p [= q iff p "\/" q = q );
Lm1:
for uu, nn being BinOp of (bool {})
for x, y being Element of LattStr(# (bool {}),uu,nn #) holds x = y
Lm2:
for n being BinOp of (bool {})
for x, y being Element of \/-SemiLattStr(# (bool {}),n #) holds x = y
Lm3:
for n being BinOp of (bool {})
for x, y being Element of /\-SemiLattStr(# (bool {}),n #) holds x = y
:: deftheorem Def4 defines join-commutative LATTICES:def 4 :
for IT being non empty \/-SemiLattStr holds
( IT is join-commutative iff for a, b being Element of IT holds a "\/" b = b "\/" a );
:: deftheorem Def5 defines join-associative LATTICES:def 5 :
for IT being non empty \/-SemiLattStr holds
( IT is join-associative iff for a, b, c being Element of IT holds a "\/" (b "\/" c) = (a "\/" b) "\/" c );
:: deftheorem Def6 defines meet-commutative LATTICES:def 6 :
for IT being non empty /\-SemiLattStr holds
( IT is meet-commutative iff for a, b being Element of IT holds a "/\" b = b "/\" a );
:: deftheorem Def7 defines meet-associative LATTICES:def 7 :
for IT being non empty /\-SemiLattStr holds
( IT is meet-associative iff for a, b, c being Element of IT holds a "/\" (b "/\" c) = (a "/\" b) "/\" c );
:: deftheorem Def8 defines meet-absorbing LATTICES:def 8 :
for IT being non empty LattStr holds
( IT is meet-absorbing iff for a, b being Element of IT holds (a "/\" b) "\/" b = b );
:: deftheorem Def9 defines join-absorbing LATTICES:def 9 :
for IT being non empty LattStr holds
( IT is join-absorbing iff for a, b being Element of IT holds a "/\" (a "\/" b) = a );
:: deftheorem Def10 defines Lattice-like LATTICES:def 10 :
for IT being non empty LattStr holds
( IT is Lattice-like iff ( IT is join-commutative & IT is join-associative & IT is meet-absorbing & IT is meet-commutative & IT is meet-associative & IT is join-absorbing ) );
:: deftheorem Def11 defines distributive LATTICES:def 11 :
for IT being non empty LattStr holds
( IT is distributive iff for a, b, c being Element of IT holds a "/\" (b "\/" c) = (a "/\" b) "\/" (a "/\" c) );
:: deftheorem Def12 defines modular LATTICES:def 12 :
for IT being non empty LattStr holds
( IT is modular iff for a, b, c being Element of IT st a [= c holds
a "\/" (b "/\" c) = (a "\/" b) "/\" c );
:: deftheorem Def13 defines lower-bounded LATTICES:def 13 :
for IT being non empty /\-SemiLattStr holds
( IT is lower-bounded iff ex c being Element of IT st
for a being Element of IT holds
( c "/\" a = c & a "/\" c = c ) );
:: deftheorem Def14 defines upper-bounded LATTICES:def 14 :
for IT being non empty \/-SemiLattStr holds
( IT is upper-bounded iff ex c being Element of IT st
for a being Element of IT holds
( c "\/" a = c & a "\/" c = c ) );
Lm4:
for n, u being BinOp of (bool {}) holds LattStr(# (bool {}),n,u #) is Lattice
Lm5:
for n, u being BinOp of (bool {}) holds LattStr(# (bool {}),n,u #) is 0_Lattice
Lm6:
for n, u being BinOp of (bool {}) holds LattStr(# (bool {}),n,u #) is 1_Lattice
:: deftheorem Def15 defines bounded LATTICES:def 15 :
for IT being non empty LattStr holds
( IT is bounded iff ( IT is lower-bounded & IT is upper-bounded ) );
:: deftheorem Def16 defines Bottom LATTICES:def 16 :
for L being non empty /\-SemiLattStr st L is lower-bounded holds
for b2 being Element of L holds
( b2 = Bottom L iff for a being Element of L holds
( b2 "/\" a = b2 & a "/\" b2 = b2 ) );
:: deftheorem Def17 defines Top LATTICES:def 17 :
for L being non empty \/-SemiLattStr st L is upper-bounded holds
for b2 being Element of L holds
( b2 = Top L iff for a being Element of L holds
( b2 "\/" a = b2 & a "\/" b2 = b2 ) );
:: deftheorem Def18 defines is_a_complement_of LATTICES:def 18 :
for L being non empty LattStr
for a, b being Element of L holds
( a is_a_complement_of b iff ( a "\/" b = Top L & b "\/" a = Top L & a "/\" b = Bottom L & b "/\" a = Bottom L ) );
:: deftheorem Def19 defines complemented LATTICES:def 19 :
for IT being non empty LattStr holds
( IT is complemented iff for b being Element of IT ex a being Element of IT st a is_a_complement_of b );
:: deftheorem Def20 defines Boolean LATTICES:def 20 :
for IT being non empty LattStr holds
( IT is Boolean iff ( IT is bounded & IT is complemented & IT is distributive ) );
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem Th17:
theorem
theorem Th19:
theorem
canceled;
theorem Th21:
theorem Th22:
theorem Th23:
theorem
canceled;
theorem
theorem Th26:
theorem Th27:
theorem
canceled;
theorem
theorem
canceled;
theorem Th31:
theorem Th32:
theorem
canceled;
theorem
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem Th39:
theorem
theorem Th41:
theorem
canceled;
theorem Th43:
theorem
theorem
:: deftheorem Def21 defines ` LATTICES:def 21 :
for L being non empty LattStr
for x being Element of L st L is complemented D_Lattice holds
for b3 being Element of L holds
( b3 = x ` iff b3 is_a_complement_of x );
theorem
canceled;
theorem Th47:
theorem Th48:
theorem Th49:
theorem Th50:
theorem
theorem Th52:
theorem
begin
:: deftheorem Defx defines initial LATTICES:def 22 :
for L being Lattice
for S being Subset of L holds
( S is initial iff for p, q being Element of L st p [= q & q in S holds
p in S );
:: deftheorem Defy defines final LATTICES:def 23 :
for L being Lattice
for S being Subset of L holds
( S is final iff for p, q being Element of L st p [= q & p in S holds
q in S );
:: deftheorem defines meet-closed LATTICES:def 24 :
for L being Lattice
for S being Subset of L holds
( S is meet-closed iff for p, q being Element of L st p in S & q in S holds
p "/\" q in S );
:: deftheorem defines join-closed LATTICES:def 25 :
for L being Lattice
for S being Subset of L holds
( S is join-closed iff for p, q being Element of L st p in S & q in S holds
p "\/" q in S );
theorem