begin
:: deftheorem defines one-to-one WAYBEL_1:def 1 :
for L1, L2 being non empty 1-sorted
for f being Function of L1,L2 holds
( f is one-to-one iff for x, y being Element of L1 st f . x = f . y holds
x = y );
:: deftheorem Def2 defines monotone WAYBEL_1:def 2 :
for L1, L2 being non empty RelStr
for f being Function of L1,L2 holds
( f is monotone iff for x, y being Element of L1 st x <= y holds
f . x <= f . y );
theorem
canceled;
theorem Th2:
theorem Th3:
theorem Th4:
theorem Th5:
:: deftheorem Def3 defines distributive WAYBEL_1:def 3 :
for L being non empty RelStr holds
( L is distributive iff for x, y, z being Element of L holds x "/\" (y "\/" z) = (x "/\" y) "\/" (x "/\" z) );
theorem Th6:
:: deftheorem defines ex_min_of WAYBEL_1:def 4 :
for S being non empty RelStr
for X being set holds
( ex_min_of X,S iff ( ex_inf_of X,S & "/\" (X,S) in X ) );
:: deftheorem Def5 defines ex_max_of WAYBEL_1:def 5 :
for S being non empty RelStr
for X being set holds
( ex_max_of X,S iff ( ex_sup_of X,S & "\/" (X,S) in X ) );
:: deftheorem Def6 defines is_minimum_of WAYBEL_1:def 6 :
for S being non empty RelStr
for s being Element of S
for X being set holds
( s is_minimum_of X iff ( ex_inf_of X,S & s = "/\" (X,S) & "/\" (X,S) in X ) );
:: deftheorem Def7 defines is_maximum_of WAYBEL_1:def 7 :
for S being non empty RelStr
for s being Element of S
for X being set holds
( s is_maximum_of X iff ( ex_sup_of X,S & s = "\/" (X,S) & "\/" (X,S) in X ) );
:: deftheorem Def8 defines are_isomorphic WAYBEL_1:def 8 :
for L1, L2 being RelStr holds
( L1,L2 are_isomorphic iff ex f being Function of L1,L2 st f is isomorphic );
theorem
theorem
begin
:: deftheorem Def9 defines Connection WAYBEL_1:def 9 :
for S, T being RelStr
for b3 being set holds
( b3 is Connection of S,T iff ex g being Function of S,T ex d being Function of T,S st b3 = [g,d] );
:: deftheorem Def10 defines Galois WAYBEL_1:def 10 :
for S, T being non empty RelStr
for gc being Connection of S,T holds
( gc is Galois iff ex g being Function of S,T ex d being Function of T,S st
( gc = [g,d] & g is monotone & d is monotone & ( for t being Element of T
for s being Element of S holds
( t <= g . s iff d . t <= s ) ) ) );
theorem Th9:
:: deftheorem Def11 defines upper_adjoint WAYBEL_1:def 11 :
for S, T being non empty RelStr
for g being Function of S,T holds
( g is upper_adjoint iff ex d being Function of T,S st [g,d] is Galois );
:: deftheorem Def12 defines lower_adjoint WAYBEL_1:def 12 :
for S, T being non empty RelStr
for d being Function of T,S holds
( d is lower_adjoint iff ex g being Function of S,T st [g,d] is Galois );
theorem Th10:
theorem Th11:
theorem Th12:
theorem Th13:
theorem Th14:
theorem Th15:
theorem Th16:
theorem
theorem Th18:
theorem Th19:
theorem Th20:
theorem Th21:
theorem Th22:
theorem Th23:
theorem Th24:
theorem
canceled;
theorem
theorem Th27:
theorem Th28:
theorem
:: deftheorem Def13 defines projection WAYBEL_1:def 13 :
for L being non empty RelStr
for p being Function of L,L holds
( p is projection iff ( p is idempotent & p is monotone ) );
:: deftheorem Def14 defines closure WAYBEL_1:def 14 :
for L being non empty RelStr
for c being Function of L,L holds
( c is closure iff ( c is projection & id L <= c ) );
Lm1:
for L1, L2 being non empty RelStr
for f being Function of L1,L2 st L2 is reflexive holds
f <= f
:: deftheorem Def15 defines kernel WAYBEL_1:def 15 :
for L being non empty RelStr
for k being Function of L,L holds
( k is kernel iff ( k is projection & k <= id L ) );
Lm2:
for L being non empty 1-sorted
for p being Function of L,L st p is idempotent holds
for x being set st x in rng p holds
p . x = x
theorem Th30:
theorem Th31:
:: deftheorem defines corestr WAYBEL_1:def 16 :
for L1, L2 being non empty RelStr
for g being Function of L1,L2 holds corestr g = the carrier of (Image g) | g;
theorem Th32:
Lm3:
for L1, L2 being non empty RelStr
for g being Function of L1,L2 holds corestr g is onto
theorem Th33:
:: deftheorem defines inclusion WAYBEL_1:def 17 :
for L1, L2 being non empty RelStr
for g being Function of L1,L2 holds inclusion g = id (Image g);
Lm4:
for L1, L2 being non empty RelStr
for g being Function of L1,L2 holds inclusion g is monotone
theorem
canceled;
theorem Th35:
theorem Th36:
theorem
theorem
theorem Th39:
theorem
theorem Th41:
theorem Th42:
theorem
theorem
theorem Th45:
theorem Th46:
theorem Th47:
theorem Th48:
theorem
theorem Th50:
theorem
theorem Th52:
theorem Th53:
theorem
theorem
theorem Th56:
theorem
theorem
theorem
begin
theorem Th60:
theorem
:: deftheorem Def18 defines "/\" WAYBEL_1:def 18 :
for S being non empty RelStr
for x being Element of S
for b3 being Function of S,S holds
( b3 = x "/\" iff for s being Element of S holds b3 . s = x "/\" s );
theorem Th62:
theorem Th63:
theorem Th64:
theorem Th65:
theorem Th66:
theorem
theorem Th68:
:: deftheorem Def19 defines Heyting WAYBEL_1:def 19 :
for H being non empty RelStr holds
( H is Heyting iff ( H is LATTICE & ( for x being Element of H holds x "/\" is lower_adjoint ) ) );
:: deftheorem Def20 defines => WAYBEL_1:def 20 :
for H being non empty RelStr
for a being Element of H st H is Heyting holds
for b3 being Function of H,H holds
( b3 = a => iff [b3,(a "/\")] is Galois );
theorem Th69:
:: deftheorem defines => WAYBEL_1:def 21 :
for H being non empty RelStr
for a, y being Element of H holds a => y = (a =>) . y;
theorem Th70:
theorem Th71:
theorem Th72:
theorem
theorem
theorem Th75:
theorem
theorem
Lm5:
for H being non empty RelStr st H is Heyting holds
for a, b being Element of H holds a "/\" (a => b) <= b
theorem Th78:
theorem
theorem Th80:
theorem Th81:
:: deftheorem defines 'not' WAYBEL_1:def 22 :
for H being non empty RelStr
for a being Element of H holds 'not' a = a => (Bottom H);
theorem
theorem Th83:
theorem
theorem Th85:
theorem
theorem
theorem
:: deftheorem Def23 defines is_a_complement_of WAYBEL_1:def 23 :
for L being non empty RelStr
for x, y being Element of L holds
( y is_a_complement_of x iff ( x "\/" y = Top L & x "/\" y = Bottom L ) );
:: deftheorem Def24 defines complemented WAYBEL_1:def 24 :
for L being non empty RelStr holds
( L is complemented iff for x being Element of L ex y being Element of L st y is_a_complement_of x );
Lm6:
for L being bounded LATTICE st L is distributive & L is complemented holds
for x being Element of L ex x9 being Element of L st
for y being Element of L holds
( (y "\/" x9) "/\" x <= y & y <= (y "/\" x) "\/" x9 )
Lm7:
for L being bounded LATTICE st ( for x being Element of L ex x9 being Element of L st
for y being Element of L holds
( (y "\/" x9) "/\" x <= y & y <= (y "/\" x) "\/" x9 ) ) holds
( L is Heyting & ( for x being Element of L holds 'not' ('not' x) = x ) )
theorem Th89:
theorem Th90:
:: deftheorem Def25 defines Boolean WAYBEL_1:def 25 :
for B being non empty RelStr holds
( B is Boolean iff ( B is LATTICE & B is bounded & B is distributive & B is complemented ) );