:: The Correspondence Between Homomorphisms of Universal Algebra & Many Sorted Algebra
:: by Adam Grabowski
::
:: Received December 13, 1994
:: Copyright (c) 1994-2021 Association of Mizar Users


theorem Th1: :: MSUHOM_1:1
for f, g being Function
for C being set st rng f c= C holds
(g | C) * f = g * f
proof end;

theorem Th2: :: MSUHOM_1:2
for I being set
for C being Subset of I holds C * c= I *
proof end;

theorem :: MSUHOM_1:3
for f being Function
for C being set st f is Function-yielding holds
f | C is Function-yielding ;

theorem Th4: :: MSUHOM_1:4
for I being set
for C being Subset of I
for M being ManySortedSet of I holds (M | C) # = (M #) | (C *)
proof end;

definition
let S, S9 be non empty ManySortedSign ;
pred S <= S9 means :: MSUHOM_1:def 1
( the carrier of S c= the carrier of S9 & the carrier' of S c= the carrier' of S9 & the Arity of S9 | the carrier' of S = the Arity of S & the ResultSort of S9 | the carrier' of S = the ResultSort of S );
reflexivity
for S being non empty ManySortedSign holds
( the carrier of S c= the carrier of S & the carrier' of S c= the carrier' of S & the Arity of S | the carrier' of S = the Arity of S & the ResultSort of S | the carrier' of S = the ResultSort of S )
;
end;

:: deftheorem defines <= MSUHOM_1:def 1 :
for S, S9 being non empty ManySortedSign holds
( S <= S9 iff ( the carrier of S c= the carrier of S9 & the carrier' of S c= the carrier' of S9 & the Arity of S9 | the carrier' of S = the Arity of S & the ResultSort of S9 | the carrier' of S = the ResultSort of S ) );

theorem :: MSUHOM_1:5
for S, S9, S99 being non empty ManySortedSign st S <= S9 & S9 <= S99 holds
S <= S99
proof end;

theorem :: MSUHOM_1:6
for S, S9 being non empty strict ManySortedSign st S <= S9 & S9 <= S holds
S = S9
proof end;

theorem :: MSUHOM_1:7
for n being Nat
for A being non empty set
for g being Function
for a being Element of A
for k being Nat st 1 <= k & k <= n holds
(a .--> g) . ((n |-> a) /. k) = g
proof end;

theorem Th8: :: MSUHOM_1:8
for I being set
for I0 being Subset of I
for A, B being ManySortedSet of I
for F being ManySortedFunction of A,B
for A0, B0 being ManySortedSet of I0 st A0 = A | I0 & B0 = B | I0 holds
F | I0 is ManySortedFunction of A0,B0
proof end;

definition
let S, S9 be non empty non void strict ManySortedSign ;
let A be strict non-empty MSAlgebra over S9;
assume A1: S <= S9 ;
func A Over S -> strict non-empty MSAlgebra over S means :Def2: :: MSUHOM_1:def 2
( the Sorts of it = the Sorts of A | the carrier of S & the Charact of it = the Charact of A | the carrier' of S );
existence
ex b1 being strict non-empty MSAlgebra over S st
( the Sorts of b1 = the Sorts of A | the carrier of S & the Charact of b1 = the Charact of A | the carrier' of S )
proof end;
uniqueness
for b1, b2 being strict non-empty MSAlgebra over S st the Sorts of b1 = the Sorts of A | the carrier of S & the Charact of b1 = the Charact of A | the carrier' of S & the Sorts of b2 = the Sorts of A | the carrier of S & the Charact of b2 = the Charact of A | the carrier' of S holds
b1 = b2
;
end;

:: deftheorem Def2 defines Over MSUHOM_1:def 2 :
for S, S9 being non empty non void strict ManySortedSign
for A being strict non-empty MSAlgebra over S9 st S <= S9 holds
for b4 being strict non-empty MSAlgebra over S holds
( b4 = A Over S iff ( the Sorts of b4 = the Sorts of A | the carrier of S & the Charact of b4 = the Charact of A | the carrier' of S ) );

theorem Th9: :: MSUHOM_1:9
for S being non empty non void strict ManySortedSign
for A being strict non-empty MSAlgebra over S holds A = A Over S
proof end;

theorem Th10: :: MSUHOM_1:10
for U1, U2 being Universal_Algebra st U1,U2 are_similar holds
MSSign U1 = MSSign U2
proof end;

definition
let U1, U2 be Universal_Algebra;
let h be Function of U1,U2;
assume A1: MSSign U1 = MSSign U2 ;
func MSAlg h -> ManySortedFunction of (MSAlg U1),((MSAlg U2) Over (MSSign U1)) equals :Def3: :: MSUHOM_1:def 3
0 .--> h;
coherence
0 .--> h is ManySortedFunction of (MSAlg U1),((MSAlg U2) Over (MSSign U1))
proof end;
end;

:: deftheorem Def3 defines MSAlg MSUHOM_1:def 3 :
for U1, U2 being Universal_Algebra
for h being Function of U1,U2 st MSSign U1 = MSSign U2 holds
MSAlg h = 0 .--> h;

theorem Th11: :: MSUHOM_1:11
for U1, U2 being Universal_Algebra
for h being Function of U1,U2 st U1,U2 are_similar holds
for o being OperSymbol of (MSSign U1) holds (MSAlg h) . (the_result_sort_of o) = h
proof end;

theorem Th12: :: MSUHOM_1:12
for U1 being Universal_Algebra
for o being OperSymbol of (MSSign U1) holds Den (o,(MSAlg U1)) = the charact of U1 . o
proof end;

Lm1: for U1 being Universal_Algebra holds dom (signature U1) = dom the charact of U1
proof end;

theorem Th13: :: MSUHOM_1:13
for U1 being Universal_Algebra
for o being OperSymbol of (MSSign U1) holds Den (o,(MSAlg U1)) is operation of U1
proof end;

Lm2: for U1, U2 being Universal_Algebra st U1,U2 are_similar holds
for o being OperSymbol of (MSSign U1) holds Den (o,((MSAlg U2) Over (MSSign U1))) is operation of U2

proof end;

theorem Th14: :: MSUHOM_1:14
for U1 being Universal_Algebra
for o being OperSymbol of (MSSign U1)
for y being Element of Args (o,(MSAlg U1)) holds y is FinSequence of the carrier of U1
proof end;

theorem Th15: :: MSUHOM_1:15
for U1, U2 being Universal_Algebra
for h being Function of U1,U2 st U1,U2 are_similar holds
for o being OperSymbol of (MSSign U1)
for y being Element of Args (o,(MSAlg U1)) holds (MSAlg h) # y = h * y
proof end;

theorem Th16: :: MSUHOM_1:16
for U1, U2 being Universal_Algebra
for h being Function of U1,U2 st h is_homomorphism holds
MSAlg h is_homomorphism MSAlg U1,(MSAlg U2) Over (MSSign U1)
proof end;

Lm3: for U1 being Universal_Algebra
for n being Nat st n in dom the charact of U1 holds
n is OperSymbol of (MSSign U1)

proof end;

theorem Th17: :: MSUHOM_1:17
for U1, U2 being Universal_Algebra
for h being Function of U1,U2 st U1,U2 are_similar holds
MSAlg h is ManySortedSet of {0}
proof end;

theorem Th18: :: MSUHOM_1:18
for U1, U2 being Universal_Algebra
for h being Function of U1,U2 st h is_epimorphism holds
MSAlg h is_epimorphism MSAlg U1,(MSAlg U2) Over (MSSign U1)
proof end;

theorem Th19: :: MSUHOM_1:19
for U1, U2 being Universal_Algebra
for h being Function of U1,U2 st h is_monomorphism holds
MSAlg h is_monomorphism MSAlg U1,(MSAlg U2) Over (MSSign U1)
proof end;

theorem :: MSUHOM_1:20
for U1, U2 being Universal_Algebra
for h being Function of U1,U2 st h is_isomorphism holds
MSAlg h is_isomorphism MSAlg U1,(MSAlg U2) Over (MSSign U1) by Th18, Th19;

theorem Th21: :: MSUHOM_1:21
for U1, U2 being Universal_Algebra
for h being Function of U1,U2 st U1,U2 are_similar & MSAlg h is_homomorphism MSAlg U1,(MSAlg U2) Over (MSSign U1) holds
h is_homomorphism
proof end;

theorem Th22: :: MSUHOM_1:22
for U1, U2 being Universal_Algebra
for h being Function of U1,U2 st U1,U2 are_similar & MSAlg h is_epimorphism MSAlg U1,(MSAlg U2) Over (MSSign U1) holds
h is_epimorphism
proof end;

theorem Th23: :: MSUHOM_1:23
for U1, U2 being Universal_Algebra
for h being Function of U1,U2 st U1,U2 are_similar & MSAlg h is_monomorphism MSAlg U1,(MSAlg U2) Over (MSSign U1) holds
h is_monomorphism
proof end;

theorem :: MSUHOM_1:24
for U1, U2 being Universal_Algebra
for h being Function of U1,U2 st U1,U2 are_similar & MSAlg h is_isomorphism MSAlg U1,(MSAlg U2) Over (MSSign U1) holds
h is_isomorphism by Th23, Th22;

theorem :: MSUHOM_1:25
for U1 being Universal_Algebra holds MSAlg (id the carrier of U1) = id the Sorts of (MSAlg U1)
proof end;

theorem :: MSUHOM_1:26
for U1, U2, U3 being Universal_Algebra st U1,U2 are_similar & U2,U3 are_similar holds
for h1 being Function of U1,U2
for h2 being Function of U2,U3 holds (MSAlg h2) ** (MSAlg h1) = MSAlg (h2 * h1)
proof end;