begin
:: deftheorem defines reflexive YELLOW_0:def 1 :
for A being non empty RelStr holds
( A is reflexive iff for x being Element of A holds x <= x );
:: deftheorem defines transitive YELLOW_0:def 2 :
for A being RelStr holds
( A is transitive iff for x, y, z being Element of A st x <= y & y <= z holds
x <= z );
:: deftheorem defines antisymmetric YELLOW_0:def 3 :
for A being RelStr holds
( A is antisymmetric iff for x, y being Element of A st x <= y & y <= x holds
x = y );
theorem Th1:
theorem Th2:
theorem
theorem Th4:
theorem Th5:
theorem Th6:
theorem Th7:
theorem Th8:
theorem Th9:
theorem Th10:
theorem Th11:
theorem Th12:
begin
:: deftheorem Def4 defines lower-bounded YELLOW_0:def 4 :
for L being RelStr holds
( L is lower-bounded iff ex x being Element of L st x is_<=_than the carrier of L );
:: deftheorem Def5 defines upper-bounded YELLOW_0:def 5 :
for L being RelStr holds
( L is upper-bounded iff ex x being Element of L st x is_>=_than the carrier of L );
:: deftheorem defines bounded YELLOW_0:def 6 :
for L being RelStr holds
( L is bounded iff ( L is lower-bounded & L is upper-bounded ) );
theorem
:: deftheorem Def7 defines ex_sup_of YELLOW_0:def 7 :
for L being RelStr
for X being set holds
( ex_sup_of X,L iff ex a being Element of L st
( X is_<=_than a & ( for b being Element of L st X is_<=_than b holds
b >= a ) & ( for c being Element of L st X is_<=_than c & ( for b being Element of L st X is_<=_than b holds
b >= c ) holds
c = a ) ) );
:: deftheorem Def8 defines ex_inf_of YELLOW_0:def 8 :
for L being RelStr
for X being set holds
( ex_inf_of X,L iff ex a being Element of L st
( X is_>=_than a & ( for b being Element of L st X is_>=_than b holds
b <= a ) & ( for c being Element of L st X is_>=_than c & ( for b being Element of L st X is_>=_than b holds
b <= c ) holds
c = a ) ) );
theorem Th14:
theorem Th15:
theorem Th16:
theorem Th17:
theorem Th18:
theorem Th19:
theorem Th20:
theorem Th21:
theorem Th22:
theorem Th23:
theorem
theorem
definition
let L be
RelStr ;
let X be
set ;
func "\/" (
X,
L)
-> Element of
L means :
Def9:
(
X is_<=_than it & ( for
a being
Element of
L st
X is_<=_than a holds
it <= a ) )
if ex_sup_of X,
L;
uniqueness
for b1, b2 being Element of L st ex_sup_of X,L & X is_<=_than b1 & ( for a being Element of L st X is_<=_than a holds
b1 <= a ) & X is_<=_than b2 & ( for a being Element of L st X is_<=_than a holds
b2 <= a ) holds
b1 = b2
existence
( ex_sup_of X,L implies ex b1 being Element of L st
( X is_<=_than b1 & ( for a being Element of L st X is_<=_than a holds
b1 <= a ) ) )
correctness
consistency
for b1 being Element of L holds verum;
;
func "/\" (
X,
L)
-> Element of
L means :
Def10:
(
X is_>=_than it & ( for
a being
Element of
L st
X is_>=_than a holds
a <= it ) )
if ex_inf_of X,
L;
uniqueness
for b1, b2 being Element of L st ex_inf_of X,L & X is_>=_than b1 & ( for a being Element of L st X is_>=_than a holds
a <= b1 ) & X is_>=_than b2 & ( for a being Element of L st X is_>=_than a holds
a <= b2 ) holds
b1 = b2
existence
( ex_inf_of X,L implies ex b1 being Element of L st
( X is_>=_than b1 & ( for a being Element of L st X is_>=_than a holds
a <= b1 ) ) )
correctness
consistency
for b1 being Element of L holds verum;
;
end;
:: deftheorem Def9 defines "\/" YELLOW_0:def 9 :
for L being RelStr
for X being set
for b3 being Element of L st ex_sup_of X,L holds
( b3 = "\/" (X,L) iff ( X is_<=_than b3 & ( for a being Element of L st X is_<=_than a holds
b3 <= a ) ) );
:: deftheorem Def10 defines "/\" YELLOW_0:def 10 :
for L being RelStr
for X being set
for b3 being Element of L st ex_inf_of X,L holds
( b3 = "/\" (X,L) iff ( X is_>=_than b3 & ( for a being Element of L st X is_>=_than a holds
a <= b3 ) ) );
theorem
theorem
theorem
theorem
theorem Th30:
theorem Th31:
theorem
theorem
theorem Th34:
theorem Th35:
theorem
theorem
theorem Th38:
theorem
theorem Th40:
theorem Th41:
theorem Th42:
theorem Th43:
:: deftheorem defines Bottom YELLOW_0:def 11 :
for L being RelStr holds Bottom L = "\/" ({},L);
:: deftheorem defines Top YELLOW_0:def 12 :
for L being RelStr holds Top L = "/\" ({},L);
theorem
theorem
theorem Th46:
theorem Th47:
theorem Th48:
theorem Th49:
theorem Th50:
theorem
theorem
theorem
theorem
theorem
begin
:: deftheorem Def13 defines SubRelStr YELLOW_0:def 13 :
for L, b2 being RelStr holds
( b2 is SubRelStr of L iff ( the carrier of b2 c= the carrier of L & the InternalRel of b2 c= the InternalRel of L ) );
:: deftheorem Def14 defines full YELLOW_0:def 14 :
for L being RelStr
for S being SubRelStr of L holds
( S is full iff the InternalRel of S = the InternalRel of L |_2 the carrier of S );
theorem
canceled;
theorem Th57:
theorem Th58:
:: deftheorem defines subrelstr YELLOW_0:def 15 :
for L being RelStr
for X being Subset of L
for b3 being strict full SubRelStr of L holds
( b3 = subrelstr X iff the carrier of b3 = X );
theorem Th59:
theorem Th60:
theorem Th61:
theorem Th62:
theorem Th63:
:: deftheorem Def16 defines meet-inheriting YELLOW_0:def 16 :
for L being non empty RelStr
for S being SubRelStr of L holds
( S is meet-inheriting iff for x, y being Element of L st x in the carrier of S & y in the carrier of S & ex_inf_of {x,y},L holds
inf {x,y} in the carrier of S );
:: deftheorem Def17 defines join-inheriting YELLOW_0:def 17 :
for L being non empty RelStr
for S being SubRelStr of L holds
( S is join-inheriting iff for x, y being Element of L st x in the carrier of S & y in the carrier of S & ex_sup_of {x,y},L holds
sup {x,y} in the carrier of S );
:: deftheorem defines infs-inheriting YELLOW_0:def 18 :
for L being non empty RelStr
for S being SubRelStr of L holds
( S is infs-inheriting iff for X being Subset of S st ex_inf_of X,L holds
"/\" (X,L) in the carrier of S );
:: deftheorem defines sups-inheriting YELLOW_0:def 19 :
for L being non empty RelStr
for S being SubRelStr of L holds
( S is sups-inheriting iff for X being Subset of S st ex_sup_of X,L holds
"\/" (X,L) in the carrier of S );
theorem Th64:
theorem Th65:
theorem
for
L being non
empty transitive RelStr for
S being non
empty full SubRelStr of
L for
x,
y being
Element of
S st
ex_inf_of {x,y},
L &
"/\" (
{x,y},
L)
in the
carrier of
S holds
(
ex_inf_of {x,y},
S &
"/\" (
{x,y},
S)
= "/\" (
{x,y},
L) )
by Th64;
theorem
for
L being non
empty transitive RelStr for
S being non
empty full SubRelStr of
L for
x,
y being
Element of
S st
ex_sup_of {x,y},
L &
"\/" (
{x,y},
L)
in the
carrier of
S holds
(
ex_sup_of {x,y},
S &
"\/" (
{x,y},
S)
= "\/" (
{x,y},
L) )
by Th65;
theorem
theorem
theorem
theorem