:: OSALG_2 semantic presentation

theorem Th1: :: OSALG_2:1
for b1 being non empty Poset
for b2, b3 being OrderSortedSet of b1 holds
b2 /\ b3 is OrderSortedSet of b1
proof end;

theorem Th2: :: OSALG_2:2
for b1 being non empty Poset
for b2, b3 being OrderSortedSet of b1 holds
b2 \/ b3 is OrderSortedSet of b1
proof end;

definition
let c1 be non empty Poset;
let c2 be OrderSortedSet of c1;
mode OrderSortedSubset of c2 -> ManySortedSubset of a2 means :Def1: :: OSALG_2:def 1
a3 is OrderSortedSet of a1;
existence
ex b1 being ManySortedSubset of c2 st
b1 is OrderSortedSet of c1
proof end;
end;

:: deftheorem Def1 defines OrderSortedSubset OSALG_2:def 1 :
for b1 being non empty Poset
for b2 being OrderSortedSet of b1
for b3 being ManySortedSubset of b2 holds
( b3 is OrderSortedSubset of b2 iff b3 is OrderSortedSet of b1 );

registration
let c1 be non empty Poset;
let c2 be V5 OrderSortedSet of c1;
cluster V5 OrderSortedSubset of a2;
existence
ex b1 being OrderSortedSubset of c2 st b1 is non-empty
proof end;
end;

definition
let c1 be OrderSortedSign;
let c2 be OSAlgebra of c1;
mode OSSubset of c2 -> ManySortedSubset of the Sorts of a2 means :Def2: :: OSALG_2:def 2
a3 is OrderSortedSet of a1;
existence
ex b1 being ManySortedSubset of the Sorts of c2 st
b1 is OrderSortedSet of c1
proof end;
end;

:: deftheorem Def2 defines OSSubset OSALG_2:def 2 :
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being ManySortedSubset of the Sorts of b2 holds
( b3 is OSSubset of b2 iff b3 is OrderSortedSet of b1 );

registration
let c1 be OrderSortedSign;
cluster strict non-empty monotone MSAlgebra of a1;
existence
ex b1 being OSAlgebra of c1 st
( b1 is monotone & b1 is strict & b1 is non-empty )
proof end;
end;

registration
let c1 be OrderSortedSign;
let c2 be non-empty OSAlgebra of c1;
cluster V5 OSSubset of a2;
existence
ex b1 being OSSubset of c2 st b1 is non-empty
proof end;
end;

theorem Th3: :: OSALG_2:3
for b1 being non empty strict non void all-with_const_op ManySortedSign holds OSSign b1 is all-with_const_op
proof end;

registration
cluster all-with_const_op strict OverloadedRSSign ;
existence
ex b1 being OrderSortedSign st
( b1 is all-with_const_op & b1 is strict )
proof end;
end;

theorem Th4: :: OSALG_2:4
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1 holds MSAlgebra(# the Sorts of b2,the Charact of b2 #) is order-sorted
proof end;

registration
let c1 be OrderSortedSign;
let c2 be OSAlgebra of c1;
cluster order-sorted MSSubAlgebra of a2;
existence
ex b1 being MSSubAlgebra of c2 st b1 is order-sorted
proof end;
end;

definition
let c1 be OrderSortedSign;
let c2 be OSAlgebra of c1;
mode OSSubAlgebra is order-sorted MSSubAlgebra of a2;
end;

registration
let c1 be OrderSortedSign;
let c2 be OSAlgebra of c1;
cluster strict MSSubAlgebra of a2;
existence
ex b1 being OSSubAlgebra of c2 st b1 is strict
proof end;
end;

registration
let c1 be OrderSortedSign;
let c2 be non-empty OSAlgebra of c1;
cluster strict non-empty MSSubAlgebra of a2;
existence
ex b1 being OSSubAlgebra of c2 st
( b1 is non-empty & b1 is strict )
proof end;
end;

theorem Th5: :: OSALG_2:5
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being MSAlgebra of b1 holds
( b3 is OSSubAlgebra of b2 iff ( the Sorts of b3 is OSSubset of b2 & ( for b4 being OSSubset of b2 holds
( b4 = the Sorts of b3 implies ( b4 is opers_closed & the Charact of b3 = Opers b2,b4 ) ) ) ) )
proof end;

definition
let c1 be OrderSortedSign;
let c2 be OSAlgebra of c1;
let c3 be SortSymbol of c1;
func OSConstants c2,c3 -> Subset of (the Sorts of a2 . a3) equals :: OSALG_2:def 3
union { (Constants a2,b1) where B is SortSymbol of a1 : b1 <= a3 } ;
coherence
union { (Constants c2,b1) where B is SortSymbol of c1 : b1 <= c3 } is Subset of (the Sorts of c2 . c3)
proof end;
end;

:: deftheorem Def3 defines OSConstants OSALG_2:def 3 :
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being SortSymbol of b1 holds OSConstants b2,b3 = union { (Constants b2,b4) where B is SortSymbol of b1 : b4 <= b3 } ;

theorem Th6: :: OSALG_2:6
canceled;

theorem Th7: :: OSALG_2:7
canceled;

theorem Th8: :: OSALG_2:8
canceled;

theorem Th9: :: OSALG_2:9
canceled;

theorem Th10: :: OSALG_2:10
canceled;

theorem Th11: :: OSALG_2:11
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being SortSymbol of b1 holds Constants b2,b3 c= OSConstants b2,b3
proof end;

definition
let c1 be OrderSortedSign;
let c2 be ManySortedSet of the carrier of c1;
func OSCl c2 -> OrderSortedSet of a1 means :Def4: :: OSALG_2:def 4
for b1 being SortSymbol of a1 holds a3 . b1 = union { (a2 . b2) where B is SortSymbol of a1 : b2 <= b1 } ;
existence
ex b1 being OrderSortedSet of c1 st
for b2 being SortSymbol of c1 holds b1 . b2 = union { (c2 . b3) where B is SortSymbol of c1 : b3 <= b2 }
proof end;
uniqueness
for b1, b2 being OrderSortedSet of c1 holds
( ( for b3 being SortSymbol of c1 holds b1 . b3 = union { (c2 . b4) where B is SortSymbol of c1 : b4 <= b3 } ) & ( for b3 being SortSymbol of c1 holds b2 . b3 = union { (c2 . b4) where B is SortSymbol of c1 : b4 <= b3 } ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def4 defines OSCl OSALG_2:def 4 :
for b1 being OrderSortedSign
for b2 being ManySortedSet of the carrier of b1
for b3 being OrderSortedSet of b1 holds
( b3 = OSCl b2 iff for b4 being SortSymbol of b1 holds b3 . b4 = union { (b2 . b5) where B is SortSymbol of b1 : b5 <= b4 } );

theorem Th12: :: OSALG_2:12
for b1 being OrderSortedSign
for b2 being ManySortedSet of the carrier of b1 holds b2 c= OSCl b2
proof end;

theorem Th13: :: OSALG_2:13
for b1 being OrderSortedSign
for b2 being ManySortedSet of the carrier of b1
for b3 being OrderSortedSet of b1 holds
( b2 c= b3 implies OSCl b2 c= b3 )
proof end;

theorem Th14: :: OSALG_2:14
for b1 being OrderSortedSign
for b2 being OrderSortedSet of b1 holds OSCl b2 = b2
proof end;

definition
let c1 be OrderSortedSign;
let c2 be OSAlgebra of c1;
func OSConstants c2 -> OSSubset of a2 means :Def5: :: OSALG_2:def 5
for b1 being SortSymbol of a1 holds a3 . b1 = OSConstants a2,b1;
existence
ex b1 being OSSubset of c2 st
for b2 being SortSymbol of c1 holds b1 . b2 = OSConstants c2,b2
proof end;
uniqueness
for b1, b2 being OSSubset of c2 holds
( ( for b3 being SortSymbol of c1 holds b1 . b3 = OSConstants c2,b3 ) & ( for b3 being SortSymbol of c1 holds b2 . b3 = OSConstants c2,b3 ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def5 defines OSConstants OSALG_2:def 5 :
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being OSSubset of b2 holds
( b3 = OSConstants b2 iff for b4 being SortSymbol of b1 holds b3 . b4 = OSConstants b2,b4 );

theorem Th15: :: OSALG_2:15
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1 holds Constants b2 c= OSConstants b2
proof end;

theorem Th16: :: OSALG_2:16
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being OSSubset of b2 holds
( Constants b2 c= b3 implies OSConstants b2 c= b3 )
proof end;

theorem Th17: :: OSALG_2:17
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being OSSubset of b2 holds OSConstants b2 = OSCl (Constants b2)
proof end;

theorem Th18: :: OSALG_2:18
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being OSSubAlgebra of b2 holds
OSConstants b2 is OSSubset of b3
proof end;

theorem Th19: :: OSALG_2:19
for b1 being all-with_const_op OrderSortedSign
for b2 being non-empty OSAlgebra of b1
for b3 being non-empty OSSubAlgebra of b2 holds
OSConstants b2 is V5 OSSubset of b3
proof end;

theorem Th20: :: OSALG_2:20
for b1 being set
for b2 being ManySortedSet of b1
for b3 being set holds
( b3 is ManySortedSubset of b2 iff b3 in product (bool b2) )
proof end;

definition
let c1 be non empty Poset;
let c2 be OrderSortedSet of c1;
func OSbool c2 -> set means :: OSALG_2:def 6
for b1 being set holds
( b1 in a3 iff b1 is OrderSortedSubset of a2 );
existence
ex b1 being set st
for b2 being set holds
( b2 in b1 iff b2 is OrderSortedSubset of c2 )
proof end;
uniqueness
for b1, b2 being set holds
( ( for b3 being set holds
( b3 in b1 iff b3 is OrderSortedSubset of c2 ) ) & ( for b3 being set holds
( b3 in b2 iff b3 is OrderSortedSubset of c2 ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def6 defines OSbool OSALG_2:def 6 :
for b1 being non empty Poset
for b2 being OrderSortedSet of b1
for b3 being set holds
( b3 = OSbool b2 iff for b4 being set holds
( b4 in b3 iff b4 is OrderSortedSubset of b2 ) );

definition
let c1 be OrderSortedSign;
let c2 be OSAlgebra of c1;
let c3 be OSSubset of c2;
func OSSubSort c3 -> set equals :: OSALG_2:def 7
{ b1 where B is Element of SubSort a3 : b1 is OrderSortedSet of a1 } ;
correctness
coherence
{ b1 where B is Element of SubSort c3 : b1 is OrderSortedSet of c1 } is set
;
;
end;

:: deftheorem Def7 defines OSSubSort OSALG_2:def 7 :
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being OSSubset of b2 holds OSSubSort b3 = { b4 where B is Element of SubSort b3 : b4 is OrderSortedSet of b1 } ;

theorem Th21: :: OSALG_2:21
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being OSSubset of b2 holds OSSubSort b3 c= SubSort b3
proof end;

theorem Th22: :: OSALG_2:22
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being OSSubset of b2 holds the Sorts of b2 in OSSubSort b3
proof end;

registration
let c1 be OrderSortedSign;
let c2 be OSAlgebra of c1;
let c3 be OSSubset of c2;
cluster OSSubSort a3 -> non empty ;
coherence
not OSSubSort c3 is empty
by Th22;
end;

definition
let c1 be OrderSortedSign;
let c2 be OSAlgebra of c1;
func OSSubSort c2 -> set equals :: OSALG_2:def 8
{ b1 where B is Element of SubSort a2 : b1 is OrderSortedSet of a1 } ;
correctness
coherence
{ b1 where B is Element of SubSort c2 : b1 is OrderSortedSet of c1 } is set
;
;
end;

:: deftheorem Def8 defines OSSubSort OSALG_2:def 8 :
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1 holds OSSubSort b2 = { b3 where B is Element of SubSort b2 : b3 is OrderSortedSet of b1 } ;

theorem Th23: :: OSALG_2:23
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being OSSubset of b2 holds OSSubSort b3 c= OSSubSort b2
proof end;

registration
let c1 be OrderSortedSign;
let c2 be OSAlgebra of c1;
cluster OSSubSort a2 -> non empty ;
coherence
not OSSubSort c2 is empty
proof end;
end;

definition
let c1 be OrderSortedSign;
let c2 be OSAlgebra of c1;
let c3 be Element of OSSubSort c2;
func @ c3 -> OSSubset of a2 equals :: OSALG_2:def 9
a3;
coherence
c3 is OSSubset of c2
proof end;
end;

:: deftheorem Def9 defines @ OSALG_2:def 9 :
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being Element of OSSubSort b2 holds @ b3 = b3;

theorem Th24: :: OSALG_2:24
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3, b4 being OSSubset of b2 holds
( b4 in OSSubSort b3 iff ( b4 is opers_closed & OSConstants b2 c= b4 & b3 c= b4 ) )
proof end;

theorem Th25: :: OSALG_2:25
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being OSSubset of b2 holds
( b3 in OSSubSort b2 iff b3 is opers_closed )
proof end;

definition
let c1 be OrderSortedSign;
let c2 be OSAlgebra of c1;
let c3 be OSSubset of c2;
let c4 be Element of c1;
func OSSubSort c3,c4 -> set means :Def10: :: OSALG_2:def 10
for b1 being set holds
( b1 in a5 iff ex b2 being OSSubset of a2 st
( b2 in OSSubSort a3 & b1 = b2 . a4 ) );
existence
ex b1 being set st
for b2 being set holds
( b2 in b1 iff ex b3 being OSSubset of c2 st
( b3 in OSSubSort c3 & b2 = b3 . c4 ) )
proof end;
uniqueness
for b1, b2 being set holds
( ( for b3 being set holds
( b3 in b1 iff ex b4 being OSSubset of c2 st
( b4 in OSSubSort c3 & b3 = b4 . c4 ) ) ) & ( for b3 being set holds
( b3 in b2 iff ex b4 being OSSubset of c2 st
( b4 in OSSubSort c3 & b3 = b4 . c4 ) ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def10 defines OSSubSort OSALG_2:def 10 :
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being OSSubset of b2
for b4 being Element of b1
for b5 being set holds
( b5 = OSSubSort b3,b4 iff for b6 being set holds
( b6 in b5 iff ex b7 being OSSubset of b2 st
( b7 in OSSubSort b3 & b6 = b7 . b4 ) ) );

theorem Th26: :: OSALG_2:26
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being OSSubset of b2
for b4, b5 being SortSymbol of b1 holds
( b4 <= b5 implies OSSubSort b3,b5 is_coarser_than OSSubSort b3,b4 )
proof end;

theorem Th27: :: OSALG_2:27
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being OSSubset of b2
for b4 being SortSymbol of b1 holds OSSubSort b3,b4 c= SubSort b3,b4
proof end;

theorem Th28: :: OSALG_2:28
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being OSSubset of b2
for b4 being SortSymbol of b1 holds the Sorts of b2 . b4 in OSSubSort b3,b4
proof end;

registration
let c1 be OrderSortedSign;
let c2 be OSAlgebra of c1;
let c3 be OSSubset of c2;
let c4 be SortSymbol of c1;
cluster OSSubSort a3,a4 -> non empty ;
coherence
not OSSubSort c3,c4 is empty
by Th28;
end;

definition
let c1 be OrderSortedSign;
let c2 be OSAlgebra of c1;
let c3 be OSSubset of c2;
func OSMSubSort c3 -> OSSubset of a2 means :Def11: :: OSALG_2:def 11
for b1 being SortSymbol of a1 holds a4 . b1 = meet (OSSubSort a3,b1);
existence
ex b1 being OSSubset of c2 st
for b2 being SortSymbol of c1 holds b1 . b2 = meet (OSSubSort c3,b2)
proof end;
uniqueness
for b1, b2 being OSSubset of c2 holds
( ( for b3 being SortSymbol of c1 holds b1 . b3 = meet (OSSubSort c3,b3) ) & ( for b3 being SortSymbol of c1 holds b2 . b3 = meet (OSSubSort c3,b3) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def11 defines OSMSubSort OSALG_2:def 11 :
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3, b4 being OSSubset of b2 holds
( b4 = OSMSubSort b3 iff for b5 being SortSymbol of b1 holds b4 . b5 = meet (OSSubSort b3,b5) );

registration
let c1 be OrderSortedSign;
let c2 be OSAlgebra of c1;
cluster opers_closed OSSubset of a2;
existence
ex b1 being OSSubset of c2 st b1 is opers_closed
proof end;
end;

theorem Th29: :: OSALG_2:29
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being OSSubset of b2 holds (OSConstants b2) \/ b3 c= OSMSubSort b3
proof end;

theorem Th30: :: OSALG_2:30
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being OSSubset of b2 holds
( (OSConstants b2) \/ b3 is non-empty implies OSMSubSort b3 is non-empty )
proof end;

theorem Th31: :: OSALG_2:31
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being OperSymbol of b1
for b4, b5 being OSSubset of b2 holds
( b5 in OSSubSort b4 implies (((OSMSubSort b4) # ) * the Arity of b1) . b3 c= ((b5 # ) * the Arity of b1) . b3 )
proof end;

theorem Th32: :: OSALG_2:32
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being OperSymbol of b1
for b4, b5 being OSSubset of b2 holds
( b5 in OSSubSort b4 implies rng ((Den b3,b2) | ((((OSMSubSort b4) # ) * the Arity of b1) . b3)) c= (b5 * the ResultSort of b1) . b3 )
proof end;

theorem Th33: :: OSALG_2:33
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being OperSymbol of b1
for b4 being OSSubset of b2 holds rng ((Den b3,b2) | ((((OSMSubSort b4) # ) * the Arity of b1) . b3)) c= ((OSMSubSort b4) * the ResultSort of b1) . b3
proof end;

theorem Th34: :: OSALG_2:34
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being OSSubset of b2 holds
( OSMSubSort b3 is opers_closed & b3 c= OSMSubSort b3 )
proof end;

registration
let c1 be OrderSortedSign;
let c2 be OSAlgebra of c1;
let c3 be OSSubset of c2;
cluster OSMSubSort a3 -> opers_closed ;
coherence
OSMSubSort c3 is opers_closed
by Th34;
end;

registration
let c1 be OrderSortedSign;
let c2 be OSAlgebra of c1;
let c3 be opers_closed OSSubset of c2;
cluster a2 | a3 -> order-sorted ;
coherence
c2 | c3 is order-sorted
proof end;
end;

registration
let c1 be OrderSortedSign;
let c2 be OSAlgebra of c1;
let c3, c4 be OSSubAlgebra of c2;
cluster a3 /\ a4 -> order-sorted ;
coherence
c3 /\ c4 is order-sorted
proof end;
end;

definition
let c1 be OrderSortedSign;
let c2 be OSAlgebra of c1;
let c3 be OSSubset of c2;
canceled;
func GenOSAlg c3 -> strict OSSubAlgebra of a2 means :Def13: :: OSALG_2:def 13
( a3 is OSSubset of a4 & ( for b1 being OSSubAlgebra of a2 holds
( a3 is OSSubset of b1 implies a4 is OSSubAlgebra of b1 ) ) );
existence
ex b1 being strict OSSubAlgebra of c2 st
( c3 is OSSubset of b1 & ( for b2 being OSSubAlgebra of c2 holds
( c3 is OSSubset of b2 implies b1 is OSSubAlgebra of b2 ) ) )
proof end;
uniqueness
for b1, b2 being strict OSSubAlgebra of c2 holds
( c3 is OSSubset of b1 & ( for b3 being OSSubAlgebra of c2 holds
( c3 is OSSubset of b3 implies b1 is OSSubAlgebra of b3 ) ) & c3 is OSSubset of b2 & ( for b3 being OSSubAlgebra of c2 holds
( c3 is OSSubset of b3 implies b2 is OSSubAlgebra of b3 ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def12 OSALG_2:def 12 :
canceled;

:: deftheorem Def13 defines GenOSAlg OSALG_2:def 13 :
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being OSSubset of b2
for b4 being strict OSSubAlgebra of b2 holds
( b4 = GenOSAlg b3 iff ( b3 is OSSubset of b4 & ( for b5 being OSSubAlgebra of b2 holds
( b3 is OSSubset of b5 implies b4 is OSSubAlgebra of b5 ) ) ) );

theorem Th35: :: OSALG_2:35
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being OSSubset of b2 holds
( GenOSAlg b3 = b2 | (OSMSubSort b3) & the Sorts of (GenOSAlg b3) = OSMSubSort b3 )
proof end;

theorem Th36: :: OSALG_2:36
for b1 being non empty non void ManySortedSign
for b2 being MSAlgebra of b1
for b3 being MSSubset of b2 holds
( GenMSAlg b3 = b2 | (MSSubSort b3) & the Sorts of (GenMSAlg b3) = MSSubSort b3 )
proof end;

theorem Th37: :: OSALG_2:37
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being OSSubset of b2 holds the Sorts of (GenMSAlg b3) c= the Sorts of (GenOSAlg b3)
proof end;

theorem Th38: :: OSALG_2:38
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being OSSubset of b2 holds
GenMSAlg b3 is MSSubAlgebra of GenOSAlg b3
proof end;

theorem Th39: :: OSALG_2:39
for b1 being OrderSortedSign
for b2 being strict OSAlgebra of b1
for b3 being OSSubset of b2 holds
( b3 = the Sorts of b2 implies GenOSAlg b3 = b2 )
proof end;

theorem Th40: :: OSALG_2:40
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being strict OSSubAlgebra of b2
for b4 being OSSubset of b2 holds
( b4 = the Sorts of b3 implies GenOSAlg b4 = b3 )
proof end;

theorem Th41: :: OSALG_2:41
for b1 being OrderSortedSign
for b2 being non-empty OSAlgebra of b1
for b3 being OSSubAlgebra of b2 holds (GenOSAlg (OSConstants b2)) /\ b3 = GenOSAlg (OSConstants b2)
proof end;

definition
let c1 be OrderSortedSign;
let c2 be non-empty OSAlgebra of c1;
let c3, c4 be OSSubAlgebra of c2;
func c3 "\/"_os c4 -> strict OSSubAlgebra of a2 means :Def14: :: OSALG_2:def 14
for b1 being OSSubset of a2 holds
( b1 = the Sorts of a3 \/ the Sorts of a4 implies a5 = GenOSAlg b1 );
existence
ex b1 being strict OSSubAlgebra of c2 st
for b2 being OSSubset of c2 holds
( b2 = the Sorts of c3 \/ the Sorts of c4 implies b1 = GenOSAlg b2 )
proof end;
uniqueness
for b1, b2 being strict OSSubAlgebra of c2 holds
( ( for b3 being OSSubset of c2 holds
( b3 = the Sorts of c3 \/ the Sorts of c4 implies b1 = GenOSAlg b3 ) ) & ( for b3 being OSSubset of c2 holds
( b3 = the Sorts of c3 \/ the Sorts of c4 implies b2 = GenOSAlg b3 ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def14 defines "\/"_os OSALG_2:def 14 :
for b1 being OrderSortedSign
for b2 being non-empty OSAlgebra of b1
for b3, b4 being OSSubAlgebra of b2
for b5 being strict OSSubAlgebra of b2 holds
( b5 = b3 "\/"_os b4 iff for b6 being OSSubset of b2 holds
( b6 = the Sorts of b3 \/ the Sorts of b4 implies b5 = GenOSAlg b6 ) );

theorem Th42: :: OSALG_2:42
for b1 being OrderSortedSign
for b2 being non-empty OSAlgebra of b1
for b3 being OSSubAlgebra of b2
for b4, b5 being OSSubset of b2 holds
( b5 = b4 \/ the Sorts of b3 implies (GenOSAlg b4) "\/"_os b3 = GenOSAlg b5 )
proof end;

theorem Th43: :: OSALG_2:43
for b1 being OrderSortedSign
for b2 being non-empty OSAlgebra of b1
for b3 being OSSubAlgebra of b2
for b4 being OSSubset of b2 holds
( b4 = the Sorts of b2 implies (GenOSAlg b4) "\/"_os b3 = GenOSAlg b4 )
proof end;

theorem Th44: :: OSALG_2:44
for b1 being OrderSortedSign
for b2 being non-empty OSAlgebra of b1
for b3, b4 being OSSubAlgebra of b2 holds b3 "\/"_os b4 = b4 "\/"_os b3
proof end;

theorem Th45: :: OSALG_2:45
for b1 being OrderSortedSign
for b2 being non-empty OSAlgebra of b1
for b3, b4 being strict OSSubAlgebra of b2 holds b3 /\ (b3 "\/"_os b4) = b3
proof end;

theorem Th46: :: OSALG_2:46
for b1 being OrderSortedSign
for b2 being non-empty OSAlgebra of b1
for b3, b4 being strict OSSubAlgebra of b2 holds (b3 /\ b4) "\/"_os b4 = b4
proof end;

definition
let c1 be OrderSortedSign;
let c2 be OSAlgebra of c1;
func OSSub c2 -> set means :Def15: :: OSALG_2:def 15
for b1 being set holds
( b1 in a3 iff b1 is strict OSSubAlgebra of a2 );
existence
ex b1 being set st
for b2 being set holds
( b2 in b1 iff b2 is strict OSSubAlgebra of c2 )
proof end;
uniqueness
for b1, b2 being set holds
( ( for b3 being set holds
( b3 in b1 iff b3 is strict OSSubAlgebra of c2 ) ) & ( for b3 being set holds
( b3 in b2 iff b3 is strict OSSubAlgebra of c2 ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def15 defines OSSub OSALG_2:def 15 :
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1
for b3 being set holds
( b3 = OSSub b2 iff for b4 being set holds
( b4 in b3 iff b4 is strict OSSubAlgebra of b2 ) );

theorem Th47: :: OSALG_2:47
for b1 being OrderSortedSign
for b2 being OSAlgebra of b1 holds OSSub b2 c= MSSub b2
proof end;

registration
let c1 be OrderSortedSign;
let c2 be OSAlgebra of c1;
cluster OSSub a2 -> non empty ;
coherence
not OSSub c2 is empty
proof end;
end;

definition
let c1 be OrderSortedSign;
let c2 be OSAlgebra of c1;
redefine func OSSub as OSSub c2 -> Subset of (MSSub a2);
coherence
OSSub c2 is Subset of (MSSub c2)
by Th47;
end;

definition
let c1 be OrderSortedSign;
let c2 be non-empty OSAlgebra of c1;
func OSAlg_join c2 -> BinOp of OSSub a2 means :Def16: :: OSALG_2:def 16
for b1, b2 being Element of OSSub a2
for b3, b4 being strict OSSubAlgebra of a2 holds
( b1 = b3 & b2 = b4 implies a3 . b1,b2 = b3 "\/"_os b4 );
existence
ex b1 being BinOp of OSSub c2 st
for b2, b3 being Element of OSSub c2
for b4, b5 being strict OSSubAlgebra of c2 holds
( b2 = b4 & b3 = b5 implies b1 . b2,b3 = b4 "\/"_os b5 )
proof end;
uniqueness
for b1, b2 being BinOp of OSSub c2 holds
( ( for b3, b4 being Element of OSSub c2
for b5, b6 being strict OSSubAlgebra of c2 holds
( b3 = b5 & b4 = b6 implies b1 . b3,b4 = b5 "\/"_os b6 ) ) & ( for b3, b4 being Element of OSSub c2
for b5, b6 being strict OSSubAlgebra of c2 holds
( b3 = b5 & b4 = b6 implies b2 . b3,b4 = b5 "\/"_os b6 ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def16 defines OSAlg_join OSALG_2:def 16 :
for b1 being OrderSortedSign
for b2 being non-empty OSAlgebra of b1
for b3 being BinOp of OSSub b2 holds
( b3 = OSAlg_join b2 iff for b4, b5 being Element of OSSub b2
for b6, b7 being strict OSSubAlgebra of b2 holds
( b4 = b6 & b5 = b7 implies b3 . b4,b5 = b6 "\/"_os b7 ) );

definition
let c1 be OrderSortedSign;
let c2 be non-empty OSAlgebra of c1;
func OSAlg_meet c2 -> BinOp of OSSub a2 means :Def17: :: OSALG_2:def 17
for b1, b2 being Element of OSSub a2
for b3, b4 being strict OSSubAlgebra of a2 holds
( b1 = b3 & b2 = b4 implies a3 . b1,b2 = b3 /\ b4 );
existence
ex b1 being BinOp of OSSub c2 st
for b2, b3 being Element of OSSub c2
for b4, b5 being strict OSSubAlgebra of c2 holds
( b2 = b4 & b3 = b5 implies b1 . b2,b3 = b4 /\ b5 )
proof end;
uniqueness
for b1, b2 being BinOp of OSSub c2 holds
( ( for b3, b4 being Element of OSSub c2
for b5, b6 being strict OSSubAlgebra of c2 holds
( b3 = b5 & b4 = b6 implies b1 . b3,b4 = b5 /\ b6 ) ) & ( for b3, b4 being Element of OSSub c2
for b5, b6 being strict OSSubAlgebra of c2 holds
( b3 = b5 & b4 = b6 implies b2 . b3,b4 = b5 /\ b6 ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def17 defines OSAlg_meet OSALG_2:def 17 :
for b1 being OrderSortedSign
for b2 being non-empty OSAlgebra of b1
for b3 being BinOp of OSSub b2 holds
( b3 = OSAlg_meet b2 iff for b4, b5 being Element of OSSub b2
for b6, b7 being strict OSSubAlgebra of b2 holds
( b4 = b6 & b5 = b7 implies b3 . b4,b5 = b6 /\ b7 ) );

theorem Th48: :: OSALG_2:48
for b1 being OrderSortedSign
for b2 being non-empty OSAlgebra of b1
for b3, b4 being Element of OSSub b2 holds (OSAlg_meet b2) . b3,b4 = (MSAlg_meet b2) . b3,b4
proof end;

theorem Th49: :: OSALG_2:49
for b1 being OrderSortedSign
for b2 being non-empty OSAlgebra of b1 holds OSAlg_join b2 is commutative
proof end;

theorem Th50: :: OSALG_2:50
for b1 being OrderSortedSign
for b2 being non-empty OSAlgebra of b1 holds OSAlg_join b2 is associative
proof end;

theorem Th51: :: OSALG_2:51
for b1 being OrderSortedSign
for b2 being non-empty OSAlgebra of b1 holds OSAlg_meet b2 is commutative
proof end;

theorem Th52: :: OSALG_2:52
for b1 being OrderSortedSign
for b2 being non-empty OSAlgebra of b1 holds OSAlg_meet b2 is associative
proof end;

definition
let c1 be OrderSortedSign;
let c2 be non-empty OSAlgebra of c1;
func OSSubAlLattice c2 -> strict Lattice equals :: OSALG_2:def 18
LattStr(# (OSSub a2),(OSAlg_join a2),(OSAlg_meet a2) #);
coherence
LattStr(# (OSSub c2),(OSAlg_join c2),(OSAlg_meet c2) #) is strict Lattice
proof end;
end;

:: deftheorem Def18 defines OSSubAlLattice OSALG_2:def 18 :
for b1 being OrderSortedSign
for b2 being non-empty OSAlgebra of b1 holds OSSubAlLattice b2 = LattStr(# (OSSub b2),(OSAlg_join b2),(OSAlg_meet b2) #);

theorem Th53: :: OSALG_2:53
for b1 being OrderSortedSign
for b2 being non-empty OSAlgebra of b1 holds OSSubAlLattice b2 is bounded
proof end;

registration
let c1 be OrderSortedSign;
let c2 be non-empty OSAlgebra of c1;
cluster OSSubAlLattice a2 -> strict bounded ;
coherence
OSSubAlLattice c2 is bounded
by Th53;
end;

theorem Th54: :: OSALG_2:54
for b1 being OrderSortedSign
for b2 being non-empty OSAlgebra of b1 holds Bottom (OSSubAlLattice b2) = GenOSAlg (OSConstants b2)
proof end;

theorem Th55: :: OSALG_2:55
for b1 being OrderSortedSign
for b2 being non-empty OSAlgebra of b1
for b3 being OSSubset of b2 holds
( b3 = the Sorts of b2 implies Top (OSSubAlLattice b2) = GenOSAlg b3 )
proof end;

theorem Th56: :: OSALG_2:56
for b1 being OrderSortedSign
for b2 being strict non-empty OSAlgebra of b1 holds Top (OSSubAlLattice b2) = b2
proof end;