begin
:: deftheorem Def1 defines pr1 PRALG_1:def 1 :
for A, B being non empty set
for h1 being FinSequence of [:A,B:]
for b4 being FinSequence of A holds
( b4 = pr1 h1 iff ( len b4 = len h1 & ( for n being Nat st n in dom b4 holds
b4 . n = (h1 . n) `1 ) ) );
:: deftheorem Def2 defines pr2 PRALG_1:def 2 :
for A, B being non empty set
for h1 being FinSequence of [:A,B:]
for b4 being FinSequence of B holds
( b4 = pr2 h1 iff ( len b4 = len h1 & ( for n being Nat st n in dom b4 holds
b4 . n = (h1 . n) `2 ) ) );
definition
let A,
B be non
empty set ;
let f1 be non
empty homogeneous quasi_total PartFunc of
(A *),
A;
let f2 be non
empty homogeneous quasi_total PartFunc of
(B *),
B;
assume A1:
arity f1 = arity f2
;
func [[:f1,f2:]] -> non
empty homogeneous quasi_total PartFunc of
([:A,B:] *),
[:A,B:] means :
Def3:
(
dom it = (arity f1) -tuples_on [:A,B:] & ( for
h being
FinSequence of
[:A,B:] st
h in dom it holds
it . h = [(f1 . (pr1 h)),(f2 . (pr2 h))] ) );
existence
ex b1 being non empty homogeneous quasi_total PartFunc of ([:A,B:] *),[:A,B:] st
( dom b1 = (arity f1) -tuples_on [:A,B:] & ( for h being FinSequence of [:A,B:] st h in dom b1 holds
b1 . h = [(f1 . (pr1 h)),(f2 . (pr2 h))] ) )
uniqueness
for b1, b2 being non empty homogeneous quasi_total PartFunc of ([:A,B:] *),[:A,B:] st dom b1 = (arity f1) -tuples_on [:A,B:] & ( for h being FinSequence of [:A,B:] st h in dom b1 holds
b1 . h = [(f1 . (pr1 h)),(f2 . (pr2 h))] ) & dom b2 = (arity f1) -tuples_on [:A,B:] & ( for h being FinSequence of [:A,B:] st h in dom b2 holds
b2 . h = [(f1 . (pr1 h)),(f2 . (pr2 h))] ) holds
b1 = b2
end;
:: deftheorem Def3 defines [[: PRALG_1:def 3 :
for A, B being non empty set
for f1 being non empty homogeneous quasi_total PartFunc of (A *),A
for f2 being non empty homogeneous quasi_total PartFunc of (B *),B st arity f1 = arity f2 holds
for b5 being non empty homogeneous quasi_total PartFunc of ([:A,B:] *),[:A,B:] holds
( b5 = [[:f1,f2:]] iff ( dom b5 = (arity f1) -tuples_on [:A,B:] & ( for h being FinSequence of [:A,B:] st h in dom b5 holds
b5 . h = [(f1 . (pr1 h)),(f2 . (pr2 h))] ) ) );
definition
let U1,
U2 be
Universal_Algebra;
assume A1:
U1,
U2 are_similar
;
func Opers (
U1,
U2)
-> PFuncFinSequence of
[: the carrier of U1, the carrier of U2:] means :
Def4:
(
len it = len the
charact of
U1 & ( for
n being
Nat st
n in dom it holds
for
h1 being non
empty homogeneous quasi_total PartFunc of
( the carrier of U1 *), the
carrier of
U1 for
h2 being non
empty homogeneous quasi_total PartFunc of
( the carrier of U2 *), the
carrier of
U2 st
h1 = the
charact of
U1 . n &
h2 = the
charact of
U2 . n holds
it . n = [[:h1,h2:]] ) );
existence
ex b1 being PFuncFinSequence of [: the carrier of U1, the carrier of U2:] st
( len b1 = len the charact of U1 & ( for n being Nat st n in dom b1 holds
for h1 being non empty homogeneous quasi_total PartFunc of ( the carrier of U1 *), the carrier of U1
for h2 being non empty homogeneous quasi_total PartFunc of ( the carrier of U2 *), the carrier of U2 st h1 = the charact of U1 . n & h2 = the charact of U2 . n holds
b1 . n = [[:h1,h2:]] ) )
uniqueness
for b1, b2 being PFuncFinSequence of [: the carrier of U1, the carrier of U2:] st len b1 = len the charact of U1 & ( for n being Nat st n in dom b1 holds
for h1 being non empty homogeneous quasi_total PartFunc of ( the carrier of U1 *), the carrier of U1
for h2 being non empty homogeneous quasi_total PartFunc of ( the carrier of U2 *), the carrier of U2 st h1 = the charact of U1 . n & h2 = the charact of U2 . n holds
b1 . n = [[:h1,h2:]] ) & len b2 = len the charact of U1 & ( for n being Nat st n in dom b2 holds
for h1 being non empty homogeneous quasi_total PartFunc of ( the carrier of U1 *), the carrier of U1
for h2 being non empty homogeneous quasi_total PartFunc of ( the carrier of U2 *), the carrier of U2 st h1 = the charact of U1 . n & h2 = the charact of U2 . n holds
b2 . n = [[:h1,h2:]] ) holds
b1 = b2
end;
:: deftheorem Def4 defines Opers PRALG_1:def 4 :
for U1, U2 being Universal_Algebra st U1,U2 are_similar holds
for b3 being PFuncFinSequence of [: the carrier of U1, the carrier of U2:] holds
( b3 = Opers (U1,U2) iff ( len b3 = len the charact of U1 & ( for n being Nat st n in dom b3 holds
for h1 being non empty homogeneous quasi_total PartFunc of ( the carrier of U1 *), the carrier of U1
for h2 being non empty homogeneous quasi_total PartFunc of ( the carrier of U2 *), the carrier of U2 st h1 = the charact of U1 . n & h2 = the charact of U2 . n holds
b3 . n = [[:h1,h2:]] ) ) );
theorem
canceled;
theorem Th2:
definition
let U1,
U2 be
Universal_Algebra;
assume A1:
U1,
U2 are_similar
;
func [:U1,U2:] -> strict Universal_Algebra equals :
Def5:
UAStr(#
[: the carrier of U1, the carrier of U2:],
(Opers (U1,U2)) #);
coherence
UAStr(# [: the carrier of U1, the carrier of U2:],(Opers (U1,U2)) #) is strict Universal_Algebra
by A1, Th2;
end;
:: deftheorem Def5 defines [: PRALG_1:def 5 :
for U1, U2 being Universal_Algebra st U1,U2 are_similar holds
[:U1,U2:] = UAStr(# [: the carrier of U1, the carrier of U2:],(Opers (U1,U2)) #);
definition
let A,
B be non
empty set ;
func Inv (
A,
B)
-> Function of
[:A,B:],
[:B,A:] means :
Def6:
for
a being
Element of
[:A,B:] holds
it . a = [(a `2),(a `1)];
existence
ex b1 being Function of [:A,B:],[:B,A:] st
for a being Element of [:A,B:] holds b1 . a = [(a `2),(a `1)]
uniqueness
for b1, b2 being Function of [:A,B:],[:B,A:] st ( for a being Element of [:A,B:] holds b1 . a = [(a `2),(a `1)] ) & ( for a being Element of [:A,B:] holds b2 . a = [(a `2),(a `1)] ) holds
b1 = b2
end;
:: deftheorem Def6 defines Inv PRALG_1:def 6 :
for A, B being non empty set
for b3 being Function of [:A,B:],[:B,A:] holds
( b3 = Inv (A,B) iff for a being Element of [:A,B:] holds b3 . a = [(a `2),(a `1)] );
theorem
theorem
theorem
theorem Th6:
theorem
theorem
begin
:: deftheorem Def7 defines TrivialOp PRALG_1:def 7 :
for k being natural number
for b2 being PartFunc of ({{}} *),{{}} holds
( b2 = TrivialOp k iff ( dom b2 = {(k |-> {})} & rng b2 = {{}} ) );
theorem
:: deftheorem Def8 defines TrivialOps PRALG_1:def 8 :
for f being FinSequence of NAT
for b2 being PFuncFinSequence of {{}} holds
( b2 = TrivialOps f iff ( len b2 = len f & ( for n being Nat st n in dom b2 holds
for m being Nat st m = f . n holds
b2 . n = TrivialOp m ) ) );
theorem Th10:
theorem Th11:
:: deftheorem defines Trivial_Algebra PRALG_1:def 9 :
for f being non empty FinSequence of NAT holds Trivial_Algebra f = UAStr(# {{}},(TrivialOps f) #);
begin
:: deftheorem Def10 defines Univ_Alg-yielding PRALG_1:def 10 :
for IT being Function holds
( IT is Univ_Alg-yielding iff for x being set st x in dom IT holds
IT . x is Universal_Algebra );
:: deftheorem Def11 defines 1-sorted-yielding PRALG_1:def 11 :
for IT being Function holds
( IT is 1-sorted-yielding iff for x being set st x in dom IT holds
IT . x is 1-sorted );
:: deftheorem Def12 defines equal-signature PRALG_1:def 12 :
for IT being Function holds
( IT is equal-signature iff for x, y being set st x in dom IT & y in dom IT holds
for U1, U2 being Universal_Algebra st U1 = IT . x & U2 = IT . y holds
signature U1 = signature U2 );
:: deftheorem Def13 defines Carrier PRALG_1:def 13 :
for J being set
for A being 1-sorted-yielding ManySortedSet of J
for b3 being ManySortedSet of J holds
( b3 = Carrier A iff for j being set st j in J holds
ex R being 1-sorted st
( R = A . j & b3 . j = the carrier of R ) );
:: deftheorem Def14 defines ComSign PRALG_1:def 14 :
for J being non empty set
for A being Univ_Alg-yielding equal-signature ManySortedSet of J
for b3 being FinSequence of NAT holds
( b3 = ComSign A iff for j being Element of J holds b3 = signature (A . j) );
:: deftheorem Def15 defines ManySortedOperation PRALG_1:def 15 :
for J being non empty set
for B being V2() ManySortedSet of J
for b3 being ManySortedFunction of J holds
( b3 is ManySortedOperation of B iff for j being Element of J holds b3 . j is non empty homogeneous quasi_total PartFunc of ((B . j) *),(B . j) );
:: deftheorem Def16 defines equal-arity PRALG_1:def 16 :
for IT being Function holds
( IT is equal-arity iff for x, y being set st x in dom IT & y in dom IT holds
for f, g being Function st IT . x = f & IT . y = g holds
for n, m being Nat
for X, Y being non empty set st dom f = n -tuples_on X & dom g = m -tuples_on Y holds
for o1 being non empty homogeneous quasi_total PartFunc of (X *),X
for o2 being non empty homogeneous quasi_total PartFunc of (Y *),Y st f = o1 & g = o2 holds
arity o1 = arity o2 );
theorem Th12:
:: deftheorem Def17 defines .. PRALG_1:def 17 :
for F being Function-yielding Function
for f, b3 being Function holds
( b3 = F .. f iff ( dom b3 = dom F & ( for x being set st x in dom F holds
b3 . x = (F . x) . (f . x) ) ) );
:: deftheorem Def18 defines .. PRALG_1:def 18 :
for I being set
for f being ManySortedFunction of I
for x being ManySortedSet of I
for b4 being Function holds
( b4 = f .. x iff ( dom b4 = I & ( for i being set st i in I holds
for g being Function st g = f . i holds
b4 . i = g . (x . i) ) ) );
:: deftheorem Def19 defines ComAr PRALG_1:def 19 :
for J being non empty set
for B being V2() ManySortedSet of J
for O being equal-arity ManySortedOperation of B
for b4 being Nat holds
( b4 = ComAr O iff for j being Element of J holds b4 = arity (O . j) );
:: deftheorem Def20 defines EmptySeq PRALG_1:def 20 :
for I being set
for A, b3 being ManySortedSet of I holds
( b3 = EmptySeq A iff for i being set st i in I holds
b3 . i = {} (A . i) );
:: deftheorem defines [[: PRALG_1:def 21 :
for J being non empty set
for B being V2() ManySortedSet of J
for O being equal-arity ManySortedOperation of B
for b4 being non empty homogeneous quasi_total PartFunc of ((product B) *),(product B) holds
( b4 = [[:O:]] iff ( dom b4 = (ComAr O) -tuples_on (product B) & ( for p being Element of (product B) * st p in dom b4 holds
( ( dom p = {} implies b4 . p = O .. (EmptySeq B) ) & ( dom p <> {} implies for Z being non empty set
for w being ManySortedSet of [:J,Z:] st Z = dom p & w = ~ (uncurry p) holds
b4 . p = O .. (curry w) ) ) ) ) );
:: deftheorem defines ProdOp PRALG_1:def 22 :
for J being non empty set
for A being Univ_Alg-yielding equal-signature ManySortedSet of J
for n being natural number st n in dom (ComSign A) holds
for b4 being equal-arity ManySortedOperation of Carrier A holds
( b4 = ProdOp (A,n) iff for j being Element of J
for o being operation of (A . j) st the charact of (A . j) . n = o holds
b4 . j = o );
:: deftheorem Def23 defines ProdOpSeq PRALG_1:def 23 :
for J being non empty set
for A being Univ_Alg-yielding equal-signature ManySortedSet of J
for b3 being PFuncFinSequence of (product (Carrier A)) holds
( b3 = ProdOpSeq A iff ( len b3 = len (ComSign A) & ( for n being Nat st n in dom b3 holds
b3 . n = [[:(ProdOp (A,n)):]] ) ) );
:: deftheorem defines ProdUnivAlg PRALG_1:def 24 :
for J being non empty set
for A being Univ_Alg-yielding equal-signature ManySortedSet of J holds ProdUnivAlg A = UAStr(# (product (Carrier A)),(ProdOpSeq A) #);