definition
let S be
OrderSortedSign;
let X be
ManySortedSet of
S;
func OSREL X -> Relation of
([: the carrier' of S,{ the carrier of S}:] \/ (Union (coprod X))),
(([: the carrier' of S,{ the carrier of S}:] \/ (Union (coprod X))) *) means :
Def4:
for
a being
Element of
[: the carrier' of S,{ the carrier of S}:] \/ (Union (coprod X)) for
b being
Element of
([: the carrier' of S,{ the carrier of S}:] \/ (Union (coprod X))) * holds
(
[a,b] in it iff (
a in [: the carrier' of S,{ the carrier of S}:] & ( for
o being
OperSymbol of
S st
[o, the carrier of S] = a holds
(
len b = len (the_arity_of o) & ( for
x being
set st
x in dom b holds
( (
b . x in [: the carrier' of S,{ the carrier of S}:] implies for
o1 being
OperSymbol of
S st
[o1, the carrier of S] = b . x holds
the_result_sort_of o1 <= (the_arity_of o) /. x ) & (
b . x in Union (coprod X) implies ex
i being
Element of
S st
(
i <= (the_arity_of o) /. x &
b . x in coprod (
i,
X) ) ) ) ) ) ) ) );
existence
ex b1 being Relation of ([: the carrier' of S,{ the carrier of S}:] \/ (Union (coprod X))),(([: the carrier' of S,{ the carrier of S}:] \/ (Union (coprod X))) *) st
for a being Element of [: the carrier' of S,{ the carrier of S}:] \/ (Union (coprod X))
for b being Element of ([: the carrier' of S,{ the carrier of S}:] \/ (Union (coprod X))) * holds
( [a,b] in b1 iff ( a in [: the carrier' of S,{ the carrier of S}:] & ( for o being OperSymbol of S st [o, the carrier of S] = a holds
( len b = len (the_arity_of o) & ( for x being set st x in dom b holds
( ( b . x in [: the carrier' of S,{ the carrier of S}:] implies for o1 being OperSymbol of S st [o1, the carrier of S] = b . x holds
the_result_sort_of o1 <= (the_arity_of o) /. x ) & ( b . x in Union (coprod X) implies ex i being Element of S st
( i <= (the_arity_of o) /. x & b . x in coprod (i,X) ) ) ) ) ) ) ) )
uniqueness
for b1, b2 being Relation of ([: the carrier' of S,{ the carrier of S}:] \/ (Union (coprod X))),(([: the carrier' of S,{ the carrier of S}:] \/ (Union (coprod X))) *) st ( for a being Element of [: the carrier' of S,{ the carrier of S}:] \/ (Union (coprod X))
for b being Element of ([: the carrier' of S,{ the carrier of S}:] \/ (Union (coprod X))) * holds
( [a,b] in b1 iff ( a in [: the carrier' of S,{ the carrier of S}:] & ( for o being OperSymbol of S st [o, the carrier of S] = a holds
( len b = len (the_arity_of o) & ( for x being set st x in dom b holds
( ( b . x in [: the carrier' of S,{ the carrier of S}:] implies for o1 being OperSymbol of S st [o1, the carrier of S] = b . x holds
the_result_sort_of o1 <= (the_arity_of o) /. x ) & ( b . x in Union (coprod X) implies ex i being Element of S st
( i <= (the_arity_of o) /. x & b . x in coprod (i,X) ) ) ) ) ) ) ) ) ) & ( for a being Element of [: the carrier' of S,{ the carrier of S}:] \/ (Union (coprod X))
for b being Element of ([: the carrier' of S,{ the carrier of S}:] \/ (Union (coprod X))) * holds
( [a,b] in b2 iff ( a in [: the carrier' of S,{ the carrier of S}:] & ( for o being OperSymbol of S st [o, the carrier of S] = a holds
( len b = len (the_arity_of o) & ( for x being set st x in dom b holds
( ( b . x in [: the carrier' of S,{ the carrier of S}:] implies for o1 being OperSymbol of S st [o1, the carrier of S] = b . x holds
the_result_sort_of o1 <= (the_arity_of o) /. x ) & ( b . x in Union (coprod X) implies ex i being Element of S st
( i <= (the_arity_of o) /. x & b . x in coprod (i,X) ) ) ) ) ) ) ) ) ) holds
b1 = b2
end;
::
deftheorem Def4 defines
OSREL OSAFREE:def 4 :
for S being OrderSortedSign
for X being ManySortedSet of S
for b3 being Relation of ([: the carrier' of S,{ the carrier of S}:] \/ (Union (coprod X))),(([: the carrier' of S,{ the carrier of S}:] \/ (Union (coprod X))) *) holds
( b3 = OSREL X iff for a being Element of [: the carrier' of S,{ the carrier of S}:] \/ (Union (coprod X))
for b being Element of ([: the carrier' of S,{ the carrier of S}:] \/ (Union (coprod X))) * holds
( [a,b] in b3 iff ( a in [: the carrier' of S,{ the carrier of S}:] & ( for o being OperSymbol of S st [o, the carrier of S] = a holds
( len b = len (the_arity_of o) & ( for x being set st x in dom b holds
( ( b . x in [: the carrier' of S,{ the carrier of S}:] implies for o1 being OperSymbol of S st [o1, the carrier of S] = b . x holds
the_result_sort_of o1 <= (the_arity_of o) /. x ) & ( b . x in Union (coprod X) implies ex i being Element of S st
( i <= (the_arity_of o) /. x & b . x in coprod (i,X) ) ) ) ) ) ) ) ) );
definition
let S be
locally_directed OrderSortedSign;
let X be
V2()
ManySortedSet of
S;
existence
ex b1 being MSEquivalence-like OrderSortedRelation of ParsedTermsOSA X st
for i being set st i in the carrier of S holds
b1 . i = { [x,y] where x, y is Element of TS (DTConOSA X) : [x,i] in (PTClasses X) . y }
uniqueness
for b1, b2 being MSEquivalence-like OrderSortedRelation of ParsedTermsOSA X st ( for i being set st i in the carrier of S holds
b1 . i = { [x,y] where x, y is Element of TS (DTConOSA X) : [x,i] in (PTClasses X) . y } ) & ( for i being set st i in the carrier of S holds
b2 . i = { [x,y] where x, y is Element of TS (DTConOSA X) : [x,i] in (PTClasses X) . y } ) holds
b1 = b2
end;
theorem Th26:
for
S being
locally_directed OrderSortedSign for
X being
V2()
ManySortedSet of
S for
R being
ManySortedRelation of
(ParsedTermsOSA X) holds
(
R = PTCongruence X iff ( ( for
s1,
s2 being
Element of
S for
x being
object st
x in X . s1 holds
( (
s1 <= s2 implies
[(root-tree [x,s1]),(root-tree [x,s1])] in R . s2 ) & ( for
y being
object st (
[(root-tree [x,s1]),y] in R . s2 or
[y,(root-tree [x,s1])] in R . s2 ) holds
(
s1 <= s2 &
y = root-tree [x,s1] ) ) ) ) & ( for
o1,
o2 being
OperSymbol of
S for
x1 being
Element of
Args (
o1,
(ParsedTermsOSA X))
for
x2 being
Element of
Args (
o2,
(ParsedTermsOSA X))
for
s3 being
Element of
S holds
(
[((Den (o1,(ParsedTermsOSA X))) . x1),((Den (o2,(ParsedTermsOSA X))) . x2)] in R . s3 iff (
o1 ~= o2 &
len (the_arity_of o1) = len (the_arity_of o2) &
the_result_sort_of o1 <= s3 &
the_result_sort_of o2 <= s3 & ex
w3 being
Element of the
carrier of
S * st
(
dom w3 = dom x1 & ( for
y being
Nat st
y in dom w3 holds
[(x1 . y),(x2 . y)] in R . (w3 /. y) ) ) ) ) ) ) )
definition
let S be
locally_directed OrderSortedSign;
let X be
V2()
ManySortedSet of
S;
let s be
Element of
S;
existence
ex b1 being Function of (OSFreeGen (s,X)),(PTVars (s,X)) st
for t being Symbol of (DTConOSA X) st ((OSNat_Hom ((ParsedTermsOSA X),(LCongruence X))) . s) . (root-tree t) in OSFreeGen (s,X) holds
b1 . (((OSNat_Hom ((ParsedTermsOSA X),(LCongruence X))) . s) . (root-tree t)) = root-tree t
uniqueness
for b1, b2 being Function of (OSFreeGen (s,X)),(PTVars (s,X)) st ( for t being Symbol of (DTConOSA X) st ((OSNat_Hom ((ParsedTermsOSA X),(LCongruence X))) . s) . (root-tree t) in OSFreeGen (s,X) holds
b1 . (((OSNat_Hom ((ParsedTermsOSA X),(LCongruence X))) . s) . (root-tree t)) = root-tree t ) & ( for t being Symbol of (DTConOSA X) st ((OSNat_Hom ((ParsedTermsOSA X),(LCongruence X))) . s) . (root-tree t) in OSFreeGen (s,X) holds
b2 . (((OSNat_Hom ((ParsedTermsOSA X),(LCongruence X))) . s) . (root-tree t)) = root-tree t ) holds
b1 = b2
end;
:: changing to MSSubset and its OSCl, to make free algebras easier
:: no good way how to retypeOSCL into OSSubset now?