:: WAYBEL23 semantic presentation

theorem Th1: :: WAYBEL23:1
for b1 being non empty Poset
for b2 being Element of b1 holds compactbelow b2 = (waybelow b2) /\ the carrier of (CompactSublatt b1)
proof end;

definition
let c1 be non empty reflexive transitive RelStr ;
let c2 be Subset of (InclPoset (Ids c1));
redefine func union as union c2 -> Subset of a1;
coherence
union c2 is Subset of c1
proof end;
end;

Lemma2: for b1 being non empty set
for b2 being Subset of (InclPoset b1) holds
( ex_sup_of b2, InclPoset b1 implies union b2 c= sup b2 )
proof end;

theorem Th2: :: WAYBEL23:2
for b1 being non empty RelStr
for b2, b3 being Subset of b1 holds
( b2 c= b3 implies finsups b2 c= finsups b3 )
proof end;

theorem Th3: :: WAYBEL23:3
for b1 being non empty transitive RelStr
for b2 being non empty full sups-inheriting SubRelStr of b1
for b3 being Subset of b1
for b4 being Subset of b2 holds
( b3 = b4 implies finsups b3 c= finsups b4 )
proof end;

theorem Th4: :: WAYBEL23:4
for b1 being non empty transitive antisymmetric complete RelStr
for b2 being non empty full sups-inheriting SubRelStr of b1
for b3 being Subset of b1
for b4 being Subset of b2 holds
( b3 = b4 implies finsups b3 = finsups b4 )
proof end;

theorem Th5: :: WAYBEL23:5
for b1 being complete sup-Semilattice
for b2 being non empty full join-inheriting SubRelStr of b1 holds
( Bottom b1 in the carrier of b2 implies for b3 being Subset of b1
for b4 being Subset of b2 holds
( b3 = b4 implies finsups b4 c= finsups b3 ) )
proof end;

theorem Th6: :: WAYBEL23:6
for b1 being lower-bounded sup-Semilattice
for b2 being Subset of (InclPoset (Ids b1)) holds sup b2 = downarrow (finsups (union b2))
proof end;

theorem Th7: :: WAYBEL23:7
for b1 being reflexive transitive RelStr
for b2 being Subset of b1 holds downarrow (downarrow b2) = downarrow b2
proof end;

theorem Th8: :: WAYBEL23:8
for b1 being reflexive transitive RelStr
for b2 being Subset of b1 holds uparrow (uparrow b2) = uparrow b2
proof end;

theorem Th9: :: WAYBEL23:9
for b1 being non empty reflexive transitive RelStr
for b2 being Element of b1 holds downarrow (downarrow b2) = downarrow b2
proof end;

theorem Th10: :: WAYBEL23:10
for b1 being non empty reflexive transitive RelStr
for b2 being Element of b1 holds uparrow (uparrow b2) = uparrow b2
proof end;

theorem Th11: :: WAYBEL23:11
for b1 being non empty RelStr
for b2 being non empty SubRelStr of b1
for b3 being Subset of b1
for b4 being Subset of b2 holds
( b3 = b4 implies downarrow b4 c= downarrow b3 )
proof end;

theorem Th12: :: WAYBEL23:12
for b1 being non empty RelStr
for b2 being non empty SubRelStr of b1
for b3 being Subset of b1
for b4 being Subset of b2 holds
( b3 = b4 implies uparrow b4 c= uparrow b3 )
proof end;

theorem Th13: :: WAYBEL23:13
for b1 being non empty RelStr
for b2 being non empty SubRelStr of b1
for b3 being Element of b1
for b4 being Element of b2 holds
( b3 = b4 implies downarrow b4 c= downarrow b3 )
proof end;

theorem Th14: :: WAYBEL23:14
for b1 being non empty RelStr
for b2 being non empty SubRelStr of b1
for b3 being Element of b1
for b4 being Element of b2 holds
( b3 = b4 implies uparrow b4 c= uparrow b3 )
proof end;

definition
let c1 be non empty RelStr ;
let c2 be Subset of c1;
attr a2 is meet-closed means :Def1: :: WAYBEL23:def 1
subrelstr a2 is meet-inheriting;
end;

:: deftheorem Def1 defines meet-closed WAYBEL23:def 1 :
for b1 being non empty RelStr
for b2 being Subset of b1 holds
( b2 is meet-closed iff subrelstr b2 is meet-inheriting );

definition
let c1 be non empty RelStr ;
let c2 be Subset of c1;
attr a2 is join-closed means :Def2: :: WAYBEL23:def 2
subrelstr a2 is join-inheriting;
end;

:: deftheorem Def2 defines join-closed WAYBEL23:def 2 :
for b1 being non empty RelStr
for b2 being Subset of b1 holds
( b2 is join-closed iff subrelstr b2 is join-inheriting );

definition
let c1 be non empty RelStr ;
let c2 be Subset of c1;
attr a2 is infs-closed means :Def3: :: WAYBEL23:def 3
subrelstr a2 is infs-inheriting;
end;

:: deftheorem Def3 defines infs-closed WAYBEL23:def 3 :
for b1 being non empty RelStr
for b2 being Subset of b1 holds
( b2 is infs-closed iff subrelstr b2 is infs-inheriting );

definition
let c1 be non empty RelStr ;
let c2 be Subset of c1;
attr a2 is sups-closed means :Def4: :: WAYBEL23:def 4
subrelstr a2 is sups-inheriting;
end;

:: deftheorem Def4 defines sups-closed WAYBEL23:def 4 :
for b1 being non empty RelStr
for b2 being Subset of b1 holds
( b2 is sups-closed iff subrelstr b2 is sups-inheriting );

registration
let c1 be non empty RelStr ;
cluster infs-closed -> meet-closed Element of K10(the carrier of a1);
coherence
for b1 being Subset of c1 holds
( b1 is infs-closed implies b1 is meet-closed )
proof end;
cluster sups-closed -> join-closed Element of K10(the carrier of a1);
coherence
for b1 being Subset of c1 holds
( b1 is sups-closed implies b1 is join-closed )
proof end;
end;

registration
let c1 be non empty RelStr ;
cluster non empty meet-closed join-closed infs-closed sups-closed Element of K10(the carrier of a1);
existence
ex b1 being Subset of c1 st
( b1 is infs-closed & b1 is sups-closed & not b1 is empty )
proof end;
end;

theorem Th15: :: WAYBEL23:15
for b1 being non empty RelStr
for b2 being Subset of b1 holds
( b2 is meet-closed iff for b3, b4 being Element of b1 holds
( b3 in b2 & b4 in b2 & ex_inf_of {b3,b4},b1 implies inf {b3,b4} in b2 ) )
proof end;

theorem Th16: :: WAYBEL23:16
for b1 being non empty RelStr
for b2 being Subset of b1 holds
( b2 is join-closed iff for b3, b4 being Element of b1 holds
( b3 in b2 & b4 in b2 & ex_sup_of {b3,b4},b1 implies sup {b3,b4} in b2 ) )
proof end;

theorem Th17: :: WAYBEL23:17
for b1 being antisymmetric with_infima RelStr
for b2 being Subset of b1 holds
( b2 is meet-closed iff for b3, b4 being Element of b1 holds
( b3 in b2 & b4 in b2 implies inf {b3,b4} in b2 ) )
proof end;

theorem Th18: :: WAYBEL23:18
for b1 being antisymmetric with_suprema RelStr
for b2 being Subset of b1 holds
( b2 is join-closed iff for b3, b4 being Element of b1 holds
( b3 in b2 & b4 in b2 implies sup {b3,b4} in b2 ) )
proof end;

theorem Th19: :: WAYBEL23:19
for b1 being non empty RelStr
for b2 being Subset of b1 holds
( b2 is infs-closed iff for b3 being Subset of b2 holds
( ex_inf_of b3,b1 implies "/\" b3,b1 in b2 ) )
proof end;

theorem Th20: :: WAYBEL23:20
for b1 being non empty RelStr
for b2 being Subset of b1 holds
( b2 is sups-closed iff for b3 being Subset of b2 holds
( ex_sup_of b3,b1 implies "\/" b3,b1 in b2 ) )
proof end;

theorem Th21: :: WAYBEL23:21
for b1 being non empty transitive RelStr
for b2 being non empty infs-closed Subset of b1
for b3 being Subset of b2 holds
( ex_inf_of b3,b1 implies ( ex_inf_of b3, subrelstr b2 & "/\" b3,(subrelstr b2) = "/\" b3,b1 ) )
proof end;

theorem Th22: :: WAYBEL23:22
for b1 being non empty transitive RelStr
for b2 being non empty sups-closed Subset of b1
for b3 being Subset of b2 holds
( ex_sup_of b3,b1 implies ( ex_sup_of b3, subrelstr b2 & "\/" b3,(subrelstr b2) = "\/" b3,b1 ) )
proof end;

theorem Th23: :: WAYBEL23:23
for b1 being non empty transitive RelStr
for b2 being non empty meet-closed Subset of b1
for b3, b4 being Element of b2 holds
( ex_inf_of {b3,b4},b1 implies ( ex_inf_of {b3,b4}, subrelstr b2 & "/\" {b3,b4},(subrelstr b2) = "/\" {b3,b4},b1 ) )
proof end;

theorem Th24: :: WAYBEL23:24
for b1 being non empty transitive RelStr
for b2 being non empty join-closed Subset of b1
for b3, b4 being Element of b2 holds
( ex_sup_of {b3,b4},b1 implies ( ex_sup_of {b3,b4}, subrelstr b2 & "\/" {b3,b4},(subrelstr b2) = "\/" {b3,b4},b1 ) )
proof end;

theorem Th25: :: WAYBEL23:25
for b1 being transitive antisymmetric with_infima RelStr
for b2 being non empty meet-closed Subset of b1 holds subrelstr b2 is with_infima
proof end;

theorem Th26: :: WAYBEL23:26
for b1 being transitive antisymmetric with_suprema RelStr
for b2 being non empty join-closed Subset of b1 holds subrelstr b2 is with_suprema
proof end;

registration
let c1 be transitive antisymmetric with_infima RelStr ;
let c2 be non empty meet-closed Subset of c1;
cluster subrelstr a2 -> with_infima ;
coherence
subrelstr c2 is with_infima
by Th25;
end;

registration
let c1 be transitive antisymmetric with_suprema RelStr ;
let c2 be non empty join-closed Subset of c1;
cluster subrelstr a2 -> with_suprema ;
coherence
subrelstr c2 is with_suprema
by Th26;
end;

theorem Th27: :: WAYBEL23:27
for b1 being non empty transitive antisymmetric complete RelStr
for b2 being non empty infs-closed Subset of b1
for b3 being Subset of b2 holds "/\" b3,(subrelstr b2) = "/\" b3,b1
proof end;

theorem Th28: :: WAYBEL23:28
for b1 being non empty transitive antisymmetric complete RelStr
for b2 being non empty sups-closed Subset of b1
for b3 being Subset of b2 holds "\/" b3,(subrelstr b2) = "\/" b3,b1
proof end;

theorem Th29: :: WAYBEL23:29
for b1 being Semilattice
for b2 being meet-closed Subset of b1 holds b2 is filtered
proof end;

theorem Th30: :: WAYBEL23:30
for b1 being sup-Semilattice
for b2 being join-closed Subset of b1 holds b2 is directed
proof end;

registration
let c1 be Semilattice;
cluster meet-closed -> filtered Element of K10(the carrier of a1);
coherence
for b1 being Subset of c1 holds
( b1 is meet-closed implies b1 is filtered )
by Th29;
end;

registration
let c1 be sup-Semilattice;
cluster join-closed -> directed Element of K10(the carrier of a1);
coherence
for b1 being Subset of c1 holds
( b1 is join-closed implies b1 is directed )
by Th30;
end;

theorem Th31: :: WAYBEL23:31
for b1 being Semilattice
for b2 being non empty upper Subset of b1 holds
( b2 is Filter of b1 iff b2 is meet-closed )
proof end;

theorem Th32: :: WAYBEL23:32
for b1 being sup-Semilattice
for b2 being non empty lower Subset of b1 holds
( b2 is Ideal of b1 iff b2 is join-closed )
proof end;

theorem Th33: :: WAYBEL23:33
for b1 being non empty RelStr
for b2, b3 being join-closed Subset of b1 holds b2 /\ b3 is join-closed
proof end;

theorem Th34: :: WAYBEL23:34
for b1 being non empty RelStr
for b2, b3 being meet-closed Subset of b1 holds b2 /\ b3 is meet-closed
proof end;

theorem Th35: :: WAYBEL23:35
for b1 being sup-Semilattice
for b2 being Element of b1 holds downarrow b2 is join-closed
proof end;

theorem Th36: :: WAYBEL23:36
for b1 being Semilattice
for b2 being Element of b1 holds downarrow b2 is meet-closed
proof end;

theorem Th37: :: WAYBEL23:37
for b1 being sup-Semilattice
for b2 being Element of b1 holds uparrow b2 is join-closed
proof end;

theorem Th38: :: WAYBEL23:38
for b1 being Semilattice
for b2 being Element of b1 holds uparrow b2 is meet-closed
proof end;

registration
let c1 be sup-Semilattice;
let c2 be Element of c1;
cluster downarrow a2 -> join-closed ;
coherence
downarrow c2 is join-closed
by Th35;
cluster uparrow a2 -> directed join-closed ;
coherence
uparrow c2 is join-closed
by Th37;
end;

registration
let c1 be Semilattice;
let c2 be Element of c1;
cluster downarrow a2 -> filtered meet-closed ;
coherence
downarrow c2 is meet-closed
by Th36;
cluster uparrow a2 -> meet-closed ;
coherence
uparrow c2 is meet-closed
by Th38;
end;

theorem Th39: :: WAYBEL23:39
for b1 being sup-Semilattice
for b2 being Element of b1 holds waybelow b2 is join-closed
proof end;

theorem Th40: :: WAYBEL23:40
for b1 being Semilattice
for b2 being Element of b1 holds waybelow b2 is meet-closed
proof end;

theorem Th41: :: WAYBEL23:41
for b1 being sup-Semilattice
for b2 being Element of b1 holds wayabove b2 is join-closed
proof end;

registration
let c1 be sup-Semilattice;
let c2 be Element of c1;
cluster waybelow a2 -> join-closed ;
coherence
waybelow c2 is join-closed
by Th39;
cluster wayabove a2 -> directed join-closed ;
coherence
wayabove c2 is join-closed
by Th41;
end;

registration
let c1 be Semilattice;
let c2 be Element of c1;
cluster waybelow a2 -> filtered meet-closed ;
coherence
waybelow c2 is meet-closed
by Th40;
end;

definition
let c1 be TopStruct ;
func weight c1 -> Cardinal equals :: WAYBEL23:def 5
meet { (Card b1) where B is Basis of a1 : verum } ;
coherence
meet { (Card b1) where B is Basis of c1 : verum } is Cardinal
proof end;
end;

:: deftheorem Def5 defines weight WAYBEL23:def 5 :
for b1 being TopStruct holds weight b1 = meet { (Card b2) where B is Basis of b1 : verum } ;

definition
let c1 be TopStruct ;
attr a1 is second-countable means :: WAYBEL23:def 6
weight a1 c= omega ;
end;

:: deftheorem Def6 defines second-countable WAYBEL23:def 6 :
for b1 being TopStruct holds
( b1 is second-countable iff weight b1 c= omega );

definition
let c1 be continuous sup-Semilattice;
mode CLbasis of c1 -> Subset of a1 means :Def7: :: WAYBEL23:def 7
( a2 is join-closed & ( for b1 being Element of a1 holds b1 = sup ((waybelow b1) /\ a2) ) );
existence
ex b1 being Subset of c1 st
( b1 is join-closed & ( for b2 being Element of c1 holds b2 = sup ((waybelow b2) /\ b1) ) )
proof end;
end;

:: deftheorem Def7 defines CLbasis WAYBEL23:def 7 :
for b1 being continuous sup-Semilattice
for b2 being Subset of b1 holds
( b2 is CLbasis of b1 iff ( b2 is join-closed & ( for b3 being Element of b1 holds b3 = sup ((waybelow b3) /\ b2) ) ) );

definition
let c1 be non empty RelStr ;
let c2 be Subset of c1;
attr a2 is with_bottom means :Def8: :: WAYBEL23:def 8
Bottom a1 in a2;
end;

:: deftheorem Def8 defines with_bottom WAYBEL23:def 8 :
for b1 being non empty RelStr
for b2 being Subset of b1 holds
( b2 is with_bottom iff Bottom b1 in b2 );

definition
let c1 be non empty RelStr ;
let c2 be Subset of c1;
attr a2 is with_top means :Def9: :: WAYBEL23:def 9
Top a1 in a2;
end;

:: deftheorem Def9 defines with_top WAYBEL23:def 9 :
for b1 being non empty RelStr
for b2 being Subset of b1 holds
( b2 is with_top iff Top b1 in b2 );

registration
let c1 be non empty RelStr ;
cluster with_bottom -> non empty Element of K10(the carrier of a1);
coherence
for b1 being Subset of c1 holds
not ( b1 is with_bottom & b1 is empty )
by Def8;
end;

registration
let c1 be non empty RelStr ;
cluster with_top -> non empty Element of K10(the carrier of a1);
coherence
for b1 being Subset of c1 holds
not ( b1 is with_top & b1 is empty )
by Def9;
end;

registration
let c1 be non empty RelStr ;
cluster non empty with_bottom Element of K10(the carrier of a1);
existence
ex b1 being Subset of c1 st b1 is with_bottom
proof end;
cluster non empty with_top Element of K10(the carrier of a1);
existence
ex b1 being Subset of c1 st b1 is with_top
proof end;
end;

registration
let c1 be continuous sup-Semilattice;
cluster non empty with_bottom CLbasis of a1;
existence
ex b1 being CLbasis of c1 st b1 is with_bottom
proof end;
cluster non empty with_top CLbasis of a1;
existence
ex b1 being CLbasis of c1 st b1 is with_top
proof end;
end;

theorem Th42: :: WAYBEL23:42
for b1 being non empty antisymmetric lower-bounded RelStr
for b2 being with_bottom Subset of b1 holds subrelstr b2 is lower-bounded
proof end;

registration
let c1 be non empty antisymmetric lower-bounded RelStr ;
let c2 be with_bottom Subset of c1;
cluster subrelstr a2 -> lower-bounded ;
coherence
subrelstr c2 is lower-bounded
by Th42;
end;

registration
let c1 be continuous sup-Semilattice;
cluster -> directed join-closed CLbasis of a1;
coherence
for b1 being CLbasis of c1 holds b1 is join-closed
by Def7;
end;

registration
cluster non trivial bounded continuous RelStr ;
existence
ex b1 being continuous LATTICE st
( b1 is bounded & not b1 is trivial )
proof end;
end;

registration
let c1 be non trivial lower-bounded continuous sup-Semilattice;
cluster -> non empty directed join-closed CLbasis of a1;
coherence
for b1 being CLbasis of c1 holds
not b1 is empty
proof end;
end;

theorem Th43: :: WAYBEL23:43
for b1 being sup-Semilattice holds
the carrier of (CompactSublatt b1) is join-closed Subset of b1
proof end;

theorem Th44: :: WAYBEL23:44
for b1 being lower-bounded algebraic LATTICE holds
the carrier of (CompactSublatt b1) is with_bottom CLbasis of b1
proof end;

theorem Th45: :: WAYBEL23:45
for b1 being lower-bounded continuous sup-Semilattice holds
( the carrier of (CompactSublatt b1) is CLbasis of b1 implies b1 is algebraic )
proof end;

theorem Th46: :: WAYBEL23:46
for b1 being lower-bounded continuous LATTICE
for b2 being join-closed Subset of b1 holds
( b2 is CLbasis of b1 iff for b3, b4 being Element of b1 holds
not ( not b4 <= b3 & ( for b5 being Element of b1 holds
not ( b5 in b2 & not b5 <= b3 & b5 << b4 ) ) ) )
proof end;

theorem Th47: :: WAYBEL23:47
for b1 being lower-bounded continuous LATTICE
for b2 being join-closed Subset of b1 holds
( Bottom b1 in b2 implies ( b2 is CLbasis of b1 iff for b3, b4 being Element of b1 holds
not ( b3 << b4 & ( for b5 being Element of b1 holds
not ( b5 in b2 & b3 <= b5 & b5 << b4 ) ) ) ) )
proof end;

Lemma39: for b1 being lower-bounded continuous LATTICE
for b2 being join-closed Subset of b1 holds
( Bottom b1 in b2 & ( for b3, b4 being Element of b1 holds
not ( b3 << b4 & ( for b5 being Element of b1 holds
not ( b5 in b2 & b3 <= b5 & b5 << b4 ) ) ) ) implies ( the carrier of (CompactSublatt b1) c= b2 & ( for b3, b4 being Element of b1 holds
not ( not b4 <= b3 & ( for b5 being Element of b1 holds
not ( b5 in b2 & not b5 <= b3 & b5 <= b4 ) ) ) ) ) )
proof end;

Lemma40: for b1 being lower-bounded continuous LATTICE
for b2 being Subset of b1 holds
( ( for b3, b4 being Element of b1 holds
not ( not b4 <= b3 & ( for b5 being Element of b1 holds
not ( b5 in b2 & not b5 <= b3 & b5 <= b4 ) ) ) ) implies for b3, b4 being Element of b1 holds
not ( not b4 <= b3 & ( for b5 being Element of b1 holds
not ( b5 in b2 & not b5 <= b3 & b5 << b4 ) ) ) )
proof end;

theorem Th48: :: WAYBEL23:48
for b1 being lower-bounded continuous LATTICE
for b2 being join-closed Subset of b1 holds
( Bottom b1 in b2 implies ( b2 is CLbasis of b1 iff ( the carrier of (CompactSublatt b1) c= b2 & ( for b3, b4 being Element of b1 holds
not ( not b4 <= b3 & ( for b5 being Element of b1 holds
not ( b5 in b2 & not b5 <= b3 & b5 <= b4 ) ) ) ) ) ) )
proof end;

theorem Th49: :: WAYBEL23:49
for b1 being lower-bounded continuous LATTICE
for b2 being join-closed Subset of b1 holds
( Bottom b1 in b2 implies ( b2 is CLbasis of b1 iff for b3, b4 being Element of b1 holds
not ( not b4 <= b3 & ( for b5 being Element of b1 holds
not ( b5 in b2 & not b5 <= b3 & b5 <= b4 ) ) ) ) )
proof end;

theorem Th50: :: WAYBEL23:50
for b1 being lower-bounded sup-Semilattice
for b2 being non empty full SubRelStr of b1 holds
( Bottom b1 in the carrier of b2 & the carrier of b2 is join-closed Subset of b1 implies for b3 being Element of b1 holds
(waybelow b3) /\ the carrier of b2 is Ideal of b2 )
proof end;

definition
let c1 be non empty reflexive transitive RelStr ;
let c2 be non empty full SubRelStr of c1;
func supMap c2 -> Function of (InclPoset (Ids a2)),a1 means :Def10: :: WAYBEL23:def 10
for b1 being Ideal of a2 holds a3 . b1 = "\/" b1,a1;
existence
ex b1 being Function of (InclPoset (Ids c2)),c1 st
for b2 being Ideal of c2 holds b1 . b2 = "\/" b2,c1
proof end;
uniqueness
for b1, b2 being Function of (InclPoset (Ids c2)),c1 holds
( ( for b3 being Ideal of c2 holds b1 . b3 = "\/" b3,c1 ) & ( for b3 being Ideal of c2 holds b2 . b3 = "\/" b3,c1 ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def10 defines supMap WAYBEL23:def 10 :
for b1 being non empty reflexive transitive RelStr
for b2 being non empty full SubRelStr of b1
for b3 being Function of (InclPoset (Ids b2)),b1 holds
( b3 = supMap b2 iff for b4 being Ideal of b2 holds b3 . b4 = "\/" b4,b1 );

definition
let c1 be non empty reflexive transitive RelStr ;
let c2 be non empty full SubRelStr of c1;
func idsMap c2 -> Function of (InclPoset (Ids a2)),(InclPoset (Ids a1)) means :Def11: :: WAYBEL23:def 11
for b1 being Ideal of a2 holds
ex b2 being Subset of a1 st
( b1 = b2 & a3 . b1 = downarrow b2 );
existence
ex b1 being Function of (InclPoset (Ids c2)),(InclPoset (Ids c1)) st
for b2 being Ideal of c2 holds
ex b3 being Subset of c1 st
( b2 = b3 & b1 . b2 = downarrow b3 )
proof end;
uniqueness
for b1, b2 being Function of (InclPoset (Ids c2)),(InclPoset (Ids c1)) holds
( ( for b3 being Ideal of c2 holds
ex b4 being Subset of c1 st
( b3 = b4 & b1 . b3 = downarrow b4 ) ) & ( for b3 being Ideal of c2 holds
ex b4 being Subset of c1 st
( b3 = b4 & b2 . b3 = downarrow b4 ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def11 defines idsMap WAYBEL23:def 11 :
for b1 being non empty reflexive transitive RelStr
for b2 being non empty full SubRelStr of b1
for b3 being Function of (InclPoset (Ids b2)),(InclPoset (Ids b1)) holds
( b3 = idsMap b2 iff for b4 being Ideal of b2 holds
ex b5 being Subset of b1 st
( b4 = b5 & b3 . b4 = downarrow b5 ) );

registration
let c1 be reflexive RelStr ;
let c2 be Subset of c1;
cluster subrelstr a2 -> reflexive ;
coherence
subrelstr c2 is reflexive
;
end;

registration
let c1 be transitive RelStr ;
let c2 be Subset of c1;
cluster subrelstr a2 -> transitive ;
coherence
subrelstr c2 is transitive
;
end;

registration
let c1 be antisymmetric RelStr ;
let c2 be Subset of c1;
cluster subrelstr a2 -> antisymmetric ;
coherence
subrelstr c2 is antisymmetric
;
end;

definition
let c1 be lower-bounded continuous sup-Semilattice;
let c2 be with_bottom CLbasis of c1;
func baseMap c2 -> Function of a1,(InclPoset (Ids (subrelstr a2))) means :Def12: :: WAYBEL23:def 12
for b1 being Element of a1 holds a3 . b1 = (waybelow b1) /\ a2;
existence
ex b1 being Function of c1,(InclPoset (Ids (subrelstr c2))) st
for b2 being Element of c1 holds b1 . b2 = (waybelow b2) /\ c2
proof end;
uniqueness
for b1, b2 being Function of c1,(InclPoset (Ids (subrelstr c2))) holds
( ( for b3 being Element of c1 holds b1 . b3 = (waybelow b3) /\ c2 ) & ( for b3 being Element of c1 holds b2 . b3 = (waybelow b3) /\ c2 ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def12 defines baseMap WAYBEL23:def 12 :
for b1 being lower-bounded continuous sup-Semilattice
for b2 being with_bottom CLbasis of b1
for b3 being Function of b1,(InclPoset (Ids (subrelstr b2))) holds
( b3 = baseMap b2 iff for b4 being Element of b1 holds b3 . b4 = (waybelow b4) /\ b2 );

theorem Th51: :: WAYBEL23:51
for b1 being non empty reflexive transitive RelStr
for b2 being non empty full SubRelStr of b1 holds
( dom (supMap b2) = Ids b2 & rng (supMap b2) is Subset of b1 )
proof end;

theorem Th52: :: WAYBEL23:52
for b1 being non empty reflexive transitive RelStr
for b2 being non empty full SubRelStr of b1
for b3 being set holds
( b3 in dom (supMap b2) iff b3 is Ideal of b2 )
proof end;

theorem Th53: :: WAYBEL23:53
for b1 being non empty reflexive transitive RelStr
for b2 being non empty full SubRelStr of b1 holds
( dom (idsMap b2) = Ids b2 & rng (idsMap b2) is Subset of (Ids b1) )
proof end;

theorem Th54: :: WAYBEL23:54
for b1 being non empty reflexive transitive RelStr
for b2 being non empty full SubRelStr of b1
for b3 being set holds
( b3 in dom (idsMap b2) iff b3 is Ideal of b2 )
proof end;

theorem Th55: :: WAYBEL23:55
for b1 being non empty reflexive transitive RelStr
for b2 being non empty full SubRelStr of b1
for b3 being set holds
( b3 in rng (idsMap b2) implies b3 is Ideal of b1 )
proof end;

theorem Th56: :: WAYBEL23:56
for b1 being lower-bounded continuous sup-Semilattice
for b2 being with_bottom CLbasis of b1 holds
( dom (baseMap b2) = the carrier of b1 & rng (baseMap b2) is Subset of (Ids (subrelstr b2)) ) by FUNCT_2:def 1, YELLOW_1:1;

theorem Th57: :: WAYBEL23:57
for b1 being lower-bounded continuous sup-Semilattice
for b2 being with_bottom CLbasis of b1
for b3 being set holds
( b3 in rng (baseMap b2) implies b3 is Ideal of (subrelstr b2) )
proof end;

theorem Th58: :: WAYBEL23:58
for b1 being non empty up-complete Poset
for b2 being non empty full SubRelStr of b1 holds supMap b2 is monotone
proof end;

theorem Th59: :: WAYBEL23:59
for b1 being non empty reflexive transitive RelStr
for b2 being non empty full SubRelStr of b1 holds idsMap b2 is monotone
proof end;

theorem Th60: :: WAYBEL23:60
for b1 being lower-bounded continuous sup-Semilattice
for b2 being with_bottom CLbasis of b1 holds baseMap b2 is monotone
proof end;

registration
let c1 be non empty up-complete Poset;
let c2 be non empty full SubRelStr of c1;
cluster supMap a2 -> monotone ;
coherence
supMap c2 is monotone
by Th58;
end;

registration
let c1 be non empty reflexive transitive RelStr ;
let c2 be non empty full SubRelStr of c1;
cluster idsMap a2 -> monotone ;
coherence
idsMap c2 is monotone
by Th59;
end;

registration
let c1 be lower-bounded continuous sup-Semilattice;
let c2 be with_bottom CLbasis of c1;
cluster baseMap a2 -> monotone ;
coherence
baseMap c2 is monotone
by Th60;
end;

theorem Th61: :: WAYBEL23:61
for b1 being lower-bounded continuous sup-Semilattice
for b2 being with_bottom CLbasis of b1 holds idsMap (subrelstr b2) is sups-preserving
proof end;

theorem Th62: :: WAYBEL23:62
for b1 being non empty up-complete Poset
for b2 being non empty full SubRelStr of b1 holds supMap b2 = (SupMap b1) * (idsMap b2)
proof end;

theorem Th63: :: WAYBEL23:63
for b1 being lower-bounded continuous sup-Semilattice
for b2 being with_bottom CLbasis of b1 holds [(supMap (subrelstr b2)),(baseMap b2)] is Galois
proof end;

theorem Th64: :: WAYBEL23:64
for b1 being lower-bounded continuous sup-Semilattice
for b2 being with_bottom CLbasis of b1 holds
( supMap (subrelstr b2) is upper_adjoint & baseMap b2 is lower_adjoint )
proof end;

theorem Th65: :: WAYBEL23:65
for b1 being lower-bounded continuous sup-Semilattice
for b2 being with_bottom CLbasis of b1 holds rng (supMap (subrelstr b2)) = the carrier of b1
proof end;

theorem Th66: :: WAYBEL23:66
for b1 being lower-bounded continuous sup-Semilattice
for b2 being with_bottom CLbasis of b1 holds
( supMap (subrelstr b2) is infs-preserving & supMap (subrelstr b2) is sups-preserving )
proof end;

theorem Th67: :: WAYBEL23:67
for b1 being lower-bounded continuous sup-Semilattice
for b2 being with_bottom CLbasis of b1 holds baseMap b2 is sups-preserving
proof end;

registration
let c1 be lower-bounded continuous sup-Semilattice;
let c2 be with_bottom CLbasis of c1;
cluster supMap (subrelstr a2) -> monotone infs-preserving sups-preserving ;
coherence
( supMap (subrelstr c2) is infs-preserving & supMap (subrelstr c2) is sups-preserving )
by Th66;
cluster baseMap a2 -> monotone sups-preserving ;
coherence
baseMap c2 is sups-preserving
by Th67;
end;

theorem Th68: :: WAYBEL23:68
canceled;

theorem Th69: :: WAYBEL23:69
for b1 being lower-bounded continuous sup-Semilattice
for b2 being with_bottom CLbasis of b1 holds the carrier of (CompactSublatt (InclPoset (Ids (subrelstr b2)))) = { (downarrow b3) where B is Element of (subrelstr b2) : verum }
proof end;

theorem Th70: :: WAYBEL23:70
for b1 being lower-bounded continuous sup-Semilattice
for b2 being with_bottom CLbasis of b1 holds CompactSublatt (InclPoset (Ids (subrelstr b2))), subrelstr b2 are_isomorphic
proof end;

Lemma62: for b1 being lower-bounded continuous LATTICE holds
( b1 is algebraic implies ( the carrier of (CompactSublatt b1) is with_bottom CLbasis of b1 & ( for b2 being with_bottom CLbasis of b1 holds the carrier of (CompactSublatt b1) c= b2 ) ) )
proof end;

theorem Th71: :: WAYBEL23:71
for b1 being lower-bounded continuous LATTICE
for b2 being with_bottom CLbasis of b1 holds
( ( for b3 being with_bottom CLbasis of b1 holds b2 c= b3 ) implies for b3 being Element of (InclPoset (Ids (subrelstr b2))) holds b3 = (waybelow ("\/" b3,b1)) /\ b2 )
proof end;

Lemma64: for b1 being lower-bounded continuous LATTICE holds
( ex b2 being with_bottom CLbasis of b1 st
for b3 being with_bottom CLbasis of b1 holds b2 c= b3 implies b1 is algebraic )
proof end;

theorem Th72: :: WAYBEL23:72
for b1 being lower-bounded continuous LATTICE holds
( b1 is algebraic iff ( the carrier of (CompactSublatt b1) is with_bottom CLbasis of b1 & ( for b2 being with_bottom CLbasis of b1 holds the carrier of (CompactSublatt b1) c= b2 ) ) ) by Lemma62, Lemma64;

theorem Th73: :: WAYBEL23:73
for b1 being lower-bounded continuous LATTICE holds
( b1 is algebraic iff ex b2 being with_bottom CLbasis of b1 st
for b3 being with_bottom CLbasis of b1 holds b2 c= b3 )
proof end;

theorem Th74: :: WAYBEL23:74
for b1 being TopStruct
for b2 being Basis of b1 holds weight b1 c= Card b2
proof end;

theorem Th75: :: WAYBEL23:75
for b1 being TopStruct holds
ex b2 being Basis of b1 st Card b2 = weight b1
proof end;