begin
Lm1:
for I being set
for f being ManySortedSet of I
for p being FinSequence of I holds
( dom (f * p) = dom p & len (f * p) = len p )
theorem Th1:
for
A,
O being non
empty set for
R being
Order of
A for
Ol being
Equivalence_Relation of
O for
f being
Function of
O,
(A *) for
g being
Function of
O,
A holds
( not
OverloadedRSSign(#
A,
R,
O,
Ol,
f,
g #) is
empty & not
OverloadedRSSign(#
A,
R,
O,
Ol,
f,
g #) is
void &
OverloadedRSSign(#
A,
R,
O,
Ol,
f,
g #) is
reflexive &
OverloadedRSSign(#
A,
R,
O,
Ol,
f,
g #) is
transitive &
OverloadedRSSign(#
A,
R,
O,
Ol,
f,
g #) is
antisymmetric )
registration
let A be non
empty set ;
let R be
Order of
A;
let O be non
empty set ;
let Ol be
Equivalence_Relation of
O;
let f be
Function of
O,
(A *);
let g be
Function of
O,
A;
cluster OverloadedRSSign(#
A,
R,
O,
Ol,
f,
g #)
-> non
empty reflexive transitive antisymmetric ;
coherence
( OverloadedRSSign(# A,R,O,Ol,f,g #) is strict & not OverloadedRSSign(# A,R,O,Ol,f,g #) is empty & OverloadedRSSign(# A,R,O,Ol,f,g #) is reflexive & OverloadedRSSign(# A,R,O,Ol,f,g #) is transitive & OverloadedRSSign(# A,R,O,Ol,f,g #) is antisymmetric )
by Th1;
end;
begin
:: deftheorem OSALG_1:def 1 :
canceled;
:: deftheorem Def2 defines order-sorted OSALG_1:def 2 :
for S being OverloadedRSSign holds
( S is order-sorted iff ( S is reflexive & S is transitive & S is antisymmetric ) );
:: deftheorem Def3 defines ~= OSALG_1:def 3 :
for S being non empty non void OverloadedMSSign
for x, y being OperSymbol of S holds
( x ~= y iff [x,y] in the Overloading of S );
theorem Th2:
:: deftheorem Def4 defines discernable OSALG_1:def 4 :
for S being non empty non void OverloadedMSSign holds
( S is discernable iff for x, y being OperSymbol of S st x ~= y & the_arity_of x = the_arity_of y & the_result_sort_of x = the_result_sort_of y holds
x = y );
:: deftheorem Def5 defines op-discrete OSALG_1:def 5 :
for S being non empty non void OverloadedMSSign holds
( S is op-discrete iff the Overloading of S = id the carrier' of S );
theorem Th3:
theorem Th4:
begin
:: deftheorem Def6 defines OSSign OSALG_1:def 6 :
for S0 being non empty non void ManySortedSign
for b2 being non empty non void strict order-sorted OverloadedRSSign holds
( b2 = OSSign S0 iff ( the carrier of S0 = the carrier of b2 & id the carrier of S0 = the InternalRel of b2 & the carrier' of S0 = the carrier' of b2 & id the carrier' of S0 = the Overloading of b2 & the Arity of S0 = the Arity of b2 & the ResultSort of S0 = the ResultSort of b2 ) );
theorem Th5:
:: deftheorem Def7 defines <= OSALG_1:def 7 :
for S being non empty Poset
for w1, w2 being Element of the carrier of S * holds
( w1 <= w2 iff ( len w1 = len w2 & ( for i being set st i in dom w1 holds
for s1, s2 being Element of S st s1 = w1 . i & s2 = w2 . i holds
s1 <= s2 ) ) );
theorem Th6:
theorem Th7:
theorem Th8:
:: deftheorem Def8 defines monotone OSALG_1:def 8 :
for S being OrderSortedSign
for o being OperSymbol of S holds
( o is monotone iff for o2 being OperSymbol of S st o ~= o2 & the_arity_of o <= the_arity_of o2 holds
the_result_sort_of o <= the_result_sort_of o2 );
:: deftheorem Def9 defines monotone OSALG_1:def 9 :
for S being OrderSortedSign holds
( S is monotone iff for o being OperSymbol of S holds o is monotone );
theorem Th9:
theorem
:: deftheorem Def10 defines has_least_args_for OSALG_1:def 10 :
for S being OrderSortedSign
for o, o1 being OperSymbol of S
for w1 being Element of the carrier of S * holds
( o1 has_least_args_for o,w1 iff ( o ~= o1 & w1 <= the_arity_of o1 & ( for o2 being OperSymbol of S st o ~= o2 & w1 <= the_arity_of o2 holds
the_arity_of o1 <= the_arity_of o2 ) ) );
:: deftheorem Def11 defines has_least_sort_for OSALG_1:def 11 :
for S being OrderSortedSign
for o, o1 being OperSymbol of S
for w1 being Element of the carrier of S * holds
( o1 has_least_sort_for o,w1 iff ( o ~= o1 & w1 <= the_arity_of o1 & ( for o2 being OperSymbol of S st o ~= o2 & w1 <= the_arity_of o2 holds
the_result_sort_of o1 <= the_result_sort_of o2 ) ) );
:: deftheorem Def12 defines has_least_rank_for OSALG_1:def 12 :
for S being OrderSortedSign
for o, o1 being OperSymbol of S
for w1 being Element of the carrier of S * holds
( o1 has_least_rank_for o,w1 iff ( o1 has_least_args_for o,w1 & o1 has_least_sort_for o,w1 ) );
:: deftheorem Def13 defines regular OSALG_1:def 13 :
for S being OrderSortedSign
for o being OperSymbol of S holds
( o is regular iff ( o is monotone & ( for w1 being Element of the carrier of S * st w1 <= the_arity_of o holds
ex o1 being OperSymbol of S st o1 has_least_args_for o,w1 ) ) );
:: deftheorem Def14 defines regular OSALG_1:def 14 :
for SM being monotone OrderSortedSign holds
( SM is regular iff for om being OperSymbol of SM holds om is regular );
theorem Th11:
theorem Th12:
theorem Th13:
:: deftheorem Def15 defines LBound OSALG_1:def 15 :
for SR being monotone regular OrderSortedSign
for o being OperSymbol of SR
for w1 being Element of the carrier of SR * st w1 <= the_arity_of o holds
for b4 being OperSymbol of SR holds
( b4 = LBound (o,w1) iff b4 has_least_args_for o,w1 );
theorem Th14:
:: deftheorem defines ConstOSSet OSALG_1:def 16 :
for R being non empty Poset
for z being non empty set holds ConstOSSet (R,z) = the carrier of R --> z;
theorem Th15:
:: deftheorem OSALG_1:def 17 :
canceled;
:: deftheorem Def18 defines order-sorted OSALG_1:def 18 :
for R being non empty Poset
for M being ManySortedSet of R holds
( M is order-sorted iff for s1, s2 being Element of R st s1 <= s2 holds
M . s1 c= M . s2 );
theorem Th16:
:: deftheorem Def19 defines order-sorted OSALG_1:def 19 :
for S being OrderSortedSign
for M being MSAlgebra of S holds
( M is order-sorted iff for s1, s2 being SortSymbol of S st s1 <= s2 holds
the Sorts of M . s1 c= the Sorts of M . s2 );
theorem Th17:
definition
let S be
OrderSortedSign;
let z be non
empty set ;
let CH be
ManySortedFunction of
((ConstOSSet (S,z)) #) * the
Arity of
S,
(ConstOSSet (S,z)) * the
ResultSort of
S;
func ConstOSA (
S,
z,
CH)
-> strict non-empty MSAlgebra of
S means :
Def20:
( the
Sorts of
it = ConstOSSet (
S,
z) & the
Charact of
it = CH );
existence
ex b1 being strict non-empty MSAlgebra of S st
( the Sorts of b1 = ConstOSSet (S,z) & the Charact of b1 = CH )
uniqueness
for b1, b2 being strict non-empty MSAlgebra of S st the Sorts of b1 = ConstOSSet (S,z) & the Charact of b1 = CH & the Sorts of b2 = ConstOSSet (S,z) & the Charact of b2 = CH holds
b1 = b2
;
end;
:: deftheorem Def20 defines ConstOSA OSALG_1:def 20 :
for S being OrderSortedSign
for z being non empty set
for CH being ManySortedFunction of ((ConstOSSet (S,z)) #) * the Arity of S,(ConstOSSet (S,z)) * the ResultSort of S
for b4 being strict non-empty MSAlgebra of S holds
( b4 = ConstOSA (S,z,CH) iff ( the Sorts of b4 = ConstOSSet (S,z) & the Charact of b4 = CH ) );
theorem Th18:
theorem Th19:
theorem Th20:
:: deftheorem defines OSAlg OSALG_1:def 21 :
for S0 being non empty non void ManySortedSign
for M being MSAlgebra of S0
for b3 being strict OSAlgebra of OSSign S0 holds
( b3 = OSAlg M iff ( the Sorts of b3 = the Sorts of M & the Charact of b3 = the Charact of M ) );
theorem Th21:
:: deftheorem Def22 defines <= OSALG_1:def 22 :
for S being OrderSortedSign
for o1, o2 being OperSymbol of S holds
( o1 <= o2 iff ( o1 ~= o2 & the_arity_of o1 <= the_arity_of o2 & the_result_sort_of o1 <= the_result_sort_of o2 ) );
theorem
theorem
theorem Th24:
theorem Th25:
theorem
:: deftheorem Def23 defines monotone OSALG_1:def 23 :
for S being OrderSortedSign
for A being OSAlgebra of S holds
( A is monotone iff for o1, o2 being OperSymbol of S st o1 <= o2 holds
(Den (o2,A)) | (Args (o1,A)) = Den (o1,A) );
theorem Th27:
theorem
definition
let S be
OrderSortedSign;
let z be non
empty set ;
let z1 be
Element of
z;
func TrivialOSA (
S,
z,
z1)
-> strict OSAlgebra of
S means :
Def24:
( the
Sorts of
it = ConstOSSet (
S,
z) & ( for
o being
OperSymbol of
S holds
Den (
o,
it)
= (Args (o,it)) --> z1 ) );
existence
ex b1 being strict OSAlgebra of S st
( the Sorts of b1 = ConstOSSet (S,z) & ( for o being OperSymbol of S holds Den (o,b1) = (Args (o,b1)) --> z1 ) )
uniqueness
for b1, b2 being strict OSAlgebra of S st the Sorts of b1 = ConstOSSet (S,z) & ( for o being OperSymbol of S holds Den (o,b1) = (Args (o,b1)) --> z1 ) & the Sorts of b2 = ConstOSSet (S,z) & ( for o being OperSymbol of S holds Den (o,b2) = (Args (o,b2)) --> z1 ) holds
b1 = b2
end;
:: deftheorem Def24 defines TrivialOSA OSALG_1:def 24 :
for S being OrderSortedSign
for z being non empty set
for z1 being Element of z
for b4 being strict OSAlgebra of S holds
( b4 = TrivialOSA (S,z,z1) iff ( the Sorts of b4 = ConstOSSet (S,z) & ( for o being OperSymbol of S holds Den (o,b4) = (Args (o,b4)) --> z1 ) ) );
theorem Th29:
:: deftheorem defines OperNames OSALG_1:def 25 :
for S being OrderSortedSign holds OperNames S = Class the Overloading of S;
:: deftheorem defines Name OSALG_1:def 26 :
for S being OrderSortedSign
for op1 being OperSymbol of S holds Name op1 = Class ( the Overloading of S,op1);
theorem Th30:
theorem Th31:
theorem
theorem Th33:
theorem Th34:
:: deftheorem defines LBound OSALG_1:def 27 :
for SR being monotone regular OrderSortedSign
for on being OperName of SR
for w being Element of the carrier of SR * st ex op being Element of on st w <= the_arity_of op holds
for b4 being Element of on holds
( b4 = LBound (on,w) iff for op being Element of on st w <= the_arity_of op holds
b4 = LBound (op,w) );
theorem