:: MSUALG_6 semantic presentation

definition
let c1 be non empty ManySortedSign ;
let c2 be MSAlgebra of c1;
let c3 be SortSymbol of c1;
mode Element is Element of the Sorts of a2 . a3;
end;

definition
let c1 be set ;
let c2 be ManySortedSet of c1;
let c3, c4 be ManySortedFunction of c2,c2;
redefine func ** as c4 ** c3 -> ManySortedFunction of a2,a2;
coherence
c4 ** c3 is ManySortedFunction of c2,c2
proof end;
end;

theorem Th1: :: MSUALG_6:1
for b1 being non empty non void ManySortedSign
for b2 being MSAlgebra of b1
for b3 being OperSymbol of b1
for b4 being set holds
( b4 in Args b3,b2 implies b4 is Function )
proof end;

theorem Th2: :: MSUALG_6:2
for b1 being non empty non void ManySortedSign
for b2 being MSAlgebra of b1
for b3 being OperSymbol of b1
for b4 being Function holds
( b4 in Args b3,b2 implies ( dom b4 = dom (the_arity_of b3) & ( for b5 being set holds
( b5 in dom (the_arity_of b3) implies b4 . b5 in the Sorts of b2 . ((the_arity_of b3) /. b5) ) ) ) )
proof end;

definition
let c1 be non empty non void ManySortedSign ;
let c2 be MSAlgebra of c1;
attr a2 is feasible means :Def1: :: MSUALG_6:def 1
for b1 being OperSymbol of a1 holds
not ( Args b1,a2 <> {} & not Result b1,a2 <> {} );
end;

:: deftheorem Def1 defines feasible MSUALG_6:def 1 :
for b1 being non empty non void ManySortedSign
for b2 being MSAlgebra of b1 holds
( b2 is feasible iff for b3 being OperSymbol of b1 holds
not ( Args b3,b2 <> {} & not Result b3,b2 <> {} ) );

theorem Th3: :: MSUALG_6:3
for b1 being non empty non void ManySortedSign
for b2 being OperSymbol of b1
for b3 being MSAlgebra of b1 holds
( Args b2,b3 <> {} iff for b4 being Nat holds
not ( b4 in dom (the_arity_of b2) & not the Sorts of b3 . ((the_arity_of b2) /. b4) <> {} ) )
proof end;

registration
let c1 be non empty non void ManySortedSign ;
cluster non-empty -> feasible MSAlgebra of a1;
coherence
for b1 being MSAlgebra of c1 holds
( b1 is non-empty implies b1 is feasible )
proof end;
end;

registration
let c1 be non empty non void ManySortedSign ;
cluster non-empty feasible MSAlgebra of a1;
existence
ex b1 being MSAlgebra of c1 st b1 is non-empty
proof end;
end;

definition
let c1 be non empty non void ManySortedSign ;
let c2 be MSAlgebra of c1;
mode Endomorphism of c2 -> ManySortedFunction of a2,a2 means :Def2: :: MSUALG_6:def 2
a3 is_homomorphism a2,a2;
existence
ex b1 being ManySortedFunction of c2,c2 st b1 is_homomorphism c2,c2
proof end;
end;

:: deftheorem Def2 defines Endomorphism MSUALG_6:def 2 :
for b1 being non empty non void ManySortedSign
for b2 being MSAlgebra of b1
for b3 being ManySortedFunction of b2,b2 holds
( b3 is Endomorphism of b2 iff b3 is_homomorphism b2,b2 );

theorem Th4: :: MSUALG_6:4
for b1 being non empty non void ManySortedSign
for b2 being MSAlgebra of b1 holds
id the Sorts of b2 is Endomorphism of b2
proof end;

theorem Th5: :: MSUALG_6:5
for b1 being non empty non void ManySortedSign
for b2 being MSAlgebra of b1
for b3, b4 being ManySortedFunction of b2,b2
for b5 being OperSymbol of b1
for b6 being Element of Args b5,b2 holds
( b6 in Args b5,b2 implies b4 # (b3 # b6) = (b4 ** b3) # b6 )
proof end;

theorem Th6: :: MSUALG_6:6
for b1 being non empty non void ManySortedSign
for b2 being MSAlgebra of b1
for b3, b4 being Endomorphism of b2 holds
b4 ** b3 is Endomorphism of b2
proof end;

definition
let c1 be non empty non void ManySortedSign ;
let c2 be MSAlgebra of c1;
let c3, c4 be Endomorphism of c2;
redefine func ** as c4 ** c3 -> Endomorphism of a2;
coherence
c4 ** c3 is Endomorphism of c2
by Th6;
end;

definition
let c1 be non empty non void ManySortedSign ;
func TranslationRel c1 -> Relation of the carrier of a1 means :Def3: :: MSUALG_6:def 3
for b1, b2 being SortSymbol of a1 holds
( [b1,b2] in a2 iff ex b3 being OperSymbol of a1 st
( the_result_sort_of b3 = b2 & ex b4 being Nat st
( b4 in dom (the_arity_of b3) & (the_arity_of b3) /. b4 = b1 ) ) );
existence
ex b1 being Relation of the carrier of c1 st
for b2, b3 being SortSymbol of c1 holds
( [b2,b3] in b1 iff ex b4 being OperSymbol of c1 st
( the_result_sort_of b4 = b3 & ex b5 being Nat st
( b5 in dom (the_arity_of b4) & (the_arity_of b4) /. b5 = b2 ) ) )
proof end;
correctness
uniqueness
for b1, b2 being Relation of the carrier of c1 holds
( ( for b3, b4 being SortSymbol of c1 holds
( [b3,b4] in b1 iff ex b5 being OperSymbol of c1 st
( the_result_sort_of b5 = b4 & ex b6 being Nat st
( b6 in dom (the_arity_of b5) & (the_arity_of b5) /. b6 = b3 ) ) ) ) & ( for b3, b4 being SortSymbol of c1 holds
( [b3,b4] in b2 iff ex b5 being OperSymbol of c1 st
( the_result_sort_of b5 = b4 & ex b6 being Nat st
( b6 in dom (the_arity_of b5) & (the_arity_of b5) /. b6 = b3 ) ) ) ) implies b1 = b2 )
;
proof end;
end;

:: deftheorem Def3 defines TranslationRel MSUALG_6:def 3 :
for b1 being non empty non void ManySortedSign
for b2 being Relation of the carrier of b1 holds
( b2 = TranslationRel b1 iff for b3, b4 being SortSymbol of b1 holds
( [b3,b4] in b2 iff ex b5 being OperSymbol of b1 st
( the_result_sort_of b5 = b4 & ex b6 being Nat st
( b6 in dom (the_arity_of b5) & (the_arity_of b5) /. b6 = b3 ) ) ) );

theorem Th7: :: MSUALG_6:7
for b1 being non empty non void ManySortedSign
for b2 being OperSymbol of b1
for b3 being MSAlgebra of b1
for b4 being Function holds
( b4 in Args b2,b3 implies for b5 being Nat
for b6 being Element of b3,((the_arity_of b2) /. b5) holds b4 +* b5,b6 in Args b2,b3 )
proof end;

theorem Th8: :: MSUALG_6:8
for b1 being non empty non void ManySortedSign
for b2, b3 being MSAlgebra of b1
for b4 being ManySortedFunction of b2,b3
for b5 being OperSymbol of b1 holds
( Args b5,b2 <> {} & Args b5,b3 <> {} implies for b6 being Nat holds
( b6 in dom (the_arity_of b5) implies for b7 being Element of b2,((the_arity_of b5) /. b6) holds (b4 . ((the_arity_of b5) /. b6)) . b7 in the Sorts of b3 . ((the_arity_of b5) /. b6) ) )
proof end;

theorem Th9: :: MSUALG_6:9
for b1 being non empty non void ManySortedSign
for b2 being OperSymbol of b1
for b3 being Nat holds
( b3 in dom (the_arity_of b2) implies for b4, b5 being MSAlgebra of b1
for b6 being ManySortedFunction of b4,b5
for b7, b8 being Element of Args b2,b4 holds
( b7 in Args b2,b4 & b6 # b7 in Args b2,b5 implies for b9, b10, b11 being Function holds
( b9 = b7 & b10 = b6 # b7 & b11 = b6 # b8 implies for b12 being Element of b4,((the_arity_of b2) /. b3) holds
( b8 = b9 +* b3,b12 implies ( b11 . b3 = (b6 . ((the_arity_of b2) /. b3)) . b12 & b6 # b8 = b10 +* b3,(b11 . b3) ) ) ) ) )
proof end;

definition
let c1 be non empty non void ManySortedSign ;
let c2 be OperSymbol of c1;
let c3 be Nat;
let c4 be MSAlgebra of c1;
let c5 be Function;
func transl c2,c3,c5,c4 -> Function means :Def4: :: MSUALG_6:def 4
( dom a6 = the Sorts of a4 . ((the_arity_of a2) /. a3) & ( for b1 being set holds
( b1 in the Sorts of a4 . ((the_arity_of a2) /. a3) implies a6 . b1 = (Den a2,a4) . (a5 +* a3,b1) ) ) );
existence
ex b1 being Function st
( dom b1 = the Sorts of c4 . ((the_arity_of c2) /. c3) & ( for b2 being set holds
( b2 in the Sorts of c4 . ((the_arity_of c2) /. c3) implies b1 . b2 = (Den c2,c4) . (c5 +* c3,b2) ) ) )
proof end;
uniqueness
for b1, b2 being Function holds
( dom b1 = the Sorts of c4 . ((the_arity_of c2) /. c3) & ( for b3 being set holds
( b3 in the Sorts of c4 . ((the_arity_of c2) /. c3) implies b1 . b3 = (Den c2,c4) . (c5 +* c3,b3) ) ) & dom b2 = the Sorts of c4 . ((the_arity_of c2) /. c3) & ( for b3 being set holds
( b3 in the Sorts of c4 . ((the_arity_of c2) /. c3) implies b2 . b3 = (Den c2,c4) . (c5 +* c3,b3) ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def4 defines transl MSUALG_6:def 4 :
for b1 being non empty non void ManySortedSign
for b2 being OperSymbol of b1
for b3 being Nat
for b4 being MSAlgebra of b1
for b5, b6 being Function holds
( b6 = transl b2,b3,b5,b4 iff ( dom b6 = the Sorts of b4 . ((the_arity_of b2) /. b3) & ( for b7 being set holds
( b7 in the Sorts of b4 . ((the_arity_of b2) /. b3) implies b6 . b7 = (Den b2,b4) . (b5 +* b3,b7) ) ) ) );

theorem Th10: :: MSUALG_6:10
for b1 being non empty non void ManySortedSign
for b2 being OperSymbol of b1
for b3 being Nat holds
( b3 in dom (the_arity_of b2) implies for b4 being feasible MSAlgebra of b1
for b5 being Function holds
( b5 in Args b2,b4 implies transl b2,b3,b5,b4 is Function of the Sorts of b4 . ((the_arity_of b2) /. b3),the Sorts of b4 . (the_result_sort_of b2) ) )
proof end;

definition
let c1 be non empty non void ManySortedSign ;
let c2, c3 be SortSymbol of c1;
let c4 be MSAlgebra of c1;
let c5 be Function;
pred c5 is_e.translation_of c4,c2,c3 means :Def5: :: MSUALG_6:def 5
ex b1 being OperSymbol of a1 st
( the_result_sort_of b1 = a3 & ex b2 being Nat st
( b2 in dom (the_arity_of b1) & (the_arity_of b1) /. b2 = a2 & ex b3 being Function st
( b3 in Args b1,a4 & a5 = transl b1,b2,b3,a4 ) ) );
end;

:: deftheorem Def5 defines is_e.translation_of MSUALG_6:def 5 :
for b1 being non empty non void ManySortedSign
for b2, b3 being SortSymbol of b1
for b4 being MSAlgebra of b1
for b5 being Function holds
( b5 is_e.translation_of b4,b2,b3 iff ex b6 being OperSymbol of b1 st
( the_result_sort_of b6 = b3 & ex b7 being Nat st
( b7 in dom (the_arity_of b6) & (the_arity_of b6) /. b7 = b2 & ex b8 being Function st
( b8 in Args b6,b4 & b5 = transl b6,b7,b8,b4 ) ) ) );

theorem Th11: :: MSUALG_6:11
for b1 being non empty non void ManySortedSign
for b2, b3 being SortSymbol of b1
for b4 being feasible MSAlgebra of b1
for b5 being Function holds
( b5 is_e.translation_of b4,b2,b3 implies ( b5 is Function of the Sorts of b4 . b2,the Sorts of b4 . b3 & the Sorts of b4 . b2 <> {} & the Sorts of b4 . b3 <> {} ) )
proof end;

theorem Th12: :: MSUALG_6:12
for b1 being non empty non void ManySortedSign
for b2, b3 being SortSymbol of b1
for b4 being MSAlgebra of b1
for b5 being Function holds
( b5 is_e.translation_of b4,b2,b3 implies [b2,b3] in TranslationRel b1 )
proof end;

theorem Th13: :: MSUALG_6:13
for b1 being non empty non void ManySortedSign
for b2, b3 being SortSymbol of b1
for b4 being non-empty MSAlgebra of b1 holds
not ( [b2,b3] in TranslationRel b1 & ( for b5 being Function holds
not b5 is_e.translation_of b4,b2,b3 ) )
proof end;

theorem Th14: :: MSUALG_6:14
for b1 being non empty non void ManySortedSign
for b2 being feasible MSAlgebra of b1
for b3, b4 being SortSymbol of b1 holds
( TranslationRel b1 reduces b3,b4 implies for b5 being RedSequence of TranslationRel b1
for b6 being Function-yielding FinSequence holds
( len b5 = (len b6) + 1 & b3 = b5 . 1 & b4 = b5 . (len b5) & ( for b7 being Nat
for b8 being Function
for b9, b10 being SortSymbol of b1 holds
( b7 in dom b6 & b8 = b6 . b7 & b9 = b5 . b7 & b10 = b5 . (b7 + 1) implies b8 is_e.translation_of b2,b9,b10 ) ) implies ( compose b6,(the Sorts of b2 . b3) is Function of the Sorts of b2 . b3,the Sorts of b2 . b4 & ( b6 <> {} implies ( the Sorts of b2 . b3 <> {} & the Sorts of b2 . b4 <> {} ) ) ) ) )
proof end;

definition
let c1 be non empty non void ManySortedSign ;
let c2 be non-empty MSAlgebra of c1;
let c3, c4 be SortSymbol of c1;
assume E19: TranslationRel c1 reduces c3,c4 ;
mode Translation of c2,c3,c4 -> Function of the Sorts of a2 . a3,the Sorts of a2 . a4 means :Def6: :: MSUALG_6:def 6
ex b1 being RedSequence of TranslationRel a1ex b2 being Function-yielding FinSequence st
( a5 = compose b2,(the Sorts of a2 . a3) & len b1 = (len b2) + 1 & a3 = b1 . 1 & a4 = b1 . (len b1) & ( for b3 being Nat
for b4 being Function
for b5, b6 being SortSymbol of a1 holds
( b3 in dom b2 & b4 = b2 . b3 & b5 = b1 . b3 & b6 = b1 . (b3 + 1) implies b4 is_e.translation_of a2,b5,b6 ) ) );
existence
ex b1 being Function of the Sorts of c2 . c3,the Sorts of c2 . c4ex b2 being RedSequence of TranslationRel c1ex b3 being Function-yielding FinSequence st
( b1 = compose b3,(the Sorts of c2 . c3) & len b2 = (len b3) + 1 & c3 = b2 . 1 & c4 = b2 . (len b2) & ( for b4 being Nat
for b5 being Function
for b6, b7 being SortSymbol of c1 holds
( b4 in dom b3 & b5 = b3 . b4 & b6 = b2 . b4 & b7 = b2 . (b4 + 1) implies b5 is_e.translation_of c2,b6,b7 ) ) )
proof end;
end;

:: deftheorem Def6 defines Translation MSUALG_6:def 6 :
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3, b4 being SortSymbol of b1 holds
( TranslationRel b1 reduces b3,b4 implies for b5 being Function of the Sorts of b2 . b3,the Sorts of b2 . b4 holds
( b5 is Translation of b2,b3,b4 iff ex b6 being RedSequence of TranslationRel b1ex b7 being Function-yielding FinSequence st
( b5 = compose b7,(the Sorts of b2 . b3) & len b6 = (len b7) + 1 & b3 = b6 . 1 & b4 = b6 . (len b6) & ( for b8 being Nat
for b9 being Function
for b10, b11 being SortSymbol of b1 holds
( b8 in dom b7 & b9 = b7 . b8 & b10 = b6 . b8 & b11 = b6 . (b8 + 1) implies b9 is_e.translation_of b2,b10,b11 ) ) ) ) );

theorem Th15: :: MSUALG_6:15
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3, b4 being SortSymbol of b1 holds
( TranslationRel b1 reduces b3,b4 implies for b5 being RedSequence of TranslationRel b1
for b6 being Function-yielding FinSequence holds
( len b5 = (len b6) + 1 & b3 = b5 . 1 & b4 = b5 . (len b5) & ( for b7 being Nat
for b8 being Function
for b9, b10 being SortSymbol of b1 holds
( b7 in dom b6 & b8 = b6 . b7 & b9 = b5 . b7 & b10 = b5 . (b7 + 1) implies b8 is_e.translation_of b2,b9,b10 ) ) implies compose b6,(the Sorts of b2 . b3) is Translation of b2,b3,b4 ) )
proof end;

theorem Th16: :: MSUALG_6:16
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being SortSymbol of b1 holds
id (the Sorts of b2 . b3) is Translation of b2,b3,b3
proof end;

theorem Th17: :: MSUALG_6:17
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3, b4 being SortSymbol of b1
for b5 being Function holds
( b5 is_e.translation_of b2,b3,b4 implies ( TranslationRel b1 reduces b3,b4 & b5 is Translation of b2,b3,b4 ) )
proof end;

theorem Th18: :: MSUALG_6:18
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3, b4, b5 being SortSymbol of b1 holds
( TranslationRel b1 reduces b3,b4 & TranslationRel b1 reduces b4,b5 implies for b6 being Translation of b2,b3,b4
for b7 being Translation of b2,b4,b5 holds
b7 * b6 is Translation of b2,b3,b5 )
proof end;

theorem Th19: :: MSUALG_6:19
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3, b4, b5 being SortSymbol of b1 holds
( TranslationRel b1 reduces b3,b4 implies for b6 being Translation of b2,b3,b4
for b7 being Function holds
( b7 is_e.translation_of b2,b4,b5 implies b7 * b6 is Translation of b2,b3,b5 ) )
proof end;

theorem Th20: :: MSUALG_6:20
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3, b4, b5 being SortSymbol of b1 holds
( TranslationRel b1 reduces b4,b5 implies for b6 being Function holds
( b6 is_e.translation_of b2,b3,b4 implies for b7 being Translation of b2,b4,b5 holds
b7 * b6 is Translation of b2,b3,b5 ) )
proof end;

scheme :: MSUALG_6:sch 1
s1{ F1() -> non empty non void ManySortedSign , F2() -> non-empty MSAlgebra of F1(), P1[ set , set , set ] } :
for b1, b2 being SortSymbol of F1() holds
( TranslationRel F1() reduces b1,b2 implies for b3 being Translation of F2(),b1,b2 holds P1[b3,b1,b2] )
provided
E24: for b1 being SortSymbol of F1() holds P1[ id (the Sorts of F2() . b1),b1,b1] and E25: for b1, b2, b3 being SortSymbol of F1() holds
( TranslationRel F1() reduces b1,b2 implies for b4 being Translation of F2(),b1,b2 holds
( P1[b4,b1,b2] implies for b5 being Function holds
( b5 is_e.translation_of F2(),b2,b3 implies P1[b5 * b4,b1,b3] ) ) )
proof end;

theorem Th21: :: MSUALG_6:21
for b1 being non empty non void ManySortedSign
for b2, b3 being non-empty MSAlgebra of b1
for b4 being ManySortedFunction of b2,b3 holds
( b4 is_homomorphism b2,b3 implies for b5 being OperSymbol of b1
for b6 being Nat holds
( b6 in dom (the_arity_of b5) implies for b7 being Element of Args b5,b2 holds (b4 . (the_result_sort_of b5)) * (transl b5,b6,b7,b2) = (transl b5,b6,(b4 # b7),b3) * (b4 . ((the_arity_of b5) /. b6)) ) )
proof end;

theorem Th22: :: MSUALG_6:22
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being Endomorphism of b2
for b4 being OperSymbol of b1
for b5 being Nat holds
( b5 in dom (the_arity_of b4) implies for b6 being Element of Args b4,b2 holds (b3 . (the_result_sort_of b4)) * (transl b4,b5,b6,b2) = (transl b4,b5,(b3 # b6),b2) * (b3 . ((the_arity_of b4) /. b5)) )
proof end;

theorem Th23: :: MSUALG_6:23
for b1 being non empty non void ManySortedSign
for b2, b3 being non-empty MSAlgebra of b1
for b4 being ManySortedFunction of b2,b3 holds
( b4 is_homomorphism b2,b3 implies for b5, b6 being SortSymbol of b1
for b7 being Function holds
not ( b7 is_e.translation_of b2,b5,b6 & ( for b8 being Function of the Sorts of b3 . b5,the Sorts of b3 . b6 holds
not ( b8 is_e.translation_of b3,b5,b6 & b8 * (b4 . b5) = (b4 . b6) * b7 ) ) ) )
proof end;

theorem Th24: :: MSUALG_6:24
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being Endomorphism of b2
for b4, b5 being SortSymbol of b1
for b6 being Function holds
not ( b6 is_e.translation_of b2,b4,b5 & ( for b7 being Function of the Sorts of b2 . b4,the Sorts of b2 . b5 holds
not ( b7 is_e.translation_of b2,b4,b5 & b7 * (b3 . b4) = (b3 . b5) * b6 ) ) )
proof end;

theorem Th25: :: MSUALG_6:25
for b1 being non empty non void ManySortedSign
for b2, b3 being non-empty MSAlgebra of b1
for b4 being ManySortedFunction of b2,b3 holds
( b4 is_homomorphism b2,b3 implies for b5, b6 being SortSymbol of b1 holds
( TranslationRel b1 reduces b5,b6 implies for b7 being Translation of b2,b5,b6 holds
ex b8 being Translation of b3,b5,b6 st b8 * (b4 . b5) = (b4 . b6) * b7 ) )
proof end;

theorem Th26: :: MSUALG_6:26
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being Endomorphism of b2
for b4, b5 being SortSymbol of b1 holds
( TranslationRel b1 reduces b4,b5 implies for b6 being Translation of b2,b4,b5 holds
ex b7 being Translation of b2,b4,b5 st b7 * (b3 . b4) = (b3 . b5) * b6 )
proof end;

definition
let c1 be non empty non void ManySortedSign ;
let c2 be MSAlgebra of c1;
let c3 be ManySortedRelation of c2;
attr a3 is compatible means :Def7: :: MSUALG_6:def 7
for b1 being OperSymbol of a1
for b2, b3 being Function holds
( b2 in Args b1,a2 & b3 in Args b1,a2 & ( for b4 being Nat holds
( b4 in dom (the_arity_of b1) implies [(b2 . b4),(b3 . b4)] in a3 . ((the_arity_of b1) /. b4) ) ) implies [((Den b1,a2) . b2),((Den b1,a2) . b3)] in a3 . (the_result_sort_of b1) );
attr a3 is invariant means :Def8: :: MSUALG_6:def 8
for b1, b2 being SortSymbol of a1
for b3 being Function holds
( b3 is_e.translation_of a2,b1,b2 implies for b4, b5 being set holds
( [b4,b5] in a3 . b1 implies [(b3 . b4),(b3 . b5)] in a3 . b2 ) );
attr a3 is stable means :Def9: :: MSUALG_6:def 9
for b1 being Endomorphism of a2
for b2 being SortSymbol of a1
for b3, b4 being set holds
( [b3,b4] in a3 . b2 implies [((b1 . b2) . b3),((b1 . b2) . b4)] in a3 . b2 );
end;

:: deftheorem Def7 defines compatible MSUALG_6:def 7 :
for b1 being non empty non void ManySortedSign
for b2 being MSAlgebra of b1
for b3 being ManySortedRelation of b2 holds
( b3 is compatible iff for b4 being OperSymbol of b1
for b5, b6 being Function holds
( b5 in Args b4,b2 & b6 in Args b4,b2 & ( for b7 being Nat holds
( b7 in dom (the_arity_of b4) implies [(b5 . b7),(b6 . b7)] in b3 . ((the_arity_of b4) /. b7) ) ) implies [((Den b4,b2) . b5),((Den b4,b2) . b6)] in b3 . (the_result_sort_of b4) ) );

:: deftheorem Def8 defines invariant MSUALG_6:def 8 :
for b1 being non empty non void ManySortedSign
for b2 being MSAlgebra of b1
for b3 being ManySortedRelation of b2 holds
( b3 is invariant iff for b4, b5 being SortSymbol of b1
for b6 being Function holds
( b6 is_e.translation_of b2,b4,b5 implies for b7, b8 being set holds
( [b7,b8] in b3 . b4 implies [(b6 . b7),(b6 . b8)] in b3 . b5 ) ) );

:: deftheorem Def9 defines stable MSUALG_6:def 9 :
for b1 being non empty non void ManySortedSign
for b2 being MSAlgebra of b1
for b3 being ManySortedRelation of b2 holds
( b3 is stable iff for b4 being Endomorphism of b2
for b5 being SortSymbol of b1
for b6, b7 being set holds
( [b6,b7] in b3 . b5 implies [((b4 . b5) . b6),((b4 . b5) . b7)] in b3 . b5 ) );

theorem Th27: :: MSUALG_6:27
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being MSEquivalence-like ManySortedRelation of b2 holds
( b3 is compatible iff b3 is MSCongruence of b2 )
proof end;

theorem Th28: :: MSUALG_6:28
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being ManySortedRelation of b2 holds
( b3 is invariant iff for b4, b5 being SortSymbol of b1 holds
( TranslationRel b1 reduces b4,b5 implies for b6 being Translation of b2,b4,b5
for b7, b8 being set holds
( [b7,b8] in b3 . b4 implies [(b6 . b7),(b6 . b8)] in b3 . b5 ) ) )
proof end;

registration
let c1 be non empty non void ManySortedSign ;
let c2 be non-empty MSAlgebra of c1;
cluster MSEquivalence-like invariant -> MSEquivalence-like compatible ManySortedRelation of the Sorts of a2,the Sorts of a2;
coherence
for b1 being MSEquivalence-like ManySortedRelation of c2 holds
( b1 is invariant implies b1 is compatible )
proof end;
cluster MSEquivalence-like compatible -> MSEquivalence-like invariant ManySortedRelation of the Sorts of a2,the Sorts of a2;
coherence
for b1 being MSEquivalence-like ManySortedRelation of c2 holds
( b1 is compatible implies b1 is invariant )
proof end;
end;

registration
let c1 be non empty set ;
cluster id a1 -> non empty ;
coherence
not id c1 is empty
;
end;

scheme :: MSUALG_6:sch 2
s2{ F1() -> non empty set , F2() -> V3 ManySortedSet of F1(), P1[ set , set , set ] } :
ex b1 being ManySortedRelation of F2() st
for b2 being Element of F1()
for b3, b4 being Element of F2() . b2 holds
( [b3,b4] in b1 . b2 iff P1[b2,b3,b4] )
proof end;

scheme :: MSUALG_6:sch 3
s3{ F1() -> set , F2() -> ManySortedSet of F1(), F3( set ) -> set } :
( ex b1 being ManySortedRelation of F2() st
for b2 being set holds
( b2 in F1() implies b1 . b2 = F3(b2) ) & ( for b1, b2 being ManySortedRelation of F2() holds
( ( for b3 being set holds
( b3 in F1() implies b1 . b3 = F3(b3) ) ) & ( for b3 being set holds
( b3 in F1() implies b2 . b3 = F3(b3) ) ) implies b1 = b2 ) ) )
provided
E32: for b1 being set holds
( b1 in F1() implies F3(b1) is Relation of F2() . b1,F2() . b1 )
proof end;

definition
let c1 be set ;
let c2 be ManySortedSet of c1;
func id c1,c2 -> ManySortedRelation of a2 means :Def10: :: MSUALG_6:def 10
for b1 being set holds
( b1 in a1 implies a3 . b1 = id (a2 . b1) );
correctness
existence
ex b1 being ManySortedRelation of c2 st
for b2 being set holds
( b2 in c1 implies b1 . b2 = id (c2 . b2) )
;
uniqueness
for b1, b2 being ManySortedRelation of c2 holds
( ( for b3 being set holds
( b3 in c1 implies b1 . b3 = id (c2 . b3) ) ) & ( for b3 being set holds
( b3 in c1 implies b2 . b3 = id (c2 . b3) ) ) implies b1 = b2 )
;
proof end;
end;

:: deftheorem Def10 defines id MSUALG_6:def 10 :
for b1 being set
for b2 being ManySortedSet of b1
for b3 being ManySortedRelation of b2 holds
( b3 = id b1,b2 iff for b4 being set holds
( b4 in b1 implies b3 . b4 = id (b2 . b4) ) );

registration
let c1 be non empty non void ManySortedSign ;
let c2 be non-empty MSAlgebra of c1;
cluster MSEquivalence-like -> V3 ManySortedRelation of the Sorts of a2,the Sorts of a2;
coherence
for b1 being ManySortedRelation of c2 holds
( b1 is MSEquivalence-like implies b1 is non-empty )
proof end;
end;

registration
let c1 be non empty non void ManySortedSign ;
let c2 be non-empty MSAlgebra of c1;
cluster V3 MSEquivalence-like compatible invariant stable ManySortedRelation of the Sorts of a2,the Sorts of a2;
existence
ex b1 being ManySortedRelation of c2 st
( b1 is invariant & b1 is stable & b1 is MSEquivalence-like )
proof end;
end;

scheme :: MSUALG_6:sch 4
s4{ F1() -> non empty non void ManySortedSign , F2() -> non-empty MSAlgebra of F1(), P1[ set , set , set ], P2[ set ], F3() -> ManySortedRelation of F2(), F4() -> ManySortedRelation of F2() } :
( P2[F4()] & F3() c= F4() & ( for b1 being ManySortedRelation of F2() holds
( P2[b1] & F3() c= b1 implies F4() c= b1 ) ) )
provided
E33: for b1 being ManySortedRelation of F2() holds
( P2[b1] iff for b2, b3 being SortSymbol of F1()
for b4 being Function of the Sorts of F2() . b2,the Sorts of F2() . b3 holds
( P1[b4,b2,b3] implies for b5, b6 being set holds
( [b5,b6] in b1 . b2 implies [(b4 . b5),(b4 . b6)] in b1 . b3 ) ) ) and E34: for b1, b2, b3 being SortSymbol of F1()
for b4 being Function of the Sorts of F2() . b1,the Sorts of F2() . b2
for b5 being Function of the Sorts of F2() . b2,the Sorts of F2() . b3 holds
( P1[b4,b1,b2] & P1[b5,b2,b3] implies P1[b5 * b4,b1,b3] ) and E35: for b1 being SortSymbol of F1() holds P1[ id (the Sorts of F2() . b1),b1,b1] and E36: for b1 being SortSymbol of F1()
for b2, b3 being Element of F2(),b1 holds
( [b2,b3] in F4() . b1 iff ex b4 being SortSymbol of F1()ex b5 being Function of the Sorts of F2() . b4,the Sorts of F2() . b1ex b6, b7 being Element of F2(),b4 st
( P1[b5,b4,b1] & [b6,b7] in F3() . b4 & b2 = b5 . b6 & b3 = b5 . b7 ) )
proof end;

E33: now
let c1 be non empty non void ManySortedSign ;
let c2 be non-empty MSAlgebra of c1;
let c3, c4 be ManySortedRelation of c2;
defpred S1[ ManySortedRelation of c2] means a1 is invariant;
defpred S2[ Function, SortSymbol of c1, SortSymbol of c1] means ( TranslationRel c1 reduces a2,a3 & a1 is Translation of c2,a2,a3 );
assume E34: for b1 being SortSymbol of c1
for b2, b3 being Element of c2,b1 holds
( [b2,b3] in c4 . b1 iff ex b4 being SortSymbol of c1ex b5 being Function of the Sorts of c2 . b4,the Sorts of c2 . b1ex b6, b7 being Element of c2,b4 st
( S2[b5,b4,b1] & [b6,b7] in c3 . b4 & b2 = b5 . b6 & b3 = b5 . b7 ) ) ;
E35: now
let c5 be ManySortedRelation of c2;
thus ( S1[c5] implies for b1, b2 being SortSymbol of c1
for b3 being Function of the Sorts of c2 . b1,the Sorts of c2 . b2 holds
( S2[b3,b1,b2] implies for b4, b5 being set holds
( [b4,b5] in c5 . b1 implies [(b3 . b4),(b3 . b5)] in c5 . b2 ) ) ) by Th28;
assume for b1, b2 being SortSymbol of c1
for b3 being Function of the Sorts of c2 . b1,the Sorts of c2 . b2 holds
( S2[b3,b1,b2] implies for b4, b5 being set holds
( [b4,b5] in c5 . b1 implies [(b3 . b4),(b3 . b5)] in c5 . b2 ) ) ;
then for b1, b2 being SortSymbol of c1 holds
( TranslationRel c1 reduces b1,b2 implies for b3 being Translation of c2,b1,b2
for b4, b5 being set holds
( [b4,b5] in c5 . b1 implies [(b3 . b4),(b3 . b5)] in c5 . b2 ) ) ;
hence S1[c5] by Th28;
end;
E36: for b1, b2, b3 being SortSymbol of c1
for b4 being Function of the Sorts of c2 . b1,the Sorts of c2 . b2
for b5 being Function of the Sorts of c2 . b2,the Sorts of c2 . b3 holds
( S2[b4,b1,b2] & S2[b5,b2,b3] implies S2[b5 * b4,b1,b3] ) by Th18, REWRITE1:17;
E37: for b1 being SortSymbol of c1 holds
S2[ id (the Sorts of c2 . b1),b1,b1] by Th16, REWRITE1:13;
thus ( S1[c4] & c3 c= c4 & ( for b1 being ManySortedRelation of c2 holds
( S1[b1] & c3 c= b1 implies c4 c= b1 ) ) ) from MSUALG_6:sch 4(E35, E36, E37, E34);
end;

definition
let c1 be non empty non void ManySortedSign ;
let c2 be non-empty MSAlgebra of c1;
let c3 be ManySortedRelation of the Sorts of c2;
func InvCl c3 -> invariant ManySortedRelation of a2 means :Def11: :: MSUALG_6:def 11
( a3 c= a4 & ( for b1 being invariant ManySortedRelation of a2 holds
( a3 c= b1 implies a4 c= b1 ) ) );
uniqueness
for b1, b2 being invariant ManySortedRelation of c2 holds
( c3 c= b1 & ( for b3 being invariant ManySortedRelation of c2 holds
( c3 c= b3 implies b1 c= b3 ) ) & c3 c= b2 & ( for b3 being invariant ManySortedRelation of c2 holds
( c3 c= b3 implies b2 c= b3 ) ) implies b1 = b2 )
proof end;
existence
ex b1 being invariant ManySortedRelation of c2 st
( c3 c= b1 & ( for b2 being invariant ManySortedRelation of c2 holds
( c3 c= b2 implies b1 c= b2 ) ) )
proof end;
end;

:: deftheorem Def11 defines InvCl MSUALG_6:def 11 :
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being ManySortedRelation of the Sorts of b2
for b4 being invariant ManySortedRelation of b2 holds
( b4 = InvCl b3 iff ( b3 c= b4 & ( for b5 being invariant ManySortedRelation of b2 holds
( b3 c= b5 implies b4 c= b5 ) ) ) );

theorem Th29: :: MSUALG_6:29
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being ManySortedRelation of the Sorts of b2
for b4 being SortSymbol of b1
for b5, b6 being Element of b2,b4 holds
( [b5,b6] in (InvCl b3) . b4 iff ex b7 being SortSymbol of b1ex b8, b9 being Element of b2,b7ex b10 being Translation of b2,b7,b4 st
( TranslationRel b1 reduces b7,b4 & [b8,b9] in b3 . b7 & b5 = b10 . b8 & b6 = b10 . b9 ) )
proof end;

theorem Th30: :: MSUALG_6:30
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being stable ManySortedRelation of b2 holds InvCl b3 is stable
proof end;

E37: now
let c1 be non empty non void ManySortedSign ;
let c2 be non-empty MSAlgebra of c1;
let c3, c4 be ManySortedRelation of c2;
defpred S1[ ManySortedRelation of c2] means a1 is stable;
defpred S2[ Function, SortSymbol of c1, SortSymbol of c1] means ( a2 = a3 & ex b1 being Endomorphism of c2 st a1 = b1 . a2 );
assume E38: for b1 being SortSymbol of c1
for b2, b3 being Element of c2,b1 holds
( [b2,b3] in c4 . b1 iff ex b4 being SortSymbol of c1ex b5 being Function of the Sorts of c2 . b4,the Sorts of c2 . b1ex b6, b7 being Element of c2,b4 st
( S2[b5,b4,b1] & [b6,b7] in c3 . b4 & b2 = b5 . b6 & b3 = b5 . b7 ) ) ;
E39: for b1 being ManySortedRelation of c2 holds
( S1[b1] iff for b2, b3 being SortSymbol of c1
for b4 being Function of the Sorts of c2 . b2,the Sorts of c2 . b3 holds
( S2[b4,b2,b3] implies for b5, b6 being set holds
( [b5,b6] in b1 . b2 implies [(b4 . b5),(b4 . b6)] in b1 . b3 ) ) )
proof
let c5 be ManySortedRelation of c2;
thus ( c5 is stable implies for b1, b2 being SortSymbol of c1
for b3 being Function of the Sorts of c2 . b1,the Sorts of c2 . b2 holds
( b1 = b2 & ex b4 being Endomorphism of c2 st b3 = b4 . b1 implies for b4, b5 being set holds
( [b4,b5] in c5 . b1 implies [(b3 . b4),(b3 . b5)] in c5 . b2 ) ) ) by Def9;
assume E40: for b1, b2 being SortSymbol of c1
for b3 being Function of the Sorts of c2 . b1,the Sorts of c2 . b2 holds
( S2[b3,b1,b2] implies for b4, b5 being set holds
( [b4,b5] in c5 . b1 implies [(b3 . b4),(b3 . b5)] in c5 . b2 ) ) ;
thus c5 is stable
proof
let c6 be Endomorphism of c2; :: according to MSUALG_6:def 9
let c7 be SortSymbol of c1;
thus for b1, b2 being set holds
( [b1,b2] in c5 . c7 implies [((c6 . c7) . b1),((c6 . c7) . b2)] in c5 . c7 ) by E40;
end;
end;
E40: for b1, b2, b3 being SortSymbol of c1
for b4 being Function of the Sorts of c2 . b1,the Sorts of c2 . b2
for b5 being Function of the Sorts of c2 . b2,the Sorts of c2 . b3 holds
( S2[b4,b1,b2] & S2[b5,b2,b3] implies S2[b5 * b4,b1,b3] )
proof
let c5, c6, c7 be SortSymbol of c1;
let c8 be Function of the Sorts of c2 . c5,the Sorts of c2 . c6;
let c9 be Function of the Sorts of c2 . c6,the Sorts of c2 . c7;
assume E41: c5 = c6 ;
given c10 being Endomorphism of c2 such that E42: c8 = c10 . c5 ;
assume E43: c6 = c7 ;
given c11 being Endomorphism of c2 such that E44: c9 = c11 . c6 ;
thus c5 = c7 by E41, E43;
reconsider c12 = c11 ** c10 as Endomorphism of c2 ;
take c12 ;
thus c9 * c8 = c12 . c5 by E41, E42, E44, MSUALG_3:2;
end;
E41: for b1 being SortSymbol of c1 holds
S2[ id (the Sorts of c2 . b1),b1,b1]
proof
let c5 be SortSymbol of c1;
thus c5 = c5 ;
reconsider c6 = id the Sorts of c2 as Endomorphism of c2 by Th4;
take c6 ;
thus id (the Sorts of c2 . c5) = c6 . c5 by MSUALG_3:def 1;
end;
thus ( S1[c4] & c3 c= c4 & ( for b1 being ManySortedRelation of c2 holds
( S1[b1] & c3 c= b1 implies c4 c= b1 ) ) ) from MSUALG_6:sch 4(E39, E40, E41, E38);
end;

definition
let c1 be non empty non void ManySortedSign ;
let c2 be non-empty MSAlgebra of c1;
let c3 be ManySortedRelation of the Sorts of c2;
func StabCl c3 -> stable ManySortedRelation of a2 means :Def12: :: MSUALG_6:def 12
( a3 c= a4 & ( for b1 being stable ManySortedRelation of a2 holds
( a3 c= b1 implies a4 c= b1 ) ) );
uniqueness
for b1, b2 being stable ManySortedRelation of c2 holds
( c3 c= b1 & ( for b3 being stable ManySortedRelation of c2 holds
( c3 c= b3 implies b1 c= b3 ) ) & c3 c= b2 & ( for b3 being stable ManySortedRelation of c2 holds
( c3 c= b3 implies b2 c= b3 ) ) implies b1 = b2 )
proof end;
existence
ex b1 being stable ManySortedRelation of c2 st
( c3 c= b1 & ( for b2 being stable ManySortedRelation of c2 holds
( c3 c= b2 implies b1 c= b2 ) ) )
proof end;
end;

:: deftheorem Def12 defines StabCl MSUALG_6:def 12 :
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being ManySortedRelation of the Sorts of b2
for b4 being stable ManySortedRelation of b2 holds
( b4 = StabCl b3 iff ( b3 c= b4 & ( for b5 being stable ManySortedRelation of b2 holds
( b3 c= b5 implies b4 c= b5 ) ) ) );

theorem Th31: :: MSUALG_6:31
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being ManySortedRelation of the Sorts of b2
for b4 being SortSymbol of b1
for b5, b6 being Element of b2,b4 holds
( [b5,b6] in (StabCl b3) . b4 iff ex b7, b8 being Element of b2,b4ex b9 being Endomorphism of b2 st
( [b7,b8] in b3 . b4 & b5 = (b9 . b4) . b7 & b6 = (b9 . b4) . b8 ) )
proof end;

theorem Th32: :: MSUALG_6:32
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being ManySortedRelation of the Sorts of b2 holds InvCl (StabCl b3) is stable by Th30;

definition
let c1 be non empty non void ManySortedSign ;
let c2 be non-empty MSAlgebra of c1;
let c3 be ManySortedRelation of the Sorts of c2;
func TRS c3 -> invariant stable ManySortedRelation of a2 means :Def13: :: MSUALG_6:def 13
( a3 c= a4 & ( for b1 being invariant stable ManySortedRelation of a2 holds
( a3 c= b1 implies a4 c= b1 ) ) );
uniqueness
for b1, b2 being invariant stable ManySortedRelation of c2 holds
( c3 c= b1 & ( for b3 being invariant stable ManySortedRelation of c2 holds
( c3 c= b3 implies b1 c= b3 ) ) & c3 c= b2 & ( for b3 being invariant stable ManySortedRelation of c2 holds
( c3 c= b3 implies b2 c= b3 ) ) implies b1 = b2 )
proof end;
existence
ex b1 being invariant stable ManySortedRelation of c2 st
( c3 c= b1 & ( for b2 being invariant stable ManySortedRelation of c2 holds
( c3 c= b2 implies b1 c= b2 ) ) )
proof end;
end;

:: deftheorem Def13 defines TRS MSUALG_6:def 13 :
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being ManySortedRelation of the Sorts of b2
for b4 being invariant stable ManySortedRelation of b2 holds
( b4 = TRS b3 iff ( b3 c= b4 & ( for b5 being invariant stable ManySortedRelation of b2 holds
( b3 c= b5 implies b4 c= b5 ) ) ) );

registration
let c1 be non empty non void ManySortedSign ;
let c2 be non-empty MSAlgebra of c1;
let c3 be V3 ManySortedRelation of c2;
cluster InvCl a3 -> V3 invariant ;
coherence
InvCl c3 is non-empty
proof end;
cluster StabCl a3 -> V3 stable ;
coherence
StabCl c3 is non-empty
proof end;
cluster TRS a3 -> V3 invariant stable ;
coherence
TRS c3 is non-empty
proof end;
end;

theorem Th33: :: MSUALG_6:33
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being invariant ManySortedRelation of b2 holds InvCl b3 = b3
proof end;

theorem Th34: :: MSUALG_6:34
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being stable ManySortedRelation of b2 holds StabCl b3 = b3
proof end;

theorem Th35: :: MSUALG_6:35
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being invariant stable ManySortedRelation of b2 holds TRS b3 = b3
proof end;

theorem Th36: :: MSUALG_6:36
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being ManySortedRelation of the Sorts of b2 holds
( StabCl b3 c= TRS b3 & InvCl b3 c= TRS b3 & StabCl (InvCl b3) c= TRS b3 )
proof end;

theorem Th37: :: MSUALG_6:37
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being ManySortedRelation of the Sorts of b2 holds InvCl (StabCl b3) = TRS b3
proof end;

theorem Th38: :: MSUALG_6:38
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being ManySortedRelation of the Sorts of b2
for b4 being SortSymbol of b1
for b5, b6 being Element of b2,b4 holds
( [b5,b6] in (TRS b3) . b4 iff ex b7 being SortSymbol of b1 st
( TranslationRel b1 reduces b7,b4 & ex b8, b9 being Element of b2,b7ex b10 being Endomorphism of b2ex b11 being Translation of b2,b7,b4 st
( [b8,b9] in b3 . b7 & b5 = b11 . ((b10 . b7) . b8) & b6 = b11 . ((b10 . b7) . b9) ) ) )
proof end;

theorem Th39: :: MSUALG_6:39
for b1 being set
for b2, b3 being Relation of b1 holds
( ( for b4, b5 being set holds
( b4 in b1 & b5 in b1 implies ( [b4,b5] in b3 iff b4,b5 are_convertible_wrt b2 ) ) ) implies ( b3 is total & b3 is symmetric & b3 is transitive ) )
proof end;

theorem Th40: :: MSUALG_6:40
for b1 being set
for b2 being Relation of b1
for b3 being Equivalence_Relation of b1 holds
( b2 c= b3 implies for b4, b5 being set holds
( b4 in b1 & b5 in b1 & b4,b5 are_convertible_wrt b2 implies [b4,b5] in b3 ) )
proof end;

theorem Th41: :: MSUALG_6:41
for b1 being non empty set
for b2 being Relation of b1
for b3, b4 being Element of b1 holds
( [b3,b4] in EqCl b2 iff b3,b4 are_convertible_wrt b2 )
proof end;

theorem Th42: :: MSUALG_6:42
for b1 being non empty set
for b2 being V3 ManySortedSet of b1
for b3 being ManySortedRelation of b2
for b4 being Element of b1
for b5, b6 being Element of b2 . b4 holds
( [b5,b6] in (EqCl b3) . b4 iff b5,b6 are_convertible_wrt b3 . b4 )
proof end;

definition
let c1 be non empty non void ManySortedSign ;
let c2 be non-empty MSAlgebra of c1;
mode EquationalTheory is MSEquivalence-like invariant stable ManySortedRelation of a2;
let c3 be ManySortedRelation of c2;
func EqCl c3,c2 -> MSEquivalence-like ManySortedRelation of a2 equals :: MSUALG_6:def 14
EqCl a3;
coherence
EqCl c3 is MSEquivalence-like ManySortedRelation of c2
by MSUALG_4:def 5;
end;

:: deftheorem Def14 defines EqCl MSUALG_6:def 14 :
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being ManySortedRelation of b2 holds EqCl b3,b2 = EqCl b3;

theorem Th43: :: MSUALG_6:43
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being ManySortedRelation of b2 holds b3 c= EqCl b3,b2
proof end;

theorem Th44: :: MSUALG_6:44
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being ManySortedRelation of b2
for b4 being MSEquivalence-like ManySortedRelation of b2 holds
( b3 c= b4 implies EqCl b3,b2 c= b4 )
proof end;

theorem Th45: :: MSUALG_6:45
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being stable ManySortedRelation of b2
for b4 being SortSymbol of b1
for b5, b6 being Element of b2,b4 holds
( b5,b6 are_convertible_wrt b3 . b4 implies for b7 being Endomorphism of b2 holds (b7 . b4) . b5,(b7 . b4) . b6 are_convertible_wrt b3 . b4 )
proof end;

theorem Th46: :: MSUALG_6:46
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being stable ManySortedRelation of b2 holds EqCl b3,b2 is stable
proof end;

registration
let c1 be non empty non void ManySortedSign ;
let c2 be non-empty MSAlgebra of c1;
let c3 be stable ManySortedRelation of c2;
cluster EqCl a3,a2 -> V3 MSEquivalence-like stable ;
coherence
EqCl c3,c2 is stable
by Th46;
end;

theorem Th47: :: MSUALG_6:47
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being invariant ManySortedRelation of b2
for b4, b5 being SortSymbol of b1
for b6, b7 being Element of b2,b4 holds
( b6,b7 are_convertible_wrt b3 . b4 implies for b8 being Function holds
( b8 is_e.translation_of b2,b4,b5 implies b8 . b6,b8 . b7 are_convertible_wrt b3 . b5 ) )
proof end;

theorem Th48: :: MSUALG_6:48
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being invariant ManySortedRelation of b2 holds EqCl b3,b2 is invariant
proof end;

registration
let c1 be non empty non void ManySortedSign ;
let c2 be non-empty MSAlgebra of c1;
let c3 be invariant ManySortedRelation of c2;
cluster EqCl a3,a2 -> V3 MSEquivalence-like compatible invariant ;
coherence
EqCl c3,c2 is invariant
by Th48;
end;

theorem Th49: :: MSUALG_6:49
for b1 being non empty set
for b2 being V3 ManySortedSet of b1
for b3, b4 being ManySortedRelation of b2 holds
( ( for b5 being Element of b1
for b6, b7 being Element of b2 . b5 holds
( [b6,b7] in b4 . b5 iff b6,b7 are_convertible_wrt b3 . b5 ) ) implies b4 is MSEquivalence_Relation-like )
proof end;

theorem Th50: :: MSUALG_6:50
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3, b4 being ManySortedRelation of b2 holds
( ( for b5 being SortSymbol of b1
for b6, b7 being Element of b2,b5 holds
( [b6,b7] in b4 . b5 iff b6,b7 are_convertible_wrt (TRS b3) . b5 ) ) implies b4 is EquationalTheory of b2 )
proof end;

theorem Th51: :: MSUALG_6:51
for b1 being non empty set
for b2 being V3 ManySortedSet of b1
for b3 being ManySortedRelation of b2
for b4 being MSEquivalence_Relation-like ManySortedRelation of b2 holds
( b3 c= b4 implies for b5 being Element of b1
for b6, b7 being Element of b2 . b5 holds
( b6,b7 are_convertible_wrt b3 . b5 implies [b6,b7] in b4 . b5 ) )
proof end;

definition
let c1 be non empty non void ManySortedSign ;
let c2 be non-empty MSAlgebra of c1;
let c3 be ManySortedRelation of the Sorts of c2;
func EqTh c3 -> EquationalTheory of a2 means :Def15: :: MSUALG_6:def 15
( a3 c= a4 & ( for b1 being EquationalTheory of a2 holds
( a3 c= b1 implies a4 c= b1 ) ) );
uniqueness
for b1, b2 being EquationalTheory of c2 holds
( c3 c= b1 & ( for b3 being EquationalTheory of c2 holds
( c3 c= b3 implies b1 c= b3 ) ) & c3 c= b2 & ( for b3 being EquationalTheory of c2 holds
( c3 c= b3 implies b2 c= b3 ) ) implies b1 = b2 )
proof end;
existence
ex b1 being EquationalTheory of c2 st
( c3 c= b1 & ( for b2 being EquationalTheory of c2 holds
( c3 c= b2 implies b1 c= b2 ) ) )
proof end;
end;

:: deftheorem Def15 defines EqTh MSUALG_6:def 15 :
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being ManySortedRelation of the Sorts of b2
for b4 being EquationalTheory of b2 holds
( b4 = EqTh b3 iff ( b3 c= b4 & ( for b5 being EquationalTheory of b2 holds
( b3 c= b5 implies b4 c= b5 ) ) ) );

theorem Th52: :: MSUALG_6:52
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being ManySortedRelation of b2 holds
( EqCl b3,b2 c= EqTh b3 & InvCl b3 c= EqTh b3 & StabCl b3 c= EqTh b3 & TRS b3 c= EqTh b3 )
proof end;

theorem Th53: :: MSUALG_6:53
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being ManySortedRelation of b2
for b4 being SortSymbol of b1
for b5, b6 being Element of b2,b4 holds
( [b5,b6] in (EqTh b3) . b4 iff b5,b6 are_convertible_wrt (TRS b3) . b4 )
proof end;

theorem Th54: :: MSUALG_6:54
for b1 being non empty non void ManySortedSign
for b2 being non-empty MSAlgebra of b1
for b3 being ManySortedRelation of b2 holds EqTh b3 = EqCl (TRS b3),b2
proof end;