:: Semigroup operations on finite subsets
:: by Czes{\l}aw Byli\'nski
::
:: Received May 4, 1990
:: Copyright (c) 1990-2021 Association of Mizar Users


Lm1: for i being Nat holds Seg i is Element of Fin NAT
by FINSUB_1:def 5;

Lm2: for i being Nat holds not i + 1 in Seg i
by FINSEQ_1:1, XREAL_1:29;

theorem Th1: :: SETWOP_2:1
for C, D being non empty set
for c1, c2 being Element of C
for F being BinOp of D
for f being Function of C,D st F is commutative & F is associative & c1 <> c2 holds
F $$ ({.c1,c2.},f) = F . ((f . c1),(f . c2))
proof end;

theorem Th2: :: SETWOP_2:2
for C, D being non empty set
for c being Element of C
for B being Element of Fin C
for F being BinOp of D
for f being Function of C,D st F is commutative & F is associative & ( B <> {} or F is having_a_unity ) & not c in B holds
F $$ ((B \/ {.c.}),f) = F . ((F $$ (B,f)),(f . c))
proof end;

theorem :: SETWOP_2:3
for C, D being non empty set
for c1, c2, c3 being Element of C
for F being BinOp of D
for f being Function of C,D st F is commutative & F is associative & c1 <> c2 & c1 <> c3 & c2 <> c3 holds
F $$ ({.c1,c2,c3.},f) = F . ((F . ((f . c1),(f . c2))),(f . c3))
proof end;

theorem :: SETWOP_2:4
for C, D being non empty set
for B1, B2 being Element of Fin C
for F being BinOp of D
for f being Function of C,D st F is commutative & F is associative & ( ( B1 <> {} & B2 <> {} ) or F is having_a_unity ) & B1 misses B2 holds
F $$ ((B1 \/ B2),f) = F . ((F $$ (B1,f)),(F $$ (B2,f)))
proof end;

theorem Th5: :: SETWOP_2:5
for C, C9, D being non empty set
for B being Element of Fin C
for A being Element of Fin C9
for F being BinOp of D
for f being Function of C,D
for g being Function of C9,D st F is commutative & F is associative & ( A <> {} or F is having_a_unity ) & ex s being Function st
( dom s = A & rng s = B & s is one-to-one & g | A = f * s ) holds
F $$ (A,g) = F $$ (B,f)
proof end;

theorem :: SETWOP_2:6
for C, D, E being non empty set
for B being Element of Fin C
for f being Function of C,D
for H being BinOp of E
for h being Function of D,E st H is commutative & H is associative & ( B <> {} or H is having_a_unity ) & f is one-to-one holds
H $$ ((f .: B),h) = H $$ (B,(h * f))
proof end;

theorem :: SETWOP_2:7
for C, D being non empty set
for B being Element of Fin C
for F being BinOp of D
for f, f9 being Function of C,D st F is commutative & F is associative & ( B <> {} or F is having_a_unity ) & f | B = f9 | B holds
F $$ (B,f) = F $$ (B,f9)
proof end;

theorem :: SETWOP_2:8
for C, D being non empty set
for B being Element of Fin C
for e being Element of D
for F being BinOp of D
for f being Function of C,D st F is commutative & F is associative & F is having_a_unity & e = the_unity_wrt F & f .: B = {e} holds
F $$ (B,f) = e
proof end;

theorem Th9: :: SETWOP_2:9
for C, D being non empty set
for B being Element of Fin C
for e being Element of D
for F, G being BinOp of D
for f, f9 being Function of C,D st F is commutative & F is associative & F is having_a_unity & e = the_unity_wrt F & G . (e,e) = e & ( for d1, d2, d3, d4 being Element of D holds F . ((G . (d1,d2)),(G . (d3,d4))) = G . ((F . (d1,d3)),(F . (d2,d4))) ) holds
G . ((F $$ (B,f)),(F $$ (B,f9))) = F $$ (B,(G .: (f,f9)))
proof end;

Lm3: for D being non empty set
for F being BinOp of D st F is commutative & F is associative holds
for d1, d2, d3, d4 being Element of D holds F . ((F . (d1,d2)),(F . (d3,d4))) = F . ((F . (d1,d3)),(F . (d2,d4)))

proof end;

theorem :: SETWOP_2:10
for C, D being non empty set
for B being Element of Fin C
for F being BinOp of D
for f, f9 being Function of C,D st F is commutative & F is associative & F is having_a_unity holds
F . ((F $$ (B,f)),(F $$ (B,f9))) = F $$ (B,(F .: (f,f9)))
proof end;

theorem :: SETWOP_2:11
for C, D being non empty set
for B being Element of Fin C
for F, G being BinOp of D
for f, f9 being Function of C,D st F is commutative & F is associative & F is having_a_unity & F is having_an_inverseOp & G = F * ((id D),(the_inverseOp_wrt F)) holds
G . ((F $$ (B,f)),(F $$ (B,f9))) = F $$ (B,(G .: (f,f9)))
proof end;

theorem Th12: :: SETWOP_2:12
for C, D being non empty set
for B being Element of Fin C
for d, e being Element of D
for F, G being BinOp of D
for f being Function of C,D st F is commutative & F is associative & F is having_a_unity & e = the_unity_wrt F & G is_distributive_wrt F & G . (d,e) = e holds
G . (d,(F $$ (B,f))) = F $$ (B,(G [;] (d,f)))
proof end;

theorem Th13: :: SETWOP_2:13
for C, D being non empty set
for B being Element of Fin C
for d, e being Element of D
for F, G being BinOp of D
for f being Function of C,D st F is commutative & F is associative & F is having_a_unity & e = the_unity_wrt F & G is_distributive_wrt F & G . (e,d) = e holds
G . ((F $$ (B,f)),d) = F $$ (B,(G [:] (f,d)))
proof end;

theorem :: SETWOP_2:14
for C, D being non empty set
for B being Element of Fin C
for d being Element of D
for F, G being BinOp of D
for f being Function of C,D st F is commutative & F is associative & F is having_a_unity & F is having_an_inverseOp & G is_distributive_wrt F holds
G . (d,(F $$ (B,f))) = F $$ (B,(G [;] (d,f)))
proof end;

theorem :: SETWOP_2:15
for C, D being non empty set
for B being Element of Fin C
for d being Element of D
for F, G being BinOp of D
for f being Function of C,D st F is commutative & F is associative & F is having_a_unity & F is having_an_inverseOp & G is_distributive_wrt F holds
G . ((F $$ (B,f)),d) = F $$ (B,(G [:] (f,d)))
proof end;

theorem Th16: :: SETWOP_2:16
for C, D, E being non empty set
for B being Element of Fin C
for F being BinOp of D
for f being Function of C,D
for H being BinOp of E
for h being Function of D,E st F is commutative & F is associative & F is having_a_unity & H is commutative & H is associative & H is having_a_unity & h . (the_unity_wrt F) = the_unity_wrt H & ( for d1, d2 being Element of D holds h . (F . (d1,d2)) = H . ((h . d1),(h . d2)) ) holds
h . (F $$ (B,f)) = H $$ (B,(h * f))
proof end;

theorem :: SETWOP_2:17
for C, D being non empty set
for B being Element of Fin C
for F being BinOp of D
for u being UnOp of D
for f being Function of C,D st F is commutative & F is associative & F is having_a_unity & u . (the_unity_wrt F) = the_unity_wrt F & u is_distributive_wrt F holds
u . (F $$ (B,f)) = F $$ (B,(u * f)) by Th16;

theorem :: SETWOP_2:18
for C, D being non empty set
for B being Element of Fin C
for d being Element of D
for F, G being BinOp of D
for f being Function of C,D st F is commutative & F is associative & F is having_a_unity & F is having_an_inverseOp & G is_distributive_wrt F holds
(G [;] (d,(id D))) . (F $$ (B,f)) = F $$ (B,((G [;] (d,(id D))) * f))
proof end;

theorem :: SETWOP_2:19
for C, D being non empty set
for B being Element of Fin C
for F being BinOp of D
for f being Function of C,D st F is commutative & F is associative & F is having_a_unity & F is having_an_inverseOp holds
(the_inverseOp_wrt F) . (F $$ (B,f)) = F $$ (B,((the_inverseOp_wrt F) * f))
proof end;

definition
let D be non empty set ;
let p be FinSequence of D;
let d be Element of D;
func [#] (p,d) -> sequence of D equals :: SETWOP_2:def 1
(NAT --> d) +* p;
coherence
(NAT --> d) +* p is sequence of D
;
end;

:: deftheorem defines [#] SETWOP_2:def 1 :
for D being non empty set
for p being FinSequence of D
for d being Element of D holds [#] (p,d) = (NAT --> d) +* p;

theorem Th20: :: SETWOP_2:20
for D being non empty set
for d being Element of D
for i being Nat
for p being FinSequence of D holds
( ( i in dom p implies ([#] (p,d)) . i = p . i ) & ( not i in dom p implies ([#] (p,d)) . i = d ) )
proof end;

theorem :: SETWOP_2:21
for D being non empty set
for d being Element of D
for p being FinSequence of D holds ([#] (p,d)) | (dom p) = p
proof end;

theorem :: SETWOP_2:22
for D being non empty set
for d being Element of D
for p, q being FinSequence of D holds ([#] ((p ^ q),d)) | (dom p) = p
proof end;

theorem :: SETWOP_2:23
for D being non empty set
for d being Element of D
for p being FinSequence of D holds rng ([#] (p,d)) = (rng p) \/ {d}
proof end;

theorem :: SETWOP_2:24
for D, E being non empty set
for d being Element of D
for h being Function of D,E
for p being FinSequence of D holds h * ([#] (p,d)) = [#] ((h * p),(h . d))
proof end;

Lm4: for D being non empty set
for e being Element of D
for F being BinOp of D
for p, q being FinSequence of D st len p = len q & F . (e,e) = e holds
F .: (([#] (p,e)),([#] (q,e))) = [#] ((F .: (p,q)),e)

proof end;

Lm5: for D being non empty set
for d, e being Element of D
for F being BinOp of D
for p being FinSequence of D st F . (e,d) = e holds
F [:] (([#] (p,e)),d) = [#] ((F [:] (p,d)),e)

proof end;

Lm6: for D being non empty set
for d, e being Element of D
for F being BinOp of D
for p being FinSequence of D st F . (d,e) = e holds
F [;] (d,([#] (p,e))) = [#] ((F [;] (d,p)),e)

proof end;

notation
let i be Nat;
synonym finSeg i for Seg i;
end;

definition
let i be Nat;
:: original: finSeg
redefine func finSeg i -> Element of Fin NAT;
coherence
finSeg i is Element of Fin NAT
by Lm1;
end;

notation
let D be non empty set ;
let p be FinSequence of D;
let F be BinOp of D;
synonym F $$ p for F "**" p;
end;

definition
let D be non empty set ;
let p be FinSequence of D;
let F be BinOp of D;
assume A1: ( ( F is having_a_unity or len p >= 1 ) & F is associative & F is commutative ) ;
redefine func F "**" p equals :Def2: :: SETWOP_2:def 2
F $$ ((findom p),([#] (p,(the_unity_wrt F))));
compatibility
for b1 being Element of D holds
( b1 = F $$ p iff b1 = F $$ ((findom p),([#] (p,(the_unity_wrt F)))) )
by A1, FINSOP_1:3;
end;

:: deftheorem Def2 defines $$ SETWOP_2:def 2 :
for D being non empty set
for p being FinSequence of D
for F being BinOp of D st ( F is having_a_unity or len p >= 1 ) & F is associative & F is commutative holds
F $$ p = F $$ ((findom p),([#] (p,(the_unity_wrt F))));

theorem Th25: :: SETWOP_2:25
for D being non empty set
for F being BinOp of D
for i being Nat st F is having_a_unity holds
F "**" (i |-> (the_unity_wrt F)) = the_unity_wrt F
proof end;

theorem Th26: :: SETWOP_2:26
for D being non empty set
for d being Element of D
for F being BinOp of D
for i, j being Nat st F is associative & ( ( i >= 1 & j >= 1 ) or F is having_a_unity ) holds
F "**" ((i + j) |-> d) = F . ((F "**" (i |-> d)),(F "**" (j |-> d)))
proof end;

theorem :: SETWOP_2:27
for D being non empty set
for d being Element of D
for F being BinOp of D
for i, j being Nat st F is commutative & F is associative & ( ( i >= 1 & j >= 1 ) or F is having_a_unity ) holds
F "**" ((i * j) |-> d) = F "**" (j |-> (F "**" (i |-> d)))
proof end;

theorem Th28: :: SETWOP_2:28
for D, E being non empty set
for F being BinOp of D
for H being BinOp of E
for h being Function of D,E
for p being FinSequence of D st F is having_a_unity & H is having_a_unity & h . (the_unity_wrt F) = the_unity_wrt H & ( for d1, d2 being Element of D holds h . (F . (d1,d2)) = H . ((h . d1),(h . d2)) ) holds
h . (F "**" p) = H "**" (h * p)
proof end;

theorem :: SETWOP_2:29
for D being non empty set
for F being BinOp of D
for u being UnOp of D
for p being FinSequence of D st F is having_a_unity & u . (the_unity_wrt F) = the_unity_wrt F & u is_distributive_wrt F holds
u . (F "**" p) = F "**" (u * p) by Th28;

theorem :: SETWOP_2:30
for D being non empty set
for d being Element of D
for F, G being BinOp of D
for p being FinSequence of D st F is associative & F is having_a_unity & F is having_an_inverseOp & G is_distributive_wrt F holds
(G [;] (d,(id D))) . (F "**" p) = F "**" ((G [;] (d,(id D))) * p)
proof end;

theorem :: SETWOP_2:31
for D being non empty set
for F being BinOp of D
for p being FinSequence of D st F is commutative & F is associative & F is having_a_unity & F is having_an_inverseOp holds
(the_inverseOp_wrt F) . (F "**" p) = F "**" ((the_inverseOp_wrt F) * p)
proof end;

theorem Th32: :: SETWOP_2:32
for D being non empty set
for e being Element of D
for F, G being BinOp of D
for p, q being FinSequence of D st F is commutative & F is associative & F is having_a_unity & e = the_unity_wrt F & G . (e,e) = e & ( for d1, d2, d3, d4 being Element of D holds F . ((G . (d1,d2)),(G . (d3,d4))) = G . ((F . (d1,d3)),(F . (d2,d4))) ) & len p = len q holds
G . ((F "**" p),(F "**" q)) = F "**" (G .: (p,q))
proof end;

theorem Th33: :: SETWOP_2:33
for D being non empty set
for e being Element of D
for F, G being BinOp of D
for i being Nat
for T1, T2 being Element of i -tuples_on D st F is commutative & F is associative & F is having_a_unity & e = the_unity_wrt F & G . (e,e) = e & ( for d1, d2, d3, d4 being Element of D holds F . ((G . (d1,d2)),(G . (d3,d4))) = G . ((F . (d1,d3)),(F . (d2,d4))) ) holds
G . ((F "**" T1),(F "**" T2)) = F "**" (G .: (T1,T2))
proof end;

theorem Th34: :: SETWOP_2:34
for D being non empty set
for F being BinOp of D
for p, q being FinSequence of D st F is commutative & F is associative & F is having_a_unity & len p = len q holds
F . ((F "**" p),(F "**" q)) = F "**" (F .: (p,q))
proof end;

theorem Th35: :: SETWOP_2:35
for D being non empty set
for F being BinOp of D
for i being Nat
for T1, T2 being Element of i -tuples_on D st F is commutative & F is associative & F is having_a_unity holds
F . ((F "**" T1),(F "**" T2)) = F "**" (F .: (T1,T2))
proof end;

theorem :: SETWOP_2:36
for D being non empty set
for d1, d2 being Element of D
for F being BinOp of D
for i being Nat st F is commutative & F is associative & F is having_a_unity holds
F "**" (i |-> (F . (d1,d2))) = F . ((F "**" (i |-> d1)),(F "**" (i |-> d2)))
proof end;

theorem :: SETWOP_2:37
for D being non empty set
for F, G being BinOp of D
for i being Nat
for T1, T2 being Element of i -tuples_on D st F is commutative & F is associative & F is having_a_unity & F is having_an_inverseOp & G = F * ((id D),(the_inverseOp_wrt F)) holds
G . ((F "**" T1),(F "**" T2)) = F "**" (G .: (T1,T2))
proof end;

theorem Th38: :: SETWOP_2:38
for D being non empty set
for d, e being Element of D
for F, G being BinOp of D
for p being FinSequence of D st F is commutative & F is associative & F is having_a_unity & e = the_unity_wrt F & G is_distributive_wrt F & G . (d,e) = e holds
G . (d,(F "**" p)) = F "**" (G [;] (d,p))
proof end;

theorem Th39: :: SETWOP_2:39
for D being non empty set
for d, e being Element of D
for F, G being BinOp of D
for p being FinSequence of D st F is commutative & F is associative & F is having_a_unity & e = the_unity_wrt F & G is_distributive_wrt F & G . (e,d) = e holds
G . ((F "**" p),d) = F "**" (G [:] (p,d))
proof end;

theorem :: SETWOP_2:40
for D being non empty set
for d being Element of D
for F, G being BinOp of D
for p being FinSequence of D st F is commutative & F is associative & F is having_a_unity & F is having_an_inverseOp & G is_distributive_wrt F holds
G . (d,(F "**" p)) = F "**" (G [;] (d,p))
proof end;

theorem :: SETWOP_2:41
for D being non empty set
for d being Element of D
for F, G being BinOp of D
for p being FinSequence of D st F is commutative & F is associative & F is having_a_unity & F is having_an_inverseOp & G is_distributive_wrt F holds
G . ((F "**" p),d) = F "**" (G [:] (p,d))
proof end;