:: Binary Operations on Finite Sequences
:: by Wojciech A. Trybulec
::
:: Received August 10, 1990
:: Copyright (c) 1990-2021 Association of Mizar Users


definition
let D be non empty set ;
let F be FinSequence of D;
let g be BinOp of D;
assume A1: ( g is having_a_unity or len F >= 1 ) ;
func g "**" F -> Element of D means :Def1: :: FINSOP_1:def 1
it = the_unity_wrt g if ( g is having_a_unity & len F = 0 )
otherwise ex f being sequence of D st
( f . 1 = F . 1 & ( for n being Nat st 0 <> n & n < len F holds
f . (n + 1) = g . ((f . n),(F . (n + 1))) ) & it = f . (len F) );
existence
( ( g is having_a_unity & len F = 0 implies ex b1 being Element of D st b1 = the_unity_wrt g ) & ( ( not g is having_a_unity or not len F = 0 ) implies ex b1 being Element of D ex f being sequence of D st
( f . 1 = F . 1 & ( for n being Nat st 0 <> n & n < len F holds
f . (n + 1) = g . ((f . n),(F . (n + 1))) ) & b1 = f . (len F) ) ) )
proof end;
uniqueness
for b1, b2 being Element of D holds
( ( g is having_a_unity & len F = 0 & b1 = the_unity_wrt g & b2 = the_unity_wrt g implies b1 = b2 ) & ( ( not g is having_a_unity or not len F = 0 ) & ex f being sequence of D st
( f . 1 = F . 1 & ( for n being Nat st 0 <> n & n < len F holds
f . (n + 1) = g . ((f . n),(F . (n + 1))) ) & b1 = f . (len F) ) & ex f being sequence of D st
( f . 1 = F . 1 & ( for n being Nat st 0 <> n & n < len F holds
f . (n + 1) = g . ((f . n),(F . (n + 1))) ) & b2 = f . (len F) ) implies b1 = b2 ) )
proof end;
consistency
for b1 being Element of D holds verum
;
end;

:: deftheorem Def1 defines "**" FINSOP_1:def 1 :
for D being non empty set
for F being FinSequence of D
for g being BinOp of D st ( g is having_a_unity or len F >= 1 ) holds
for b4 being Element of D holds
( ( g is having_a_unity & len F = 0 implies ( b4 = g "**" F iff b4 = the_unity_wrt g ) ) & ( ( not g is having_a_unity or not len F = 0 ) implies ( b4 = g "**" F iff ex f being sequence of D st
( f . 1 = F . 1 & ( for n being Nat st 0 <> n & n < len F holds
f . (n + 1) = g . ((f . n),(F . (n + 1))) ) & b4 = f . (len F) ) ) ) );

theorem :: FINSOP_1:1
for D being non empty set
for F being FinSequence of D
for g being BinOp of D st len F >= 1 holds
ex f being sequence of D st
( f . 1 = F . 1 & ( for n being Nat st 0 <> n & n < len F holds
f . (n + 1) = g . ((f . n),(F . (n + 1))) ) & g "**" F = f . (len F) ) by Def1;

theorem :: FINSOP_1:2
for D being non empty set
for d being Element of D
for F being FinSequence of D
for g being BinOp of D st len F >= 1 & ex f being sequence of D st
( f . 1 = F . 1 & ( for n being Nat st 0 <> n & n < len F holds
f . (n + 1) = g . ((f . n),(F . (n + 1))) ) & d = f . (len F) ) holds
d = g "**" F by Def1;

definition
let A be non empty set ;
let F be sequence of A;
let p be FinSequence of A;
:: original: +*
redefine func F +* p -> sequence of A;
coherence
F +* p is sequence of A
proof end;
end;

notation
let f be FinSequence;
synonym findom f for dom f;
end;

definition
let f be FinSequence;
:: original: findom
redefine func findom f -> Element of Fin NAT;
coherence
findom f is Element of Fin NAT
proof end;
end;

Lm1: for D being non empty set
for F being FinSequence of D
for g being BinOp of D st len F >= 1 & g is associative & g is commutative holds
g "**" F = g $$ ((findom F),((NAT --> (the_unity_wrt g)) +* F))

proof end;

Lm2: for D being non empty set
for F being FinSequence of D
for g being BinOp of D st len F = 0 & g is having_a_unity & g is associative & g is commutative holds
g "**" F = g $$ ((findom F),((NAT --> (the_unity_wrt g)) +* F))

proof end;

theorem :: FINSOP_1:3
for D being non empty set
for F being FinSequence of D
for g being BinOp of D st ( g is having_a_unity or len F >= 1 ) & g is associative & g is commutative holds
g "**" F = g $$ ((findom F),((NAT --> (the_unity_wrt g)) +* F))
proof end;

Lm3: for D being non empty set
for g being BinOp of D st g is having_a_unity holds
g "**" (<*> D) = the_unity_wrt g

proof end;

Lm4: for D being non empty set
for d being Element of D
for g being BinOp of D holds g "**" <*d*> = d

proof end;

Lm5: for D being non empty set
for d being Element of D
for F being FinSequence of D
for g being BinOp of D st len F >= 1 holds
g "**" (F ^ <*d*>) = g . ((g "**" F),d)

proof end;

Lm6: for D being non empty set
for d being Element of D
for F being FinSequence of D
for g being BinOp of D st g is having_a_unity & len F = 0 holds
g "**" (F ^ <*d*>) = g . ((g "**" F),d)

proof end;

theorem Th4: :: FINSOP_1:4
for D being non empty set
for d being Element of D
for F being FinSequence of D
for g being BinOp of D st ( g is having_a_unity or len F >= 1 ) holds
g "**" (F ^ <*d*>) = g . ((g "**" F),d)
proof end;

theorem Th5: :: FINSOP_1:5
for D being non empty set
for F, G being FinSequence of D
for g being BinOp of D st g is associative & ( g is having_a_unity or ( len F >= 1 & len G >= 1 ) ) holds
g "**" (F ^ G) = g . ((g "**" F),(g "**" G))
proof end;

theorem Th6: :: FINSOP_1:6
for D being non empty set
for d being Element of D
for F being FinSequence of D
for g being BinOp of D st g is associative & ( g is having_a_unity or len F >= 1 ) holds
g "**" (<*d*> ^ F) = g . (d,(g "**" F))
proof end;

Lm7: for D being non empty set
for F, G being FinSequence of D
for g being BinOp of D st g is associative & g is commutative holds
for f being Permutation of (dom F) st len F >= 1 & len F = len G & ( for i being Nat st i in dom G holds
G . i = F . (f . i) ) holds
g "**" F = g "**" G

proof end;

Lm8: for D being non empty set
for F, G being FinSequence of D
for g being BinOp of D
for P being Permutation of (dom F) st g is having_a_unity & len F = 0 & G = F * P holds
g "**" F = g "**" G

proof end;

theorem Th7: :: FINSOP_1:7
for D being non empty set
for F, G being FinSequence of D
for g being BinOp of D
for P being Permutation of (dom F) st g is commutative & g is associative & ( g is having_a_unity or len F >= 1 ) & G = F * P holds
g "**" F = g "**" G
proof end;

Lm9: for D being non empty set
for F, G being FinSequence of D
for g being BinOp of D st g is associative & g is commutative & F is one-to-one & G is one-to-one & rng F = rng G & len F >= 1 holds
g "**" F = g "**" G

proof end;

Lm10: for D being non empty set
for F, G being FinSequence of D
for g being BinOp of D st len F = 0 & g is having_a_unity & F is one-to-one & G is one-to-one & rng F = rng G holds
g "**" F = g "**" G

proof end;

theorem :: FINSOP_1:8
for D being non empty set
for F, G being FinSequence of D
for g being BinOp of D st ( g is having_a_unity or len F >= 1 ) & g is associative & g is commutative & F is one-to-one & G is one-to-one & rng F = rng G holds
g "**" F = g "**" G
proof end;

Lm11: for D being non empty set
for F being FinSequence of D
for g being BinOp of D st len F = 1 holds
g "**" F = F . 1

proof end;

Lm12: for D being non empty set
for F, G, H being FinSequence of D
for g being BinOp of D st g is associative & g is commutative & len F >= 1 & len F = len G & len F = len H & ( for k being Nat st k in dom F holds
H . k = g . ((F . k),(G . k)) ) holds
g "**" H = g . ((g "**" F),(g "**" G))

proof end;

Lm13: for D being non empty set
for F, G, H being FinSequence of D
for g being BinOp of D st g is having_a_unity & len F = 0 & len F = len G & len F = len H holds
g "**" F = g . ((g "**" G),(g "**" H))

proof end;

theorem :: FINSOP_1:9
for D being non empty set
for F, G, H being FinSequence of D
for g being BinOp of D st g is associative & g is commutative & ( g is having_a_unity or len F >= 1 ) & len F = len G & len F = len H & ( for k being Nat st k in dom F holds
F . k = g . ((G . k),(H . k)) ) holds
g "**" F = g . ((g "**" G),(g "**" H))
proof end;

theorem :: FINSOP_1:10
for D being non empty set
for g being BinOp of D st g is having_a_unity holds
g "**" (<*> D) = the_unity_wrt g by Lm3;

theorem :: FINSOP_1:11
for D being non empty set
for d being Element of D
for g being BinOp of D holds g "**" <*d*> = d by Lm4;

theorem Th12: :: FINSOP_1:12
for D being non empty set
for d1, d2 being Element of D
for g being BinOp of D holds g "**" <*d1,d2*> = g . (d1,d2)
proof end;

theorem :: FINSOP_1:13
for D being non empty set
for d1, d2 being Element of D
for g being BinOp of D st g is commutative holds
g "**" <*d1,d2*> = g "**" <*d2,d1*>
proof end;

theorem Th14: :: FINSOP_1:14
for D being non empty set
for d1, d2, d3 being Element of D
for g being BinOp of D holds g "**" <*d1,d2,d3*> = g . ((g . (d1,d2)),d3)
proof end;

theorem :: FINSOP_1:15
for D being non empty set
for d1, d2, d3 being Element of D
for g being BinOp of D st g is commutative holds
g "**" <*d1,d2,d3*> = g "**" <*d2,d1,d3*>
proof end;

theorem Th16: :: FINSOP_1:16
for D being non empty set
for d being Element of D
for g being BinOp of D holds g "**" (1 |-> d) = d
proof end;

theorem :: FINSOP_1:17
for D being non empty set
for d being Element of D
for g being BinOp of D holds g "**" (2 |-> d) = g . (d,d)
proof end;

theorem Th18: :: FINSOP_1:18
for D being non empty set
for d being Element of D
for g being BinOp of D
for k, l being Nat st g is associative & ( g is having_a_unity or ( k <> 0 & l <> 0 ) ) holds
g "**" ((k + l) |-> d) = g . ((g "**" (k |-> d)),(g "**" (l |-> d)))
proof end;

theorem :: FINSOP_1:19
for D being non empty set
for d being Element of D
for g being BinOp of D
for k, l being Nat st g is associative & ( g is having_a_unity or ( k <> 0 & l <> 0 ) ) holds
g "**" ((k * l) |-> d) = g "**" (l |-> (g "**" (k |-> d)))
proof end;

theorem :: FINSOP_1:20
for D being non empty set
for F being FinSequence of D
for g being BinOp of D st len F = 1 holds
g "**" F = F . 1 by Lm11;

theorem :: FINSOP_1:21
for D being non empty set
for F being FinSequence of D
for g being BinOp of D st len F = 2 holds
g "**" F = g . ((F . 1),(F . 2))
proof end;