begin
deffunc H1( LattStr ) -> set = the carrier of $1;
deffunc H2( LattStr ) -> Element of bool [:[: the carrier of $1, the carrier of $1:], the carrier of $1:] = the L_join of $1;
deffunc H3( LattStr ) -> Element of bool [:[: the carrier of $1, the carrier of $1:], the carrier of $1:] = the L_meet of $1;
definition
let X be
set ;
func BooleLatt X -> strict LattStr means :
Def1:
( the
carrier of
it = bool X & ( for
Y,
Z being
Subset of
X holds
( the
L_join of
it . (
Y,
Z)
= Y \/ Z & the
L_meet of
it . (
Y,
Z)
= Y /\ Z ) ) );
existence
ex b1 being strict LattStr st
( the carrier of b1 = bool X & ( for Y, Z being Subset of X holds
( the L_join of b1 . (Y,Z) = Y \/ Z & the L_meet of b1 . (Y,Z) = Y /\ Z ) ) )
uniqueness
for b1, b2 being strict LattStr st the carrier of b1 = bool X & ( for Y, Z being Subset of X holds
( the L_join of b1 . (Y,Z) = Y \/ Z & the L_meet of b1 . (Y,Z) = Y /\ Z ) ) & the carrier of b2 = bool X & ( for Y, Z being Subset of X holds
( the L_join of b2 . (Y,Z) = Y \/ Z & the L_meet of b2 . (Y,Z) = Y /\ Z ) ) holds
b1 = b2
end;
:: deftheorem Def1 defines BooleLatt LATTICE3:def 1 :
for X being set
for b2 being strict LattStr holds
( b2 = BooleLatt X iff ( the carrier of b2 = bool X & ( for Y, Z being Subset of X holds
( the L_join of b2 . (Y,Z) = Y \/ Z & the L_meet of b2 . (Y,Z) = Y /\ Z ) ) ) );
theorem Th1:
theorem Th2:
theorem Th3:
theorem Th4:
theorem
begin
:: deftheorem defines LattPOSet LATTICE3:def 2 :
for L being Lattice holds LattPOSet L = RelStr(# the carrier of L,(LattRel L) #);
theorem Th6:
:: deftheorem defines % LATTICE3:def 3 :
for L being Lattice
for p being Element of L holds p % = p;
:: deftheorem defines % LATTICE3:def 4 :
for L being Lattice
for p being Element of (LattPOSet L) holds % p = p;
theorem Th7:
:: deftheorem defines ~ LATTICE3:def 5 :
for A being RelStr holds A ~ = RelStr(# the carrier of A,( the InternalRel of A ~) #);
theorem
:: deftheorem defines ~ LATTICE3:def 6 :
for A being RelStr
for a being Element of A holds a ~ = a;
:: deftheorem defines ~ LATTICE3:def 7 :
for A being RelStr
for a being Element of (A ~) holds ~ a = a;
theorem Th9:
:: deftheorem defines is_<=_than LATTICE3:def 8 :
for A being RelStr
for X being set
for a being Element of A holds
( a is_<=_than X iff for b being Element of A st b in X holds
a <= b );
:: deftheorem Def9 defines is_<=_than LATTICE3:def 9 :
for A being RelStr
for X being set
for a being Element of A holds
( X is_<=_than a iff for b being Element of A st b in X holds
b <= a );
:: deftheorem Def10 defines with_suprema LATTICE3:def 10 :
for IT being RelStr holds
( IT is with_suprema iff for x, y being Element of IT ex z being Element of IT st
( x <= z & y <= z & ( for z9 being Element of IT st x <= z9 & y <= z9 holds
z <= z9 ) ) );
:: deftheorem Def11 defines with_infima LATTICE3:def 11 :
for IT being RelStr holds
( IT is with_infima iff for x, y being Element of IT ex z being Element of IT st
( z <= x & z <= y & ( for z9 being Element of IT st z9 <= x & z9 <= y holds
z9 <= z ) ) );
theorem
theorem
:: deftheorem Def12 defines complete LATTICE3:def 12 :
for IT being RelStr holds
( IT is complete iff for X being set ex a being Element of IT st
( X is_<=_than a & ( for b being Element of IT st X is_<=_than b holds
a <= b ) ) );
theorem Th12:
:: deftheorem Def13 defines "\/" LATTICE3:def 13 :
for A being RelStr st A is antisymmetric holds
for a, b being Element of A st ex x being Element of A st
( a <= x & b <= x & ( for c being Element of A st a <= c & b <= c holds
x <= c ) ) holds
for b4 being Element of A holds
( b4 = a "\/" b iff ( a <= b4 & b <= b4 & ( for c being Element of A st a <= c & b <= c holds
b4 <= c ) ) );
:: deftheorem Def14 defines "/\" LATTICE3:def 14 :
for A being RelStr st A is antisymmetric holds
for a, b being Element of A st ex x being Element of A st
( a >= x & b >= x & ( for c being Element of A st a >= c & b >= c holds
x >= c ) ) holds
for b4 being Element of A holds
( b4 = a "/\" b iff ( b4 <= a & b4 <= b & ( for c being Element of A st c <= a & c <= b holds
c <= b4 ) ) );
theorem Th13:
theorem Th14:
theorem Th15:
theorem Th16:
theorem Th17:
theorem Th18:
theorem Th19:
:: deftheorem Def15 defines latt LATTICE3:def 15 :
for A being RelStr st A is with_suprema with_infima Poset holds
for b2 being strict Lattice holds
( b2 = latt A iff RelStr(# the carrier of A, the InternalRel of A #) = LattPOSet b2 );
theorem
begin
:: deftheorem Def16 defines is_less_than LATTICE3:def 16 :
for L being non empty LattStr
for p being Element of L
for X being set holds
( p is_less_than X iff for q being Element of L st q in X holds
p [= q );
:: deftheorem Def17 defines is_less_than LATTICE3:def 17 :
for L being non empty LattStr
for p being Element of L
for X being set holds
( X is_less_than p iff for q being Element of L st q in X holds
q [= p );
theorem
theorem
:: deftheorem Def18 defines complete LATTICE3:def 18 :
for IT being non empty LattStr holds
( IT is complete iff for X being set ex p being Element of IT st
( X is_less_than p & ( for r being Element of IT st X is_less_than r holds
p [= r ) ) );
:: deftheorem Def19 defines \/-distributive LATTICE3:def 19 :
for IT being non empty LattStr holds
( IT is \/-distributive iff for X being set
for a, b, c being Element of IT st X is_less_than a & ( for d being Element of IT st X is_less_than d holds
a [= d ) & { (b "/\" a9) where a9 is Element of IT : a9 in X } is_less_than c & ( for d being Element of IT st { (b "/\" a9) where a9 is Element of IT : a9 in X } is_less_than d holds
c [= d ) holds
b "/\" a [= c );
:: deftheorem defines /\-distributive LATTICE3:def 20 :
for IT being non empty LattStr holds
( IT is /\-distributive iff for X being set
for a, b, c being Element of IT st X is_greater_than a & ( for d being Element of IT st X is_greater_than d holds
d [= a ) & { (b "\/" a9) where a9 is Element of IT : a9 in X } is_greater_than c & ( for d being Element of IT st { (b "\/" a9) where a9 is Element of IT : a9 in X } is_greater_than d holds
d [= c ) holds
c [= b "\/" a );
theorem
theorem Th24:
theorem Th25:
theorem Th26:
theorem Th27:
theorem Th28:
theorem
theorem Th30:
theorem Th31:
:: deftheorem Def21 defines "\/" LATTICE3:def 21 :
for L being non empty LattStr st L is complete Lattice holds
for X being set
for b3 being Element of L holds
( b3 = "\/" (X,L) iff ( X is_less_than b3 & ( for r being Element of L st X is_less_than r holds
b3 [= r ) ) );
:: deftheorem defines "/\" LATTICE3:def 22 :
for L being non empty LattStr
for X being set holds "/\" (X,L) = "\/" ( { p where p is Element of L : p is_less_than X } ,L);
theorem Th32:
theorem Th33:
theorem Th34:
theorem Th35:
theorem Th36:
theorem
theorem Th38:
theorem
canceled;
theorem Th40:
theorem Th41:
theorem Th42:
theorem
theorem
theorem
theorem Th46:
theorem Th47:
theorem
theorem
theorem
theorem
theorem Th52:
theorem
theorem
theorem
Lm3:
now
let D be non
empty set ;
for f being Function of (bool D),D st ( for a being Element of D holds f . {a} = a ) & ( for X being Subset-Family of D holds f . (f .: X) = f . (union X) ) holds
ex L being strict Lattice st
( H1(L) = D & ( for X being Subset of L holds "\/" X = f . X ) )let f be
Function of
(bool D),
D;
( ( for a being Element of D holds f . {a} = a ) & ( for X being Subset-Family of D holds f . (f .: X) = f . (union X) ) implies ex L being strict Lattice st
( H1(L) = D & ( for X being Subset of L holds "\/" X = f . X ) ) )assume that A1:
for
a being
Element of
D holds
f . {a} = a
and A2:
for
X being
Subset-Family of
D holds
f . (f .: X) = f . (union X)
;
ex L being strict Lattice st
( H1(L) = D & ( for X being Subset of L holds "\/" X = f . X ) )defpred S1[
set ,
set ]
means f . {$1,$2} = $2;
consider R being
Relation of
D such that A3:
for
x,
y being
set holds
(
[x,y] in R iff (
x in D &
y in D &
S1[
x,
y] ) )
from RELSET_1:sch 1();
A4:
dom f = bool D
by FUNCT_2:def 1;
A6:
for
x,
y being
Element of
D for
X being
Subset of
D st
y in X holds
f . (X \/ {x}) = f . { (f . {t,x}) where t is Element of D : t in X }
proof
let x,
y be
Element of
D;
for X being Subset of D st y in X holds
f . (X \/ {x}) = f . { (f . {t,x}) where t is Element of D : t in X } let X be
Subset of
D;
( y in X implies f . (X \/ {x}) = f . { (f . {t,x}) where t is Element of D : t in X } )
assume A7:
y in X
;
f . (X \/ {x}) = f . { (f . {t,x}) where t is Element of D : t in X }
set Y =
{ {t,x} where t is Element of D : t in X } ;
A8:
X \/ {x} = union { {t,x} where t is Element of D : t in X }
proof
thus
X \/ {x} c= union { {t,x} where t is Element of D : t in X }
XBOOLE_0:def 10 union { {t,x} where t is Element of D : t in X } c= X \/ {x}
proof
let s be
set ;
TARSKI:def 3 ( not s in X \/ {x} or s in union { {t,x} where t is Element of D : t in X } )
assume
s in X \/ {x}
;
s in union { {t,x} where t is Element of D : t in X }
then
( (
s in X &
X c= D ) or
s in {x} )
by XBOOLE_0:def 3;
then
( (
s in X &
s is
Element of
D ) or
s = x )
by TARSKI:def 1;
then
( (
s in {s,x} &
{s,x} in { {t,x} where t is Element of D : t in X } ) or (
s in {y,x} &
{y,x} in { {t,x} where t is Element of D : t in X } ) )
by A7, TARSKI:def 2;
hence
s in union { {t,x} where t is Element of D : t in X }
by TARSKI:def 4;
verum
end;
let s be
set ;
TARSKI:def 3 ( not s in union { {t,x} where t is Element of D : t in X } or s in X \/ {x} )
assume
s in union { {t,x} where t is Element of D : t in X }
;
s in X \/ {x}
then consider Z being
set such that A9:
s in Z
and A10:
Z in { {t,x} where t is Element of D : t in X }
by TARSKI:def 4;
consider t being
Element of
D such that A11:
Z = {t,x}
and A12:
t in X
by A10;
(
s = t or (
s = x &
x in {x} ) )
by A9, A11, TARSKI:def 1, TARSKI:def 2;
hence
s in X \/ {x}
by A12, XBOOLE_0:def 3;
verum
end;
{ {t,x} where t is Element of D : t in X } c= bool D
then reconsider Y =
{ {t,x} where t is Element of D : t in X } as
Subset-Family of
D ;
defpred S2[
set ]
means $1
in X;
deffunc H4(
Element of
D)
-> Element of
bool D =
{$1,x};
A13:
bool D c= dom f
by FUNCT_2:def 1;
f .: { H4(t) where t is Element of D : S2[t] } = { (f . H4(t)) where t is Element of D : S2[t] }
from LATTICE3:sch 2(A13);
then
f . (union Y) = f . { (f . {t,x}) where t is Element of D : t in X }
by A2;
hence
f . (X \/ {x}) = f . { (f . {t,x}) where t is Element of D : t in X }
by A8;
verum
end;
A14:
R is_reflexive_in D
A16:
R is_antisymmetric_in D
A19:
R is_transitive_in D
proof
let x,
y,
z be
set ;
RELAT_2:def 8 ( not x in D or not y in D or not z in D or not [x,y] in R or not [y,z] in R or [x,z] in R )
assume that A20:
x in D
and A21:
y in D
and A22:
z in D
and A23:
[x,y] in R
and A24:
[y,z] in R
;
[x,z] in R
reconsider a =
x,
b =
y,
c =
z as
Element of
D by A20, A21, A22;
A25:
f . {x,y} = y
by A3, A23;
A26:
f . {y,z} = z
by A3, A24;
then f . {a,c} =
f . {(f . {a}),(f . {b,c})}
by A1
.=
f . ({a} \/ {b,c})
by A5
.=
f . {a,b,c}
by ENUMSET1:42
.=
f . ({a,b} \/ {c})
by ENUMSET1:43
.=
f . {(f . {a,b}),(f . {c})}
by A5
.=
c
by A1, A25, A26
;
hence
[x,z] in R
by A3;
verum
end;
A27:
dom R = D
by A14, ORDERS_1:98;
field R = D
by A14, ORDERS_1:98;
then reconsider R =
R as
Order of
D by A14, A16, A19, A27, PARTFUN1:def 4, RELAT_2:def 9, RELAT_2:def 12, RELAT_2:def 16;
set A =
RelStr(#
D,
R #);
RelStr(#
D,
R #) is
complete
proof
let X be
set ;
LATTICE3:def 12 ex a being Element of RelStr(# D,R #) st
( X is_<=_than a & ( for b being Element of RelStr(# D,R #) st X is_<=_than b holds
a <= b ) )
reconsider Y =
X /\ D as
Subset of
D by XBOOLE_1:17;
reconsider a =
f . Y as
Element of
RelStr(#
D,
R #) ;
take
a
;
( X is_<=_than a & ( for b being Element of RelStr(# D,R #) st X is_<=_than b holds
a <= b ) )
thus
X is_<=_than a
for b being Element of RelStr(# D,R #) st X is_<=_than b holds
a <= b
let b be
Element of
RelStr(#
D,
R #);
( X is_<=_than b implies a <= b )
assume A28:
X is_<=_than b
;
a <= b
A29:
f . {a,b} =
f . {a,(f . {b})}
by A1
.=
f . (Y \/ {b})
by A5
;
hence
[a,b] in the
InternalRel of
RelStr(#
D,
R #)
by A3;
ORDERS_2:def 9 verum
end;
then reconsider A =
RelStr(#
D,
R #) as non
empty strict complete Poset ;
take L =
latt A;
( H1(L) = D & ( for X being Subset of L holds "\/" X = f . X ) )A34:
(
A is
with_suprema &
A is
with_infima )
by Th12;
then A35:
A = LattPOSet L
by Def15;
hence
H1(
L)
= D
;
for X being Subset of L holds "\/" X = f . Xlet X be
Subset of
L;
"\/" X = f . Xreconsider Y =
X as
Subset of
D by A35;
reconsider a =
f . Y as
Element of
(LattPOSet L) by A34, Def15;
set p =
% a;
X is_<=_than a
then A37:
X is_less_than % a
by Th31;
hence
"\/" X = f . X
by A37, Def21;
verum
end;
theorem