src/HOL/Algebra/FiniteProduct.thy
author nipkow
Sun, 28 Nov 2010 15:20:51 +0100
changeset 41030 0a54cfc9add3
parent 35849 b5522b51cb1e
child 41681 1b8ff770f02c
permissions -rw-r--r--
gave more standard finite set rules simp and intro attribute
     1 (*  Title:      HOL/Algebra/FiniteProduct.thy
     2     Author:     Clemens Ballarin, started 19 November 2002
     3 
     4 This file is largely based on HOL/Finite_Set.thy.
     5 *)
     6 
     7 theory FiniteProduct
     8 imports Group
     9 begin
    10 
    11 subsection {* Product Operator for Commutative Monoids *}
    12 
    13 subsubsection {* Inductive Definition of a Relation for Products over Sets *}
    14 
    15 text {* Instantiation of locale @{text LC} of theory @{text Finite_Set} is not
    16   possible, because here we have explicit typing rules like 
    17   @{text "x \<in> carrier G"}.  We introduce an explicit argument for the domain
    18   @{text D}. *}
    19 
    20 inductive_set
    21   foldSetD :: "['a set, 'b => 'a => 'a, 'a] => ('b set * 'a) set"
    22   for D :: "'a set" and f :: "'b => 'a => 'a" and e :: 'a
    23   where
    24     emptyI [intro]: "e \<in> D ==> ({}, e) \<in> foldSetD D f e"
    25   | insertI [intro]: "[| x ~: A; f x y \<in> D; (A, y) \<in> foldSetD D f e |] ==>
    26                       (insert x A, f x y) \<in> foldSetD D f e"
    27 
    28 inductive_cases empty_foldSetDE [elim!]: "({}, x) \<in> foldSetD D f e"
    29 
    30 definition
    31   foldD :: "['a set, 'b => 'a => 'a, 'a, 'b set] => 'a"
    32   where "foldD D f e A = (THE x. (A, x) \<in> foldSetD D f e)"
    33 
    34 lemma foldSetD_closed:
    35   "[| (A, z) \<in> foldSetD D f e ; e \<in> D; !!x y. [| x \<in> A; y \<in> D |] ==> f x y \<in> D 
    36       |] ==> z \<in> D";
    37   by (erule foldSetD.cases) auto
    38 
    39 lemma Diff1_foldSetD:
    40   "[| (A - {x}, y) \<in> foldSetD D f e; x \<in> A; f x y \<in> D |] ==>
    41    (A, f x y) \<in> foldSetD D f e"
    42   apply (erule insert_Diff [THEN subst], rule foldSetD.intros)
    43     apply auto
    44   done
    45 
    46 lemma foldSetD_imp_finite [simp]: "(A, x) \<in> foldSetD D f e ==> finite A"
    47   by (induct set: foldSetD) auto
    48 
    49 lemma finite_imp_foldSetD:
    50   "[| finite A; e \<in> D; !!x y. [| x \<in> A; y \<in> D |] ==> f x y \<in> D |] ==>
    51    EX x. (A, x) \<in> foldSetD D f e"
    52 proof (induct set: finite)
    53   case empty then show ?case by auto
    54 next
    55   case (insert x F)
    56   then obtain y where y: "(F, y) \<in> foldSetD D f e" by auto
    57   with insert have "y \<in> D" by (auto dest: foldSetD_closed)
    58   with y and insert have "(insert x F, f x y) \<in> foldSetD D f e"
    59     by (intro foldSetD.intros) auto
    60   then show ?case ..
    61 qed
    62 
    63 
    64 text {* Left-Commutative Operations *}
    65 
    66 locale LCD =
    67   fixes B :: "'b set"
    68   and D :: "'a set"
    69   and f :: "'b => 'a => 'a"    (infixl "\<cdot>" 70)
    70   assumes left_commute:
    71     "[| x \<in> B; y \<in> B; z \<in> D |] ==> x \<cdot> (y \<cdot> z) = y \<cdot> (x \<cdot> z)"
    72   and f_closed [simp, intro!]: "!!x y. [| x \<in> B; y \<in> D |] ==> f x y \<in> D"
    73 
    74 lemma (in LCD) foldSetD_closed [dest]:
    75   "(A, z) \<in> foldSetD D f e ==> z \<in> D";
    76   by (erule foldSetD.cases) auto
    77 
    78 lemma (in LCD) Diff1_foldSetD:
    79   "[| (A - {x}, y) \<in> foldSetD D f e; x \<in> A; A \<subseteq> B |] ==>
    80   (A, f x y) \<in> foldSetD D f e"
    81   apply (subgoal_tac "x \<in> B")
    82    prefer 2 apply fast
    83   apply (erule insert_Diff [THEN subst], rule foldSetD.intros)
    84     apply auto
    85   done
    86 
    87 lemma (in LCD) foldSetD_imp_finite [simp]:
    88   "(A, x) \<in> foldSetD D f e ==> finite A"
    89   by (induct set: foldSetD) auto
    90 
    91 lemma (in LCD) finite_imp_foldSetD:
    92   "[| finite A; A \<subseteq> B; e \<in> D |] ==> EX x. (A, x) \<in> foldSetD D f e"
    93 proof (induct set: finite)
    94   case empty then show ?case by auto
    95 next
    96   case (insert x F)
    97   then obtain y where y: "(F, y) \<in> foldSetD D f e" by auto
    98   with insert have "y \<in> D" by auto
    99   with y and insert have "(insert x F, f x y) \<in> foldSetD D f e"
   100     by (intro foldSetD.intros) auto
   101   then show ?case ..
   102 qed
   103 
   104 lemma (in LCD) foldSetD_determ_aux:
   105   "e \<in> D ==> \<forall>A x. A \<subseteq> B & card A < n --> (A, x) \<in> foldSetD D f e -->
   106     (\<forall>y. (A, y) \<in> foldSetD D f e --> y = x)"
   107   apply (induct n)
   108    apply (auto simp add: less_Suc_eq) (* slow *)
   109   apply (erule foldSetD.cases)
   110    apply blast
   111   apply (erule foldSetD.cases)
   112    apply blast
   113   apply clarify
   114   txt {* force simplification of @{text "card A < card (insert ...)"}. *}
   115   apply (erule rev_mp)
   116   apply (simp add: less_Suc_eq_le)
   117   apply (rule impI)
   118   apply (rename_tac xa Aa ya xb Ab yb, case_tac "xa = xb")
   119    apply (subgoal_tac "Aa = Ab")
   120     prefer 2 apply (blast elim!: equalityE)
   121    apply blast
   122   txt {* case @{prop "xa \<notin> xb"}. *}
   123   apply (subgoal_tac "Aa - {xb} = Ab - {xa} & xb \<in> Aa & xa \<in> Ab")
   124    prefer 2 apply (blast elim!: equalityE)
   125   apply clarify
   126   apply (subgoal_tac "Aa = insert xb Ab - {xa}")
   127    prefer 2 apply blast
   128   apply (subgoal_tac "card Aa \<le> card Ab")
   129    prefer 2
   130    apply (rule Suc_le_mono [THEN subst])
   131    apply (simp add: card_Suc_Diff1)
   132   apply (rule_tac A1 = "Aa - {xb}" in finite_imp_foldSetD [THEN exE])
   133      apply (blast intro: foldSetD_imp_finite)
   134     apply best
   135    apply assumption
   136   apply (frule (1) Diff1_foldSetD)
   137    apply best
   138   apply (subgoal_tac "ya = f xb x")
   139    prefer 2
   140    apply (subgoal_tac "Aa \<subseteq> B")
   141     prefer 2 apply best (* slow *)
   142    apply (blast del: equalityCE)
   143   apply (subgoal_tac "(Ab - {xa}, x) \<in> foldSetD D f e")
   144    prefer 2 apply simp
   145   apply (subgoal_tac "yb = f xa x")
   146    prefer 2 
   147    apply (blast del: equalityCE dest: Diff1_foldSetD)
   148   apply (simp (no_asm_simp))
   149   apply (rule left_commute)
   150     apply assumption
   151    apply best (* slow *)
   152   apply best
   153   done
   154 
   155 lemma (in LCD) foldSetD_determ:
   156   "[| (A, x) \<in> foldSetD D f e; (A, y) \<in> foldSetD D f e; e \<in> D; A \<subseteq> B |]
   157   ==> y = x"
   158   by (blast intro: foldSetD_determ_aux [rule_format])
   159 
   160 lemma (in LCD) foldD_equality:
   161   "[| (A, y) \<in> foldSetD D f e; e \<in> D; A \<subseteq> B |] ==> foldD D f e A = y"
   162   by (unfold foldD_def) (blast intro: foldSetD_determ)
   163 
   164 lemma foldD_empty [simp]:
   165   "e \<in> D ==> foldD D f e {} = e"
   166   by (unfold foldD_def) blast
   167 
   168 lemma (in LCD) foldD_insert_aux:
   169   "[| x ~: A; x \<in> B; e \<in> D; A \<subseteq> B |] ==>
   170     ((insert x A, v) \<in> foldSetD D f e) =
   171     (EX y. (A, y) \<in> foldSetD D f e & v = f x y)"
   172   apply auto
   173   apply (rule_tac A1 = A in finite_imp_foldSetD [THEN exE])
   174      apply (fastsimp dest: foldSetD_imp_finite)
   175     apply assumption
   176    apply assumption
   177   apply (blast intro: foldSetD_determ)
   178   done
   179 
   180 lemma (in LCD) foldD_insert:
   181     "[| finite A; x ~: A; x \<in> B; e \<in> D; A \<subseteq> B |] ==>
   182      foldD D f e (insert x A) = f x (foldD D f e A)"
   183   apply (unfold foldD_def)
   184   apply (simp add: foldD_insert_aux)
   185   apply (rule the_equality)
   186    apply (auto intro: finite_imp_foldSetD
   187      cong add: conj_cong simp add: foldD_def [symmetric] foldD_equality)
   188   done
   189 
   190 lemma (in LCD) foldD_closed [simp]:
   191   "[| finite A; e \<in> D; A \<subseteq> B |] ==> foldD D f e A \<in> D"
   192 proof (induct set: finite)
   193   case empty then show ?case by (simp add: foldD_empty)
   194 next
   195   case insert then show ?case by (simp add: foldD_insert)
   196 qed
   197 
   198 lemma (in LCD) foldD_commute:
   199   "[| finite A; x \<in> B; e \<in> D; A \<subseteq> B |] ==>
   200    f x (foldD D f e A) = foldD D f (f x e) A"
   201   apply (induct set: finite)
   202    apply simp
   203   apply (auto simp add: left_commute foldD_insert)
   204   done
   205 
   206 lemma Int_mono2:
   207   "[| A \<subseteq> C; B \<subseteq> C |] ==> A Int B \<subseteq> C"
   208   by blast
   209 
   210 lemma (in LCD) foldD_nest_Un_Int:
   211   "[| finite A; finite C; e \<in> D; A \<subseteq> B; C \<subseteq> B |] ==>
   212    foldD D f (foldD D f e C) A = foldD D f (foldD D f e (A Int C)) (A Un C)"
   213   apply (induct set: finite)
   214    apply simp
   215   apply (simp add: foldD_insert foldD_commute Int_insert_left insert_absorb
   216     Int_mono2)
   217   done
   218 
   219 lemma (in LCD) foldD_nest_Un_disjoint:
   220   "[| finite A; finite B; A Int B = {}; e \<in> D; A \<subseteq> B; C \<subseteq> B |]
   221     ==> foldD D f e (A Un B) = foldD D f (foldD D f e B) A"
   222   by (simp add: foldD_nest_Un_Int)
   223 
   224 -- {* Delete rules to do with @{text foldSetD} relation. *}
   225 
   226 declare foldSetD_imp_finite [simp del]
   227   empty_foldSetDE [rule del]
   228   foldSetD.intros [rule del]
   229 declare (in LCD)
   230   foldSetD_closed [rule del]
   231 
   232 
   233 text {* Commutative Monoids *}
   234 
   235 text {*
   236   We enter a more restrictive context, with @{text "f :: 'a => 'a => 'a"}
   237   instead of @{text "'b => 'a => 'a"}.
   238 *}
   239 
   240 locale ACeD =
   241   fixes D :: "'a set"
   242     and f :: "'a => 'a => 'a"    (infixl "\<cdot>" 70)
   243     and e :: 'a
   244   assumes ident [simp]: "x \<in> D ==> x \<cdot> e = x"
   245     and commute: "[| x \<in> D; y \<in> D |] ==> x \<cdot> y = y \<cdot> x"
   246     and assoc: "[| x \<in> D; y \<in> D; z \<in> D |] ==> (x \<cdot> y) \<cdot> z = x \<cdot> (y \<cdot> z)"
   247     and e_closed [simp]: "e \<in> D"
   248     and f_closed [simp]: "[| x \<in> D; y \<in> D |] ==> x \<cdot> y \<in> D"
   249 
   250 lemma (in ACeD) left_commute:
   251   "[| x \<in> D; y \<in> D; z \<in> D |] ==> x \<cdot> (y \<cdot> z) = y \<cdot> (x \<cdot> z)"
   252 proof -
   253   assume D: "x \<in> D" "y \<in> D" "z \<in> D"
   254   then have "x \<cdot> (y \<cdot> z) = (y \<cdot> z) \<cdot> x" by (simp add: commute)
   255   also from D have "... = y \<cdot> (z \<cdot> x)" by (simp add: assoc)
   256   also from D have "z \<cdot> x = x \<cdot> z" by (simp add: commute)
   257   finally show ?thesis .
   258 qed
   259 
   260 lemmas (in ACeD) AC = assoc commute left_commute
   261 
   262 lemma (in ACeD) left_ident [simp]: "x \<in> D ==> e \<cdot> x = x"
   263 proof -
   264   assume "x \<in> D"
   265   then have "x \<cdot> e = x" by (rule ident)
   266   with `x \<in> D` show ?thesis by (simp add: commute)
   267 qed
   268 
   269 lemma (in ACeD) foldD_Un_Int:
   270   "[| finite A; finite B; A \<subseteq> D; B \<subseteq> D |] ==>
   271     foldD D f e A \<cdot> foldD D f e B =
   272     foldD D f e (A Un B) \<cdot> foldD D f e (A Int B)"
   273   apply (induct set: finite)
   274    apply (simp add: left_commute LCD.foldD_closed [OF LCD.intro [of D]])
   275   apply (simp add: AC insert_absorb Int_insert_left
   276     LCD.foldD_insert [OF LCD.intro [of D]]
   277     LCD.foldD_closed [OF LCD.intro [of D]]
   278     Int_mono2)
   279   done
   280 
   281 lemma (in ACeD) foldD_Un_disjoint:
   282   "[| finite A; finite B; A Int B = {}; A \<subseteq> D; B \<subseteq> D |] ==>
   283     foldD D f e (A Un B) = foldD D f e A \<cdot> foldD D f e B"
   284   by (simp add: foldD_Un_Int
   285     left_commute LCD.foldD_closed [OF LCD.intro [of D]])
   286 
   287 
   288 subsubsection {* Products over Finite Sets *}
   289 
   290 definition
   291   finprod :: "[('b, 'm) monoid_scheme, 'a => 'b, 'a set] => 'b"
   292   where "finprod G f A =
   293    (if finite A
   294     then foldD (carrier G) (mult G o f) \<one>\<^bsub>G\<^esub> A
   295     else undefined)"
   296 
   297 syntax
   298   "_finprod" :: "index => idt => 'a set => 'b => 'b"
   299       ("(3\<Otimes>__:_. _)" [1000, 0, 51, 10] 10)
   300 syntax (xsymbols)
   301   "_finprod" :: "index => idt => 'a set => 'b => 'b"
   302       ("(3\<Otimes>__\<in>_. _)" [1000, 0, 51, 10] 10)
   303 syntax (HTML output)
   304   "_finprod" :: "index => idt => 'a set => 'b => 'b"
   305       ("(3\<Otimes>__\<in>_. _)" [1000, 0, 51, 10] 10)
   306 translations
   307   "\<Otimes>\<index>i:A. b" == "CONST finprod \<struct>\<index> (%i. b) A"
   308   -- {* Beware of argument permutation! *}
   309 
   310 lemma (in comm_monoid) finprod_empty [simp]: 
   311   "finprod G f {} = \<one>"
   312   by (simp add: finprod_def)
   313 
   314 declare funcsetI [intro]
   315   funcset_mem [dest]
   316 
   317 context comm_monoid begin
   318 
   319 lemma finprod_insert [simp]:
   320   "[| finite F; a \<notin> F; f \<in> F -> carrier G; f a \<in> carrier G |] ==>
   321    finprod G f (insert a F) = f a \<otimes> finprod G f F"
   322   apply (rule trans)
   323    apply (simp add: finprod_def)
   324   apply (rule trans)
   325    apply (rule LCD.foldD_insert [OF LCD.intro [of "insert a F"]])
   326          apply simp
   327          apply (rule m_lcomm)
   328            apply fast
   329           apply fast
   330          apply assumption
   331         apply (fastsimp intro: m_closed)
   332        apply simp+
   333    apply fast
   334   apply (auto simp add: finprod_def)
   335   done
   336 
   337 lemma finprod_one [simp]:
   338   "finite A ==> (\<Otimes>i:A. \<one>) = \<one>"
   339 proof (induct set: finite)
   340   case empty show ?case by simp
   341 next
   342   case (insert a A)
   343   have "(%i. \<one>) \<in> A -> carrier G" by auto
   344   with insert show ?case by simp
   345 qed
   346 
   347 lemma finprod_closed [simp]:
   348   fixes A
   349   assumes fin: "finite A" and f: "f \<in> A -> carrier G" 
   350   shows "finprod G f A \<in> carrier G"
   351 using fin f
   352 proof induct
   353   case empty show ?case by simp
   354 next
   355   case (insert a A)
   356   then have a: "f a \<in> carrier G" by fast
   357   from insert have A: "f \<in> A -> carrier G" by fast
   358   from insert A a show ?case by simp
   359 qed
   360 
   361 lemma funcset_Int_left [simp, intro]:
   362   "[| f \<in> A -> C; f \<in> B -> C |] ==> f \<in> A Int B -> C"
   363   by fast
   364 
   365 lemma funcset_Un_left [iff]:
   366   "(f \<in> A Un B -> C) = (f \<in> A -> C & f \<in> B -> C)"
   367   by fast
   368 
   369 lemma finprod_Un_Int:
   370   "[| finite A; finite B; g \<in> A -> carrier G; g \<in> B -> carrier G |] ==>
   371      finprod G g (A Un B) \<otimes> finprod G g (A Int B) =
   372      finprod G g A \<otimes> finprod G g B"
   373 -- {* The reversed orientation looks more natural, but LOOPS as a simprule! *}
   374 proof (induct set: finite)
   375   case empty then show ?case by (simp add: finprod_closed)
   376 next
   377   case (insert a A)
   378   then have a: "g a \<in> carrier G" by fast
   379   from insert have A: "g \<in> A -> carrier G" by fast
   380   from insert A a show ?case
   381     by (simp add: m_ac Int_insert_left insert_absorb finprod_closed
   382           Int_mono2) 
   383 qed
   384 
   385 lemma finprod_Un_disjoint:
   386   "[| finite A; finite B; A Int B = {};
   387       g \<in> A -> carrier G; g \<in> B -> carrier G |]
   388    ==> finprod G g (A Un B) = finprod G g A \<otimes> finprod G g B"
   389   apply (subst finprod_Un_Int [symmetric])
   390       apply (auto simp add: finprod_closed)
   391   done
   392 
   393 lemma finprod_multf:
   394   "[| finite A; f \<in> A -> carrier G; g \<in> A -> carrier G |] ==>
   395    finprod G (%x. f x \<otimes> g x) A = (finprod G f A \<otimes> finprod G g A)"
   396 proof (induct set: finite)
   397   case empty show ?case by simp
   398 next
   399   case (insert a A) then
   400   have fA: "f \<in> A -> carrier G" by fast
   401   from insert have fa: "f a \<in> carrier G" by fast
   402   from insert have gA: "g \<in> A -> carrier G" by fast
   403   from insert have ga: "g a \<in> carrier G" by fast
   404   from insert have fgA: "(%x. f x \<otimes> g x) \<in> A -> carrier G"
   405     by (simp add: Pi_def)
   406   show ?case
   407     by (simp add: insert fA fa gA ga fgA m_ac)
   408 qed
   409 
   410 lemma finprod_cong':
   411   "[| A = B; g \<in> B -> carrier G;
   412       !!i. i \<in> B ==> f i = g i |] ==> finprod G f A = finprod G g B"
   413 proof -
   414   assume prems: "A = B" "g \<in> B -> carrier G"
   415     "!!i. i \<in> B ==> f i = g i"
   416   show ?thesis
   417   proof (cases "finite B")
   418     case True
   419     then have "!!A. [| A = B; g \<in> B -> carrier G;
   420       !!i. i \<in> B ==> f i = g i |] ==> finprod G f A = finprod G g B"
   421     proof induct
   422       case empty thus ?case by simp
   423     next
   424       case (insert x B)
   425       then have "finprod G f A = finprod G f (insert x B)" by simp
   426       also from insert have "... = f x \<otimes> finprod G f B"
   427       proof (intro finprod_insert)
   428         show "finite B" by fact
   429       next
   430         show "x ~: B" by fact
   431       next
   432         assume "x ~: B" "!!i. i \<in> insert x B \<Longrightarrow> f i = g i"
   433           "g \<in> insert x B \<rightarrow> carrier G"
   434         thus "f \<in> B -> carrier G" by fastsimp
   435       next
   436         assume "x ~: B" "!!i. i \<in> insert x B \<Longrightarrow> f i = g i"
   437           "g \<in> insert x B \<rightarrow> carrier G"
   438         thus "f x \<in> carrier G" by fastsimp
   439       qed
   440       also from insert have "... = g x \<otimes> finprod G g B" by fastsimp
   441       also from insert have "... = finprod G g (insert x B)"
   442       by (intro finprod_insert [THEN sym]) auto
   443       finally show ?case .
   444     qed
   445     with prems show ?thesis by simp
   446   next
   447     case False with prems show ?thesis by (simp add: finprod_def)
   448   qed
   449 qed
   450 
   451 lemma finprod_cong:
   452   "[| A = B; f \<in> B -> carrier G = True;
   453       !!i. i \<in> B ==> f i = g i |] ==> finprod G f A = finprod G g B"
   454   (* This order of prems is slightly faster (3%) than the last two swapped. *)
   455   by (rule finprod_cong') force+
   456 
   457 text {*Usually, if this rule causes a failed congruence proof error,
   458   the reason is that the premise @{text "g \<in> B -> carrier G"} cannot be shown.
   459   Adding @{thm [source] Pi_def} to the simpset is often useful.
   460   For this reason, @{thm [source] comm_monoid.finprod_cong}
   461   is not added to the simpset by default.
   462 *}
   463 
   464 end
   465 
   466 declare funcsetI [rule del]
   467   funcset_mem [rule del]
   468 
   469 context comm_monoid begin
   470 
   471 lemma finprod_0 [simp]:
   472   "f \<in> {0::nat} -> carrier G ==> finprod G f {..0} = f 0"
   473 by (simp add: Pi_def)
   474 
   475 lemma finprod_Suc [simp]:
   476   "f \<in> {..Suc n} -> carrier G ==>
   477    finprod G f {..Suc n} = (f (Suc n) \<otimes> finprod G f {..n})"
   478 by (simp add: Pi_def atMost_Suc)
   479 
   480 lemma finprod_Suc2:
   481   "f \<in> {..Suc n} -> carrier G ==>
   482    finprod G f {..Suc n} = (finprod G (%i. f (Suc i)) {..n} \<otimes> f 0)"
   483 proof (induct n)
   484   case 0 thus ?case by (simp add: Pi_def)
   485 next
   486   case Suc thus ?case by (simp add: m_assoc Pi_def)
   487 qed
   488 
   489 lemma finprod_mult [simp]:
   490   "[| f \<in> {..n} -> carrier G; g \<in> {..n} -> carrier G |] ==>
   491      finprod G (%i. f i \<otimes> g i) {..n::nat} =
   492      finprod G f {..n} \<otimes> finprod G g {..n}"
   493   by (induct n) (simp_all add: m_ac Pi_def)
   494 
   495 (* The following two were contributed by Jeremy Avigad. *)
   496 
   497 lemma finprod_reindex:
   498   assumes fin: "finite A"
   499     shows "f : (h ` A) \<rightarrow> carrier G \<Longrightarrow> 
   500         inj_on h A ==> finprod G f (h ` A) = finprod G (%x. f (h x)) A"
   501   using fin apply induct
   502   apply (auto simp add: finprod_insert Pi_def)
   503 done
   504 
   505 lemma finprod_const:
   506   assumes fin [simp]: "finite A"
   507       and a [simp]: "a : carrier G"
   508     shows "finprod G (%x. a) A = a (^) card A"
   509   using fin apply induct
   510   apply force
   511   apply (subst finprod_insert)
   512   apply auto
   513   apply (subst m_comm)
   514   apply auto
   515 done
   516 
   517 (* The following lemma was contributed by Jesus Aransay. *)
   518 
   519 lemma finprod_singleton:
   520   assumes i_in_A: "i \<in> A" and fin_A: "finite A" and f_Pi: "f \<in> A \<rightarrow> carrier G"
   521   shows "(\<Otimes>j\<in>A. if i = j then f j else \<one>) = f i"
   522   using i_in_A finprod_insert [of "A - {i}" i "(\<lambda>j. if i = j then f j else \<one>)"]
   523     fin_A f_Pi finprod_one [of "A - {i}"]
   524     finprod_cong [of "A - {i}" "A - {i}" "(\<lambda>j. if i = j then f j else \<one>)" "(\<lambda>i. \<one>)"] 
   525   unfolding Pi_def simp_implies_def by (force simp add: insert_absorb)
   526 
   527 end
   528 
   529 end