Up to index of Isabelle/HOL/HOL-Algebra
theory FiniteProduct(* Title: HOL/Algebra/FiniteProduct.thy ID: $Id: FiniteProduct.thy,v 1.15 2007/07/11 09:14:51 berghofe Exp $ Author: Clemens Ballarin, started 19 November 2002 This file is largely based on HOL/Finite_Set.thy. *) theory FiniteProduct imports Group begin section {* Product Operator for Commutative Monoids *} subsection {* Inductive Definition of a Relation for Products over Sets *} text {* Instantiation of locale @{text LC} of theory @{text Finite_Set} is not possible, because here we have explicit typing rules like @{text "x ∈ carrier G"}. We introduce an explicit argument for the domain @{text D}. *} inductive_set foldSetD :: "['a set, 'b => 'a => 'a, 'a] => ('b set * 'a) set" for D :: "'a set" and f :: "'b => 'a => 'a" and e :: 'a where emptyI [intro]: "e ∈ D ==> ({}, e) ∈ foldSetD D f e" | insertI [intro]: "[| x ~: A; f x y ∈ D; (A, y) ∈ foldSetD D f e |] ==> (insert x A, f x y) ∈ foldSetD D f e" inductive_cases empty_foldSetDE [elim!]: "({}, x) ∈ foldSetD D f e" constdefs foldD :: "['a set, 'b => 'a => 'a, 'a, 'b set] => 'a" "foldD D f e A == THE x. (A, x) ∈ foldSetD D f e" lemma foldSetD_closed: "[| (A, z) ∈ foldSetD D f e ; e ∈ D; !!x y. [| x ∈ A; y ∈ D |] ==> f x y ∈ D |] ==> z ∈ D"; by (erule foldSetD.cases) auto lemma Diff1_foldSetD: "[| (A - {x}, y) ∈ foldSetD D f e; x ∈ A; f x y ∈ D |] ==> (A, f x y) ∈ foldSetD D f e" apply (erule insert_Diff [THEN subst], rule foldSetD.intros) apply auto done lemma foldSetD_imp_finite [simp]: "(A, x) ∈ foldSetD D f e ==> finite A" by (induct set: foldSetD) auto lemma finite_imp_foldSetD: "[| finite A; e ∈ D; !!x y. [| x ∈ A; y ∈ D |] ==> f x y ∈ D |] ==> EX x. (A, x) ∈ foldSetD D f e" proof (induct set: finite) case empty then show ?case by auto next case (insert x F) then obtain y where y: "(F, y) ∈ foldSetD D f e" by auto with insert have "y ∈ D" by (auto dest: foldSetD_closed) with y and insert have "(insert x F, f x y) ∈ foldSetD D f e" by (intro foldSetD.intros) auto then show ?case .. qed subsection {* Left-Commutative Operations *} locale LCD = fixes B :: "'b set" and D :: "'a set" and f :: "'b => 'a => 'a" (infixl "·" 70) assumes left_commute: "[| x ∈ B; y ∈ B; z ∈ D |] ==> x · (y · z) = y · (x · z)" and f_closed [simp, intro!]: "!!x y. [| x ∈ B; y ∈ D |] ==> f x y ∈ D" lemma (in LCD) foldSetD_closed [dest]: "(A, z) ∈ foldSetD D f e ==> z ∈ D"; by (erule foldSetD.cases) auto lemma (in LCD) Diff1_foldSetD: "[| (A - {x}, y) ∈ foldSetD D f e; x ∈ A; A ⊆ B |] ==> (A, f x y) ∈ foldSetD D f e" apply (subgoal_tac "x ∈ B") prefer 2 apply fast apply (erule insert_Diff [THEN subst], rule foldSetD.intros) apply auto done lemma (in LCD) foldSetD_imp_finite [simp]: "(A, x) ∈ foldSetD D f e ==> finite A" by (induct set: foldSetD) auto lemma (in LCD) finite_imp_foldSetD: "[| finite A; A ⊆ B; e ∈ D |] ==> EX x. (A, x) ∈ foldSetD D f e" proof (induct set: finite) case empty then show ?case by auto next case (insert x F) then obtain y where y: "(F, y) ∈ foldSetD D f e" by auto with insert have "y ∈ D" by auto with y and insert have "(insert x F, f x y) ∈ foldSetD D f e" by (intro foldSetD.intros) auto then show ?case .. qed lemma (in LCD) foldSetD_determ_aux: "e ∈ D ==> ∀A x. A ⊆ B & card A < n --> (A, x) ∈ foldSetD D f e --> (∀y. (A, y) ∈ foldSetD D f e --> y = x)" apply (induct n) apply (auto simp add: less_Suc_eq) (* slow *) apply (erule foldSetD.cases) apply blast apply (erule foldSetD.cases) apply blast apply clarify txt {* force simplification of @{text "card A < card (insert ...)"}. *} apply (erule rev_mp) apply (simp add: less_Suc_eq_le) apply (rule impI) apply (rename_tac xa Aa ya xb Ab yb, case_tac "xa = xb") apply (subgoal_tac "Aa = Ab") prefer 2 apply (blast elim!: equalityE) apply blast txt {* case @{prop "xa ∉ xb"}. *} apply (subgoal_tac "Aa - {xb} = Ab - {xa} & xb ∈ Aa & xa ∈ Ab") prefer 2 apply (blast elim!: equalityE) apply clarify apply (subgoal_tac "Aa = insert xb Ab - {xa}") prefer 2 apply blast apply (subgoal_tac "card Aa ≤ card Ab") prefer 2 apply (rule Suc_le_mono [THEN subst]) apply (simp add: card_Suc_Diff1) apply (rule_tac A1 = "Aa - {xb}" in finite_imp_foldSetD [THEN exE]) apply (blast intro: foldSetD_imp_finite finite_Diff) apply best apply assumption apply (frule (1) Diff1_foldSetD) apply best apply (subgoal_tac "ya = f xb x") prefer 2 apply (subgoal_tac "Aa ⊆ B") prefer 2 apply best (* slow *) apply (blast del: equalityCE) apply (subgoal_tac "(Ab - {xa}, x) ∈ foldSetD D f e") prefer 2 apply simp apply (subgoal_tac "yb = f xa x") prefer 2 apply (blast del: equalityCE dest: Diff1_foldSetD) apply (simp (no_asm_simp)) apply (rule left_commute) apply assumption apply best (* slow *) apply best done lemma (in LCD) foldSetD_determ: "[| (A, x) ∈ foldSetD D f e; (A, y) ∈ foldSetD D f e; e ∈ D; A ⊆ B |] ==> y = x" by (blast intro: foldSetD_determ_aux [rule_format]) lemma (in LCD) foldD_equality: "[| (A, y) ∈ foldSetD D f e; e ∈ D; A ⊆ B |] ==> foldD D f e A = y" by (unfold foldD_def) (blast intro: foldSetD_determ) lemma foldD_empty [simp]: "e ∈ D ==> foldD D f e {} = e" by (unfold foldD_def) blast lemma (in LCD) foldD_insert_aux: "[| x ~: A; x ∈ B; e ∈ D; A ⊆ B |] ==> ((insert x A, v) ∈ foldSetD D f e) = (EX y. (A, y) ∈ foldSetD D f e & v = f x y)" apply auto apply (rule_tac A1 = A in finite_imp_foldSetD [THEN exE]) apply (fastsimp dest: foldSetD_imp_finite) apply assumption apply assumption apply (blast intro: foldSetD_determ) done lemma (in LCD) foldD_insert: "[| finite A; x ~: A; x ∈ B; e ∈ D; A ⊆ B |] ==> foldD D f e (insert x A) = f x (foldD D f e A)" apply (unfold foldD_def) apply (simp add: foldD_insert_aux) apply (rule the_equality) apply (auto intro: finite_imp_foldSetD cong add: conj_cong simp add: foldD_def [symmetric] foldD_equality) done lemma (in LCD) foldD_closed [simp]: "[| finite A; e ∈ D; A ⊆ B |] ==> foldD D f e A ∈ D" proof (induct set: finite) case empty then show ?case by (simp add: foldD_empty) next case insert then show ?case by (simp add: foldD_insert) qed lemma (in LCD) foldD_commute: "[| finite A; x ∈ B; e ∈ D; A ⊆ B |] ==> f x (foldD D f e A) = foldD D f (f x e) A" apply (induct set: finite) apply simp apply (auto simp add: left_commute foldD_insert) done lemma Int_mono2: "[| A ⊆ C; B ⊆ C |] ==> A Int B ⊆ C" by blast lemma (in LCD) foldD_nest_Un_Int: "[| finite A; finite C; e ∈ D; A ⊆ B; C ⊆ B |] ==> foldD D f (foldD D f e C) A = foldD D f (foldD D f e (A Int C)) (A Un C)" apply (induct set: finite) apply simp apply (simp add: foldD_insert foldD_commute Int_insert_left insert_absorb Int_mono2 Un_subset_iff) done lemma (in LCD) foldD_nest_Un_disjoint: "[| finite A; finite B; A Int B = {}; e ∈ D; A ⊆ B; C ⊆ B |] ==> foldD D f e (A Un B) = foldD D f (foldD D f e B) A" by (simp add: foldD_nest_Un_Int) -- {* Delete rules to do with @{text foldSetD} relation. *} declare foldSetD_imp_finite [simp del] empty_foldSetDE [rule del] foldSetD.intros [rule del] declare (in LCD) foldSetD_closed [rule del] subsection {* Commutative Monoids *} text {* We enter a more restrictive context, with @{text "f :: 'a => 'a => 'a"} instead of @{text "'b => 'a => 'a"}. *} locale ACeD = fixes D :: "'a set" and f :: "'a => 'a => 'a" (infixl "·" 70) and e :: 'a assumes ident [simp]: "x ∈ D ==> x · e = x" and commute: "[| x ∈ D; y ∈ D |] ==> x · y = y · x" and assoc: "[| x ∈ D; y ∈ D; z ∈ D |] ==> (x · y) · z = x · (y · z)" and e_closed [simp]: "e ∈ D" and f_closed [simp]: "[| x ∈ D; y ∈ D |] ==> x · y ∈ D" lemma (in ACeD) left_commute: "[| x ∈ D; y ∈ D; z ∈ D |] ==> x · (y · z) = y · (x · z)" proof - assume D: "x ∈ D" "y ∈ D" "z ∈ D" then have "x · (y · z) = (y · z) · x" by (simp add: commute) also from D have "... = y · (z · x)" by (simp add: assoc) also from D have "z · x = x · z" by (simp add: commute) finally show ?thesis . qed lemmas (in ACeD) AC = assoc commute left_commute lemma (in ACeD) left_ident [simp]: "x ∈ D ==> e · x = x" proof - assume "x ∈ D" then have "x · e = x" by (rule ident) with `x ∈ D` show ?thesis by (simp add: commute) qed lemma (in ACeD) foldD_Un_Int: "[| finite A; finite B; A ⊆ D; B ⊆ D |] ==> foldD D f e A · foldD D f e B = foldD D f e (A Un B) · foldD D f e (A Int B)" apply (induct set: finite) apply (simp add: left_commute LCD.foldD_closed [OF LCD.intro [of D]]) apply (simp add: AC insert_absorb Int_insert_left LCD.foldD_insert [OF LCD.intro [of D]] LCD.foldD_closed [OF LCD.intro [of D]] Int_mono2 Un_subset_iff) done lemma (in ACeD) foldD_Un_disjoint: "[| finite A; finite B; A Int B = {}; A ⊆ D; B ⊆ D |] ==> foldD D f e (A Un B) = foldD D f e A · foldD D f e B" by (simp add: foldD_Un_Int left_commute LCD.foldD_closed [OF LCD.intro [of D]] Un_subset_iff) subsection {* Products over Finite Sets *} constdefs (structure G) finprod :: "[('b, 'm) monoid_scheme, 'a => 'b, 'a set] => 'b" "finprod G f A == if finite A then foldD (carrier G) (mult G o f) \<one> A else arbitrary" syntax "_finprod" :: "index => idt => 'a set => 'b => 'b" ("(3\<Otimes>__:_. _)" [1000, 0, 51, 10] 10) syntax (xsymbols) "_finprod" :: "index => idt => 'a set => 'b => 'b" ("(3\<Otimes>__∈_. _)" [1000, 0, 51, 10] 10) syntax (HTML output) "_finprod" :: "index => idt => 'a set => 'b => 'b" ("(3\<Otimes>__∈_. _)" [1000, 0, 51, 10] 10) translations "\<Otimes>\<index>i:A. b" == "finprod \<struct>\<index> (%i. b) A" -- {* Beware of argument permutation! *} lemma (in comm_monoid) finprod_empty [simp]: "finprod G f {} = \<one>" by (simp add: finprod_def) declare funcsetI [intro] funcset_mem [dest] lemma (in comm_monoid) finprod_insert [simp]: "[| finite F; a ∉ F; f ∈ F -> carrier G; f a ∈ carrier G |] ==> finprod G f (insert a F) = f a ⊗ finprod G f F" apply (rule trans) apply (simp add: finprod_def) apply (rule trans) apply (rule LCD.foldD_insert [OF LCD.intro [of "insert a F"]]) apply simp apply (rule m_lcomm) apply fast apply fast apply assumption apply (fastsimp intro: m_closed) apply simp+ apply fast apply (auto simp add: finprod_def) done lemma (in comm_monoid) finprod_one [simp]: "finite A ==> (\<Otimes>i:A. \<one>) = \<one>" proof (induct set: finite) case empty show ?case by simp next case (insert a A) have "(%i. \<one>) ∈ A -> carrier G" by auto with insert show ?case by simp qed lemma (in comm_monoid) finprod_closed [simp]: fixes A assumes fin: "finite A" and f: "f ∈ A -> carrier G" shows "finprod G f A ∈ carrier G" using fin f proof induct case empty show ?case by simp next case (insert a A) then have a: "f a ∈ carrier G" by fast from insert have A: "f ∈ A -> carrier G" by fast from insert A a show ?case by simp qed lemma funcset_Int_left [simp, intro]: "[| f ∈ A -> C; f ∈ B -> C |] ==> f ∈ A Int B -> C" by fast lemma funcset_Un_left [iff]: "(f ∈ A Un B -> C) = (f ∈ A -> C & f ∈ B -> C)" by fast lemma (in comm_monoid) finprod_Un_Int: "[| finite A; finite B; g ∈ A -> carrier G; g ∈ B -> carrier G |] ==> finprod G g (A Un B) ⊗ finprod G g (A Int B) = finprod G g A ⊗ finprod G g B" -- {* The reversed orientation looks more natural, but LOOPS as a simprule! *} proof (induct set: finite) case empty then show ?case by (simp add: finprod_closed) next case (insert a A) then have a: "g a ∈ carrier G" by fast from insert have A: "g ∈ A -> carrier G" by fast from insert A a show ?case by (simp add: m_ac Int_insert_left insert_absorb finprod_closed Int_mono2 Un_subset_iff) qed lemma (in comm_monoid) finprod_Un_disjoint: "[| finite A; finite B; A Int B = {}; g ∈ A -> carrier G; g ∈ B -> carrier G |] ==> finprod G g (A Un B) = finprod G g A ⊗ finprod G g B" apply (subst finprod_Un_Int [symmetric]) apply (auto simp add: finprod_closed) done lemma (in comm_monoid) finprod_multf: "[| finite A; f ∈ A -> carrier G; g ∈ A -> carrier G |] ==> finprod G (%x. f x ⊗ g x) A = (finprod G f A ⊗ finprod G g A)" proof (induct set: finite) case empty show ?case by simp next case (insert a A) then have fA: "f ∈ A -> carrier G" by fast from insert have fa: "f a ∈ carrier G" by fast from insert have gA: "g ∈ A -> carrier G" by fast from insert have ga: "g a ∈ carrier G" by fast from insert have fgA: "(%x. f x ⊗ g x) ∈ A -> carrier G" by (simp add: Pi_def) show ?case by (simp add: insert fA fa gA ga fgA m_ac) qed lemma (in comm_monoid) finprod_cong': "[| A = B; g ∈ B -> carrier G; !!i. i ∈ B ==> f i = g i |] ==> finprod G f A = finprod G g B" proof - assume prems: "A = B" "g ∈ B -> carrier G" "!!i. i ∈ B ==> f i = g i" show ?thesis proof (cases "finite B") case True then have "!!A. [| A = B; g ∈ B -> carrier G; !!i. i ∈ B ==> f i = g i |] ==> finprod G f A = finprod G g B" proof induct case empty thus ?case by simp next case (insert x B) then have "finprod G f A = finprod G f (insert x B)" by simp also from insert have "... = f x ⊗ finprod G f B" proof (intro finprod_insert) show "finite B" by fact next show "x ~: B" by fact next assume "x ~: B" "!!i. i ∈ insert x B ==> f i = g i" "g ∈ insert x B -> carrier G" thus "f ∈ B -> carrier G" by fastsimp next assume "x ~: B" "!!i. i ∈ insert x B ==> f i = g i" "g ∈ insert x B -> carrier G" thus "f x ∈ carrier G" by fastsimp qed also from insert have "... = g x ⊗ finprod G g B" by fastsimp also from insert have "... = finprod G g (insert x B)" by (intro finprod_insert [THEN sym]) auto finally show ?case . qed with prems show ?thesis by simp next case False with prems show ?thesis by (simp add: finprod_def) qed qed lemma (in comm_monoid) finprod_cong: "[| A = B; f ∈ B -> carrier G = True; !!i. i ∈ B ==> f i = g i |] ==> finprod G f A = finprod G g B" (* This order of prems is slightly faster (3%) than the last two swapped. *) by (rule finprod_cong') force+ text {*Usually, if this rule causes a failed congruence proof error, the reason is that the premise @{text "g ∈ B -> carrier G"} cannot be shown. Adding @{thm [source] Pi_def} to the simpset is often useful. For this reason, @{thm [source] comm_monoid.finprod_cong} is not added to the simpset by default. *} declare funcsetI [rule del] funcset_mem [rule del] lemma (in comm_monoid) finprod_0 [simp]: "f ∈ {0::nat} -> carrier G ==> finprod G f {..0} = f 0" by (simp add: Pi_def) lemma (in comm_monoid) finprod_Suc [simp]: "f ∈ {..Suc n} -> carrier G ==> finprod G f {..Suc n} = (f (Suc n) ⊗ finprod G f {..n})" by (simp add: Pi_def atMost_Suc) lemma (in comm_monoid) finprod_Suc2: "f ∈ {..Suc n} -> carrier G ==> finprod G f {..Suc n} = (finprod G (%i. f (Suc i)) {..n} ⊗ f 0)" proof (induct n) case 0 thus ?case by (simp add: Pi_def) next case Suc thus ?case by (simp add: m_assoc Pi_def) qed lemma (in comm_monoid) finprod_mult [simp]: "[| f ∈ {..n} -> carrier G; g ∈ {..n} -> carrier G |] ==> finprod G (%i. f i ⊗ g i) {..n::nat} = finprod G f {..n} ⊗ finprod G g {..n}" by (induct n) (simp_all add: m_ac Pi_def) end
lemma foldSetD_closed:
[| (A, z) ∈ foldSetD D f e; e ∈ D; !!x y. [| x ∈ A; y ∈ D |] ==> f x y ∈ D |]
==> z ∈ D
lemma Diff1_foldSetD:
[| (A - {x}, y) ∈ foldSetD D f e; x ∈ A; f x y ∈ D |]
==> (A, f x y) ∈ foldSetD D f e
lemma foldSetD_imp_finite:
(A, x) ∈ foldSetD D f e ==> finite A
lemma finite_imp_foldSetD:
[| finite A; e ∈ D; !!x y. [| x ∈ A; y ∈ D |] ==> f x y ∈ D |]
==> ∃x. (A, x) ∈ foldSetD D f e
lemma foldSetD_closed:
(A, z) ∈ foldSetD D op · e ==> z ∈ D
lemma Diff1_foldSetD:
[| (A - {x}, y) ∈ foldSetD D op · e; x ∈ A; A ⊆ B |]
==> (A, x · y) ∈ foldSetD D op · e
lemma foldSetD_imp_finite:
(A, x) ∈ foldSetD D op · e ==> finite A
lemma finite_imp_foldSetD:
[| finite A; A ⊆ B; e ∈ D |] ==> ∃x. (A, x) ∈ foldSetD D op · e
lemma foldSetD_determ_aux:
e ∈ D
==> ∀A x. A ⊆ B ∧ card A < n -->
(A, x) ∈ foldSetD D op · e -->
(∀y. (A, y) ∈ foldSetD D op · e --> y = x)
lemma foldSetD_determ:
[| (A, x) ∈ foldSetD D op · e; (A, y) ∈ foldSetD D op · e; e ∈ D; A ⊆ B |]
==> y = x
lemma foldD_equality:
[| (A, y) ∈ foldSetD D op · e; e ∈ D; A ⊆ B |] ==> foldD D op · e A = y
lemma foldD_empty:
e ∈ D ==> foldD D f e {} = e
lemma foldD_insert_aux:
[| x ∉ A; x ∈ B; e ∈ D; A ⊆ B |]
==> ((insert x A, v) ∈ foldSetD D op · e) =
(∃y. (A, y) ∈ foldSetD D op · e ∧ v = x · y)
lemma foldD_insert:
[| finite A; x ∉ A; x ∈ B; e ∈ D; A ⊆ B |]
==> foldD D op · e (insert x A) = x · foldD D op · e A
lemma foldD_closed:
[| finite A; e ∈ D; A ⊆ B |] ==> foldD D op · e A ∈ D
lemma foldD_commute:
[| finite A; x ∈ B; e ∈ D; A ⊆ B |]
==> x · foldD D op · e A = foldD D op · (x · e) A
lemma Int_mono2:
[| A ⊆ C; B ⊆ C |] ==> A ∩ B ⊆ C
lemma foldD_nest_Un_Int:
[| finite A; finite C; e ∈ D; A ⊆ B; C ⊆ B |]
==> foldD D op · (foldD D op · e C) A =
foldD D op · (foldD D op · e (A ∩ C)) (A ∪ C)
lemma foldD_nest_Un_disjoint:
[| finite A; finite B; A ∩ B = {}; e ∈ D; A ⊆ B; C ⊆ B |]
==> foldD D op · e (A ∪ B) = foldD D op · (foldD D op · e B) A
lemma left_commute:
[| x ∈ D; y ∈ D; z ∈ D |] ==> x · (y · z) = y · (x · z)
lemma AC:
[| x ∈ D; y ∈ D; z ∈ D |] ==> x · y · z = x · (y · z)
[| x ∈ D; y ∈ D |] ==> x · y = y · x
[| x ∈ D; y ∈ D; z ∈ D |] ==> x · (y · z) = y · (x · z)
lemma left_ident:
x ∈ D ==> e · x = x
lemma foldD_Un_Int:
[| finite A; finite B; A ⊆ D; B ⊆ D |]
==> foldD D op · e A · foldD D op · e B =
foldD D op · e (A ∪ B) · foldD D op · e (A ∩ B)
lemma foldD_Un_disjoint:
[| finite A; finite B; A ∩ B = {}; A ⊆ D; B ⊆ D |]
==> foldD D op · e (A ∪ B) = foldD D op · e A · foldD D op · e B
lemma finprod_empty:
finprod G f {} = \<one>
lemma finprod_insert:
[| finite F; a ∉ F; f ∈ F -> carrier G; f a ∈ carrier G |]
==> finprod G f (insert a F) = f a ⊗ finprod G f F
lemma finprod_one:
finite A ==> (\<Otimes>i∈A. \<one>) = \<one>
lemma finprod_closed:
[| finite A; f ∈ A -> carrier G |] ==> finprod G f A ∈ carrier G
lemma funcset_Int_left:
[| f ∈ A -> C; f ∈ B -> C |] ==> f ∈ A ∩ B -> C
lemma funcset_Un_left:
(f ∈ A ∪ B -> C) = (f ∈ A -> C ∧ f ∈ B -> C)
lemma finprod_Un_Int:
[| finite A; finite B; g ∈ A -> carrier G; g ∈ B -> carrier G |]
==> finprod G g (A ∪ B) ⊗ finprod G g (A ∩ B) = finprod G g A ⊗ finprod G g B
lemma finprod_Un_disjoint:
[| finite A; finite B; A ∩ B = {}; g ∈ A -> carrier G; g ∈ B -> carrier G |]
==> finprod G g (A ∪ B) = finprod G g A ⊗ finprod G g B
lemma finprod_multf:
[| finite A; f ∈ A -> carrier G; g ∈ A -> carrier G |]
==> (\<Otimes>x∈A. f x ⊗ g x) = finprod G f A ⊗ finprod G g A
lemma finprod_cong':
[| A = B; g ∈ B -> carrier G; !!i. i ∈ B ==> f i = g i |]
==> finprod G f A = finprod G g B
lemma finprod_cong:
[| A = B; (f ∈ B -> carrier G) = True; !!i. i ∈ B ==> f i = g i |]
==> finprod G f A = finprod G g B
lemma finprod_0:
f ∈ {0} -> carrier G ==> finprod G f {..0} = f 0
lemma finprod_Suc:
f ∈ {..Suc n} -> carrier G
==> finprod G f {..Suc n} = f (Suc n) ⊗ finprod G f {..n}
lemma finprod_Suc2:
f ∈ {..Suc n} -> carrier G
==> finprod G f {..Suc n} = (\<Otimes>i∈{..n}. f (Suc i)) ⊗ f 0
lemma finprod_mult:
[| f ∈ {..n} -> carrier G; g ∈ {..n} -> carrier G |]
==> (\<Otimes>i∈{..n}. f i ⊗ g i) = finprod G f {..n} ⊗ finprod G g {..n}