| author | Fabian Huch <huch@in.tum.de> | 
| Fri, 01 Dec 2023 20:32:34 +0100 | |
| changeset 79101 | 4e47b34fbb8e | 
| parent 73932 | fd21b4a93043 | 
| permissions | -rw-r--r-- | 
| 70045 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1 | section\<open>Free Abelian Groups\<close> | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3 | theory Free_Abelian_Groups | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4 | imports | 
| 72630 
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
 paulson <lp15@cam.ac.uk> parents: 
70660diff
changeset | 5 | Product_Groups FiniteProduct "HOL-Cardinals.Cardinal_Arithmetic" | 
| 70045 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 6 | "HOL-Library.Countable_Set" "HOL-Library.Poly_Mapping" "HOL-Library.Equipollence" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 7 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 8 | begin | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 9 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 10 | (*Move? But where?*) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 11 | lemma eqpoll_Fpow: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 12 | assumes "infinite A" shows "Fpow A \<approx> A" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 13 | unfolding eqpoll_iff_card_of_ordIso | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 14 | by (metis assms card_of_Fpow_infinite) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 15 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 16 | lemma infinite_iff_card_of_countable: "\<lbrakk>countable B; infinite B\<rbrakk> \<Longrightarrow> infinite A \<longleftrightarrow> ( |B| \<le>o |A| )" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 17 | unfolding infinite_iff_countable_subset card_of_ordLeq countable_def | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 18 | by (force intro: card_of_ordLeqI ordLeq_transitive) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 19 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 20 | lemma iso_imp_eqpoll_carrier: "G \<cong> H \<Longrightarrow> carrier G \<approx> carrier H" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 21 | by (auto simp: is_iso_def iso_def eqpoll_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 22 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 23 | subsection\<open>Generalised finite product\<close> | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 24 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 25 | definition | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 26 |   gfinprod :: "[('b, 'm) monoid_scheme, 'a \<Rightarrow> 'b, 'a set] \<Rightarrow> 'b"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 27 | where "gfinprod G f A = | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 28 |    (if finite {x \<in> A. f x \<noteq> \<one>\<^bsub>G\<^esub>} then finprod G f {x \<in> A. f x \<noteq> \<one>\<^bsub>G\<^esub>} else \<one>\<^bsub>G\<^esub>)"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 29 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 30 | context comm_monoid begin | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 31 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 32 | lemma gfinprod_closed [simp]: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 33 | "f \<in> A \<rightarrow> carrier G \<Longrightarrow> gfinprod G f A \<in> carrier G" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 34 | unfolding gfinprod_def | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 35 | by (auto simp: image_subset_iff_funcset intro: finprod_closed) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 36 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 37 | lemma gfinprod_cong: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 38 | "\<lbrakk>A = B; f \<in> B \<rightarrow> carrier G; | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 39 | \<And>i. i \<in> B =simp=> f i = g i\<rbrakk> \<Longrightarrow> gfinprod G f A = gfinprod G g B" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 40 | unfolding gfinprod_def | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 41 | by (auto simp: simp_implies_def cong: conj_cong intro: finprod_cong) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 42 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 43 | lemma gfinprod_eq_finprod [simp]: "\<lbrakk>finite A; f \<in> A \<rightarrow> carrier G\<rbrakk> \<Longrightarrow> gfinprod G f A = finprod G f A" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 44 | by (auto simp: gfinprod_def intro: finprod_mono_neutral_cong_left) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 45 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 46 | lemma gfinprod_insert [simp]: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 47 |   assumes "finite {x \<in> A. f x \<noteq> \<one>\<^bsub>G\<^esub>}" "f \<in> A \<rightarrow> carrier G" "f i \<in> carrier G"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 48 | shows "gfinprod G f (insert i A) = (if i \<in> A then gfinprod G f A else f i \<otimes> gfinprod G f A)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 49 | proof - | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 50 |   have f: "f \<in> {x \<in> A. f x \<noteq> \<one>} \<rightarrow> carrier G"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 51 | using assms by (auto simp: image_subset_iff_funcset) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 52 |   have "{x. x = i \<and> f x \<noteq> \<one> \<or> x \<in> A \<and> f x \<noteq> \<one>} = (if f i = \<one> then {x \<in> A. f x \<noteq> \<one>} else insert i {x \<in> A. f x \<noteq> \<one>})"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 53 | by auto | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 54 | then show ?thesis | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 55 | using assms | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 56 | unfolding gfinprod_def by (simp add: conj_disj_distribR insert_absorb f split: if_split_asm) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 57 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 58 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 59 | lemma gfinprod_distrib: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 60 |   assumes fin: "finite {x \<in> A. f x \<noteq> \<one>\<^bsub>G\<^esub>}" "finite {x \<in> A. g x \<noteq> \<one>\<^bsub>G\<^esub>}"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 61 | and "f \<in> A \<rightarrow> carrier G" "g \<in> A \<rightarrow> carrier G" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 62 | shows "gfinprod G (\<lambda>i. f i \<otimes> g i) A = gfinprod G f A \<otimes> gfinprod G g A" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 63 | proof - | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 64 |   have "finite {x \<in> A. f x \<otimes> g x \<noteq> \<one>}"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 65 | by (auto intro: finite_subset [OF _ finite_UnI [OF fin]]) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 66 |   then have "gfinprod G (\<lambda>i. f i \<otimes> g i) A = gfinprod G (\<lambda>i. f i \<otimes> g i) ({i \<in> A. f i \<noteq> \<one>\<^bsub>G\<^esub>} \<union> {i \<in> A. g i \<noteq> \<one>\<^bsub>G\<^esub>})"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 67 | unfolding gfinprod_def | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 68 | using assms by (force intro: finprod_mono_neutral_cong) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 69 | also have "\<dots> = gfinprod G f A \<otimes> gfinprod G g A" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 70 | proof - | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 71 |     have "finprod G f ({i \<in> A. f i \<noteq> \<one>\<^bsub>G\<^esub>} \<union> {i \<in> A. g i \<noteq> \<one>\<^bsub>G\<^esub>}) = gfinprod G f A"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 72 |       "finprod G g ({i \<in> A. f i \<noteq> \<one>\<^bsub>G\<^esub>} \<union> {i \<in> A. g i \<noteq> \<one>\<^bsub>G\<^esub>}) = gfinprod G g A"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 73 | using assms by (auto simp: gfinprod_def intro: finprod_mono_neutral_cong_right) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 74 |     moreover have "(\<lambda>i. f i \<otimes> g i) \<in> {i \<in> A. f i \<noteq> \<one>} \<union> {i \<in> A. g i \<noteq> \<one>} \<rightarrow> carrier G"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 75 | using assms by (force simp: image_subset_iff_funcset) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 76 | ultimately show ?thesis | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 77 | using assms | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 78 | apply simp | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 79 | apply (subst finprod_multf, auto) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 80 | done | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 81 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 82 | finally show ?thesis . | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 83 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 84 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 85 | lemma gfinprod_mono_neutral_cong_left: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 86 | assumes "A \<subseteq> B" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 87 | and 1: "\<And>i. i \<in> B - A \<Longrightarrow> h i = \<one>" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 88 | and gh: "\<And>x. x \<in> A \<Longrightarrow> g x = h x" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 89 | and h: "h \<in> B \<rightarrow> carrier G" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 90 | shows "gfinprod G g A = gfinprod G h B" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 91 | proof (cases "finite {x \<in> B. h x \<noteq> \<one>}")
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 92 | case True | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 93 |   then have "finite {x \<in> A. h x \<noteq> \<one>}"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 94 | apply (rule rev_finite_subset) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 95 | using \<open>A \<subseteq> B\<close> by auto | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 96 | with True assms show ?thesis | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 97 | apply (simp add: gfinprod_def cong: conj_cong) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 98 | apply (auto intro!: finprod_mono_neutral_cong_left) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 99 | done | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 100 | next | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 101 | case False | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 102 |   have "{x \<in> B. h x \<noteq> \<one>} \<subseteq> {x \<in> A. h x \<noteq> \<one>}"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 103 | using 1 by auto | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 104 |   with False have "infinite {x \<in> A. h x \<noteq> \<one>}"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 105 | using infinite_super by blast | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 106 | with False assms show ?thesis | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 107 | by (simp add: gfinprod_def cong: conj_cong) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 108 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 109 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 110 | lemma gfinprod_mono_neutral_cong_right: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 111 | assumes "A \<subseteq> B" "\<And>i. i \<in> B - A \<Longrightarrow> g i = \<one>" "\<And>x. x \<in> A \<Longrightarrow> g x = h x" "g \<in> B \<rightarrow> carrier G" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 112 | shows "gfinprod G g B = gfinprod G h A" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 113 | using assms by (auto intro!: gfinprod_mono_neutral_cong_left [symmetric]) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 114 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 115 | lemma gfinprod_mono_neutral_cong: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 116 | assumes [simp]: "finite B" "finite A" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 117 | and *: "\<And>i. i \<in> B - A \<Longrightarrow> h i = \<one>" "\<And>i. i \<in> A - B \<Longrightarrow> g i = \<one>" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 118 | and gh: "\<And>x. x \<in> A \<inter> B \<Longrightarrow> g x = h x" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 119 | and g: "g \<in> A \<rightarrow> carrier G" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 120 | and h: "h \<in> B \<rightarrow> carrier G" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 121 | shows "gfinprod G g A = gfinprod G h B" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 122 | proof- | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 123 | have "gfinprod G g A = gfinprod G g (A \<inter> B)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 124 | by (rule gfinprod_mono_neutral_cong_right) (use assms in auto) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 125 | also have "\<dots> = gfinprod G h (A \<inter> B)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 126 | by (rule gfinprod_cong) (use assms in auto) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 127 | also have "\<dots> = gfinprod G h B" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 128 | by (rule gfinprod_mono_neutral_cong_left) (use assms in auto) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 129 | finally show ?thesis . | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 130 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 131 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 132 | end | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 133 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 134 | lemma (in comm_group) hom_group_sum: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 135 | assumes hom: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> hom (A i) G" and grp: "\<And>i. i \<in> I \<Longrightarrow> group (A i)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 136 | shows "(\<lambda>x. gfinprod G (\<lambda>i. (f i) (x i)) I) \<in> hom (sum_group I A) G" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 137 | unfolding hom_def | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 138 | proof (intro CollectI conjI ballI) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 139 | show "(\<lambda>x. gfinprod G (\<lambda>i. f i (x i)) I) \<in> carrier (sum_group I A) \<rightarrow> carrier G" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 140 | using assms | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 141 | by (force simp: hom_def carrier_sum_group intro: gfinprod_closed simp flip: image_subset_iff_funcset) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 142 | next | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 143 | fix x y | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 144 | assume x: "x \<in> carrier (sum_group I A)" and y: "y \<in> carrier (sum_group I A)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 145 |   then have finx: "finite {i \<in> I. x i \<noteq> \<one>\<^bsub>A i\<^esub>}" and finy: "finite {i \<in> I. y i \<noteq> \<one>\<^bsub>A i\<^esub>}"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 146 | using assms by (auto simp: carrier_sum_group) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 147 |   have finfx: "finite {i \<in> I. f i (x i) \<noteq> \<one>}"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 148 | using assms by (auto simp: is_group hom_one [OF hom] intro: finite_subset [OF _ finx]) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 149 |   have finfy: "finite {i \<in> I. f i (y i) \<noteq> \<one>}"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 150 | using assms by (auto simp: is_group hom_one [OF hom] intro: finite_subset [OF _ finy]) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 151 | have carr: "f i (x i) \<in> carrier G" "f i (y i) \<in> carrier G" if "i \<in> I" for i | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 152 | using hom_carrier [OF hom] that x y assms | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 153 | by (fastforce simp add: carrier_sum_group)+ | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 154 | have lam: "(\<lambda>i. f i ( x i \<otimes>\<^bsub>A i\<^esub> y i)) \<in> I \<rightarrow> carrier G" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 155 | using x y assms by (auto simp: hom_def carrier_sum_group PiE_def Pi_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 156 | have lam': "(\<lambda>i. f i (if i \<in> I then x i \<otimes>\<^bsub>A i\<^esub> y i else undefined)) \<in> I \<rightarrow> carrier G" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 157 | by (simp add: lam Pi_cong) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 158 | with lam x y assms | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 159 | show "gfinprod G (\<lambda>i. f i ((x \<otimes>\<^bsub>sum_group I A\<^esub> y) i)) I | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 160 | = gfinprod G (\<lambda>i. f i (x i)) I \<otimes> gfinprod G (\<lambda>i. f i (y i)) I" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 161 | by (simp add: carrier_sum_group PiE_def Pi_def hom_mult [OF hom] | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 162 | gfinprod_distrib finfx finfy carr cong: gfinprod_cong) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 163 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 164 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 165 | subsection\<open>Free Abelian groups on a set, using the "frag" type constructor. \<close> | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 166 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 167 | definition free_Abelian_group :: "'a set \<Rightarrow> ('a \<Rightarrow>\<^sub>0 int) monoid"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 168 |   where "free_Abelian_group S = \<lparr>carrier = {c. Poly_Mapping.keys c \<subseteq> S}, monoid.mult = (+), one  = 0\<rparr>"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 169 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 170 | lemma group_free_Abelian_group [simp]: "group (free_Abelian_group S)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 171 | proof - | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 172 | have "\<And>x. Poly_Mapping.keys x \<subseteq> S \<Longrightarrow> x \<in> Units (free_Abelian_group S)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 173 | unfolding free_Abelian_group_def Units_def | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 174 | by clarsimp (metis eq_neg_iff_add_eq_0 neg_eq_iff_add_eq_0 keys_minus) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 175 | then show ?thesis | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 176 | unfolding free_Abelian_group_def | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 177 | by unfold_locales (auto simp: dest: subsetD [OF keys_add]) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 178 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 179 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 180 | lemma carrier_free_Abelian_group_iff [simp]: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 181 | shows "x \<in> carrier (free_Abelian_group S) \<longleftrightarrow> Poly_Mapping.keys x \<subseteq> S" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 182 | by (auto simp: free_Abelian_group_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 183 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 184 | lemma one_free_Abelian_group [simp]: "\<one>\<^bsub>free_Abelian_group S\<^esub> = 0" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 185 | by (auto simp: free_Abelian_group_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 186 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 187 | lemma mult_free_Abelian_group [simp]: "x \<otimes>\<^bsub>free_Abelian_group S\<^esub> y = x + y" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 188 | by (auto simp: free_Abelian_group_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 189 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 190 | lemma inv_free_Abelian_group [simp]: "Poly_Mapping.keys x \<subseteq> S \<Longrightarrow> inv\<^bsub>free_Abelian_group S\<^esub> x = -x" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 191 | by (rule group.inv_equality [OF group_free_Abelian_group]) auto | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 192 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 193 | lemma abelian_free_Abelian_group: "comm_group(free_Abelian_group S)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 194 | apply (rule group.group_comm_groupI [OF group_free_Abelian_group]) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 195 | by (simp add: free_Abelian_group_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 196 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 197 | lemma pow_free_Abelian_group [simp]: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 198 | fixes n::nat | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 199 | shows "Group.pow (free_Abelian_group S) x n = frag_cmul (int n) x" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 200 | by (induction n) (auto simp: nat_pow_def free_Abelian_group_def frag_cmul_distrib) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 201 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 202 | lemma int_pow_free_Abelian_group [simp]: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 203 | fixes n::int | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 204 | assumes "Poly_Mapping.keys x \<subseteq> S" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 205 | shows "Group.pow (free_Abelian_group S) x n = frag_cmul n x" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 206 | proof (induction n) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 207 | case (nonneg n) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 208 | then show ?case | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 209 | by (simp add: int_pow_int) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 210 | next | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 211 | case (neg n) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 212 | have "x [^]\<^bsub>free_Abelian_group S\<^esub> - int (Suc n) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 213 | = inv\<^bsub>free_Abelian_group S\<^esub> (x [^]\<^bsub>free_Abelian_group S\<^esub> int (Suc n))" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 214 | by (rule group.int_pow_neg [OF group_free_Abelian_group]) (use assms in \<open>simp add: free_Abelian_group_def\<close>) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 215 | also have "\<dots> = frag_cmul (- int (Suc n)) x" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 216 | by (metis assms inv_free_Abelian_group pow_free_Abelian_group int_pow_int minus_frag_cmul | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 217 | order_trans keys_cmul) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 218 | finally show ?case . | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 219 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 220 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 221 | lemma frag_of_in_free_Abelian_group [simp]: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 222 | "frag_of x \<in> carrier(free_Abelian_group S) \<longleftrightarrow> x \<in> S" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 223 | by simp | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 224 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 225 | lemma free_Abelian_group_induct: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 226 | assumes major: "Poly_Mapping.keys x \<subseteq> S" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 227 | and minor: "P(0)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 228 | "\<And>x y. \<lbrakk>Poly_Mapping.keys x \<subseteq> S; Poly_Mapping.keys y \<subseteq> S; P x; P y\<rbrakk> \<Longrightarrow> P(x-y)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 229 | "\<And>a. a \<in> S \<Longrightarrow> P(frag_of a)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 230 | shows "P x" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 231 | proof - | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 232 | have "Poly_Mapping.keys x \<subseteq> S \<and> P x" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 233 | using major | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 234 | proof (induction x rule: frag_induction) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 235 | case (diff a b) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 236 | then show ?case | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 237 | by (meson Un_least minor(2) order.trans keys_diff) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 238 | qed (auto intro: minor) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 239 | then show ?thesis .. | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 240 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 241 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 242 | lemma sum_closed_free_Abelian_group: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 243 | "(\<And>i. i \<in> I \<Longrightarrow> x i \<in> carrier (free_Abelian_group S)) \<Longrightarrow> sum x I \<in> carrier (free_Abelian_group S)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 244 | apply (induction I rule: infinite_finite_induct, auto) | 
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
73348diff
changeset | 245 | by (metis (no_types, opaque_lifting) UnE subsetCE keys_add) | 
| 70045 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 246 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 247 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 248 | lemma (in comm_group) free_Abelian_group_universal: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 249 | fixes f :: "'c \<Rightarrow> 'a" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 250 | assumes "f ` S \<subseteq> carrier G" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 251 | obtains h where "h \<in> hom (free_Abelian_group S) G" "\<And>x. x \<in> S \<Longrightarrow> h(frag_of x) = f x" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 252 | proof | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 253 |   have fin: "Poly_Mapping.keys u \<subseteq> S \<Longrightarrow> finite {x \<in> S. f x [^] poly_mapping.lookup u x \<noteq> \<one>}" for u :: "'c \<Rightarrow>\<^sub>0 int"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 254 | apply (rule finite_subset [OF _ finite_keys [of u]]) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 255 | unfolding keys.rep_eq by force | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 256 |   define h :: "('c \<Rightarrow>\<^sub>0 int) \<Rightarrow> 'a"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 257 | where "h \<equiv> \<lambda>x. gfinprod G (\<lambda>a. f a [^] poly_mapping.lookup x a) S" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 258 | show "h \<in> hom (free_Abelian_group S) G" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 259 | proof (rule homI) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 260 | fix x y | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 261 | assume xy: "x \<in> carrier (free_Abelian_group S)" "y \<in> carrier (free_Abelian_group S)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 262 | then show "h (x \<otimes>\<^bsub>free_Abelian_group S\<^esub> y) = h x \<otimes> h y" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 263 | using assms unfolding h_def free_Abelian_group_def | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 264 | by (simp add: fin gfinprod_distrib image_subset_iff Poly_Mapping.lookup_add int_pow_mult cong: gfinprod_cong) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 265 | qed (use assms in \<open>force simp: free_Abelian_group_def h_def intro: gfinprod_closed\<close>) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 266 | show "h(frag_of x) = f x" if "x \<in> S" for x | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 267 | proof - | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 268 |     have fin: "(\<lambda>a. f x [^] (1::int)) \<in> {x} \<rightarrow> carrier G" "f x [^] (1::int) \<in> carrier G"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 269 | using assms that by force+ | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 270 | show ?thesis | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 271 | by (cases " f x [^] (1::int) = \<one>") (use assms that in \<open>auto simp: h_def gfinprod_def finprod_singleton\<close>) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 272 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 273 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 274 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 275 | lemma eqpoll_free_Abelian_group_infinite: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 276 | assumes "infinite A" shows "carrier(free_Abelian_group A) \<approx> A" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 277 | proof (rule lepoll_antisym) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 278 |   have "carrier (free_Abelian_group A) \<lesssim> {f::'a\<Rightarrow>int. f ` A \<subseteq> UNIV \<and> {x. f x \<noteq> 0} \<subseteq> A \<and> finite {x. f x \<noteq> 0}}"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 279 | unfolding lepoll_def | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 280 | by (rule_tac x="Poly_Mapping.lookup" in exI) (auto simp: poly_mapping_eqI lookup_not_eq_zero_eq_in_keys inj_onI) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 281 | also have "\<dots> \<lesssim> Fpow (A \<times> (UNIV::int set))" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 282 | by (rule lepoll_restricted_funspace) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 283 | also have "\<dots> \<approx> A \<times> (UNIV::int set)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 284 | proof (rule eqpoll_Fpow) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 285 | show "infinite (A \<times> (UNIV::int set))" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 286 | using assms finite_cartesian_productD1 by fastforce | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 287 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 288 | also have "\<dots> \<approx> A" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 289 | unfolding eqpoll_iff_card_of_ordIso | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 290 | proof - | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 291 | have "|A \<times> (UNIV::int set)| <=o |A|" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 292 | by (simp add: assms card_of_Times_ordLeq_infinite flip: infinite_iff_card_of_countable) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 293 | moreover have "|A| \<le>o |A \<times> (UNIV::int set)|" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 294 | by simp | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 295 | ultimately have "|A| *c |(UNIV::int set)| =o |A|" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 296 | by (simp add: cprod_def ordIso_iff_ordLeq) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 297 | then show "|A \<times> (UNIV::int set)| =o |A|" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 298 | by (metis Times_cprod ordIso_transitive) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 299 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 300 | finally show "carrier (free_Abelian_group A) \<lesssim> A" . | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 301 | have "inj_on frag_of A" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 302 | by (simp add: frag_of_eq inj_on_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 303 | moreover have "frag_of ` A \<subseteq> carrier (free_Abelian_group A)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 304 | by (simp add: image_subsetI) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 305 | ultimately show "A \<lesssim> carrier (free_Abelian_group A)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 306 | by (force simp: lepoll_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 307 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 308 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 309 | proposition (in comm_group) eqpoll_homomorphisms_from_free_Abelian_group: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 310 |    "{f. f \<in> extensional (carrier(free_Abelian_group S)) \<and> f \<in> hom (free_Abelian_group S) G}
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 311 | \<approx> (S \<rightarrow>\<^sub>E carrier G)" (is "?lhs \<approx> ?rhs") | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 312 | unfolding eqpoll_def bij_betw_def | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 313 | proof (intro exI conjI) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 314 | let ?f = "\<lambda>f. restrict (f \<circ> frag_of) S" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 315 | show "inj_on ?f ?lhs" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 316 | proof (clarsimp simp: inj_on_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 317 | fix g h | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 318 | assume | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 319 | g: "g \<in> extensional (carrier (free_Abelian_group S))" "g \<in> hom (free_Abelian_group S) G" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 320 | and h: "h \<in> extensional (carrier (free_Abelian_group S))" "h \<in> hom (free_Abelian_group S) G" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 321 | and eq: "restrict (g \<circ> frag_of) S = restrict (h \<circ> frag_of) S" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 322 | have 0: "0 \<in> carrier (free_Abelian_group S)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 323 | by simp | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 324 | interpret hom_g: group_hom "free_Abelian_group S" G g | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 325 | using g by (auto simp: group_hom_def group_hom_axioms_def is_group) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 326 | interpret hom_h: group_hom "free_Abelian_group S" G h | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 327 | using h by (auto simp: group_hom_def group_hom_axioms_def is_group) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 328 | have "Poly_Mapping.keys c \<subseteq> S \<Longrightarrow> Poly_Mapping.keys c \<subseteq> S \<and> g c = h c" for c | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 329 | proof (induction c rule: frag_induction) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 330 | case zero | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 331 | show ?case | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 332 | using hom_g.hom_one hom_h.hom_one by auto | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 333 | next | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 334 | case (one x) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 335 | then show ?case | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 336 | using eq by (simp add: fun_eq_iff) (metis comp_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 337 | next | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 338 | case (diff a b) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 339 | then show ?case | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 340 | using hom_g.hom_mult hom_h.hom_mult hom_g.hom_inv hom_h.hom_inv | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 341 | apply (auto simp: dest: subsetD [OF keys_diff]) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 342 | by (metis keys_minus uminus_add_conv_diff) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 343 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 344 | then show "g = h" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 345 | by (meson g h carrier_free_Abelian_group_iff extensionalityI) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 346 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 347 | have "f \<in> (\<lambda>f. restrict (f \<circ> frag_of) S) ` | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 348 |             {f \<in> extensional (carrier (free_Abelian_group S)). f \<in> hom (free_Abelian_group S) G}"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 349 | if f: "f \<in> S \<rightarrow>\<^sub>E carrier G" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 350 | for f :: "'c \<Rightarrow> 'a" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 351 | proof - | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 352 | obtain h where h: "h \<in> hom (free_Abelian_group S) G" "\<And>x. x \<in> S \<Longrightarrow> h(frag_of x) = f x" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 353 | proof (rule free_Abelian_group_universal) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 354 | show "f ` S \<subseteq> carrier G" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 355 | using f by blast | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 356 | qed auto | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 357 | let ?h = "restrict h (carrier (free_Abelian_group S))" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 358 | show ?thesis | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 359 | proof | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 360 | show "f = restrict (?h \<circ> frag_of) S" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 361 | using f by (force simp: h) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 362 |       show "?h \<in> {f \<in> extensional (carrier (free_Abelian_group S)). f \<in> hom (free_Abelian_group S) G}"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 363 | using h by (auto simp: hom_def dest!: subsetD [OF keys_add]) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 364 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 365 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 366 | then show "?f ` ?lhs = S \<rightarrow>\<^sub>E carrier G" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 367 | by (auto simp: hom_def Ball_def Pi_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 368 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 369 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 370 | lemma hom_frag_minus: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 371 | assumes "h \<in> hom (free_Abelian_group S) (free_Abelian_group T)" "Poly_Mapping.keys a \<subseteq> S" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 372 | shows "h (-a) = - (h a)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 373 | proof - | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 374 | have "Poly_Mapping.keys (h a) \<subseteq> T" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 375 | by (meson assms carrier_free_Abelian_group_iff hom_in_carrier) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 376 | then show ?thesis | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 377 | by (metis (no_types) assms carrier_free_Abelian_group_iff group_free_Abelian_group group_hom.hom_inv group_hom_axioms_def group_hom_def inv_free_Abelian_group) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 378 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 379 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 380 | lemma hom_frag_add: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 381 | assumes "h \<in> hom (free_Abelian_group S) (free_Abelian_group T)" "Poly_Mapping.keys a \<subseteq> S" "Poly_Mapping.keys b \<subseteq> S" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 382 | shows "h (a+b) = h a + h b" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 383 | proof - | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 384 | have "Poly_Mapping.keys (h a) \<subseteq> T" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 385 | by (meson assms carrier_free_Abelian_group_iff hom_in_carrier) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 386 | moreover | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 387 | have "Poly_Mapping.keys (h b) \<subseteq> T" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 388 | by (meson assms carrier_free_Abelian_group_iff hom_in_carrier) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 389 | ultimately show ?thesis | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 390 | using assms hom_mult by fastforce | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 391 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 392 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 393 | lemma hom_frag_diff: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 394 | assumes "h \<in> hom (free_Abelian_group S) (free_Abelian_group T)" "Poly_Mapping.keys a \<subseteq> S" "Poly_Mapping.keys b \<subseteq> S" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 395 | shows "h (a-b) = h a - h b" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 396 | by (metis (no_types, lifting) assms diff_conv_add_uminus hom_frag_add hom_frag_minus keys_minus) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 397 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 398 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 399 | proposition isomorphic_free_Abelian_groups: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 400 | "free_Abelian_group S \<cong> free_Abelian_group T \<longleftrightarrow> S \<approx> T" (is "(?FS \<cong> ?FT) = ?rhs") | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 401 | proof | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 402 | interpret S: group "?FS" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 403 | by simp | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 404 | interpret T: group "?FT" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 405 | by simp | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 406 | interpret G2: comm_group "integer_mod_group 2" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 407 | by (rule abelian_integer_mod_group) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 408 |   let ?Two = "{0..<2::int}"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 409 |   have [simp]: "\<not> ?Two \<subseteq> {a}" for a
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 410 | by (simp add: subset_iff) presburger | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 411 | assume L: "?FS \<cong> ?FT" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 412 |   let ?HS = "{h \<in> extensional (carrier ?FS). h \<in> hom ?FS (integer_mod_group 2)}"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 413 |   let ?HT = "{h \<in> extensional (carrier ?FT). h \<in> hom ?FT (integer_mod_group 2)}"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 414 | have "S \<rightarrow>\<^sub>E ?Two \<approx> ?HS" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 415 | apply (rule eqpoll_sym) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 416 | using G2.eqpoll_homomorphisms_from_free_Abelian_group by (simp add: carrier_integer_mod_group) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 417 | also have "\<dots> \<approx> ?HT" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 418 | proof - | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 419 | obtain f g where "group_isomorphisms ?FS ?FT f g" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 420 | using L S.iso_iff_group_isomorphisms by (force simp: is_iso_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 421 | then have f: "f \<in> hom ?FS ?FT" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 422 | and g: "g \<in> hom ?FT ?FS" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 423 | and gf: "\<forall>x \<in> carrier ?FS. g(f x) = x" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 424 | and fg: "\<forall>y \<in> carrier ?FT. f(g y) = y" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 425 | by (auto simp: group_isomorphisms_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 426 | let ?f = "\<lambda>h. restrict (h \<circ> g) (carrier ?FT)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 427 | let ?g = "\<lambda>h. restrict (h \<circ> f) (carrier ?FS)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 428 | show ?thesis | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 429 | proof (rule lepoll_antisym) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 430 | show "?HS \<lesssim> ?HT" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 431 | unfolding lepoll_def | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 432 | proof (intro exI conjI) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 433 | show "inj_on ?f ?HS" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 434 | apply (rule inj_on_inverseI [where g = ?g]) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 435 | using hom_in_carrier [OF f] | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 436 | by (auto simp: gf fun_eq_iff carrier_integer_mod_group Ball_def Pi_def extensional_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 437 | show "?f ` ?HS \<subseteq> ?HT" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 438 | proof clarsimp | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 439 | fix h | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 440 | assume h: "h \<in> hom ?FS (integer_mod_group 2)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 441 | have "h \<circ> g \<in> hom ?FT (integer_mod_group 2)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 442 | by (rule hom_compose [OF g h]) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 443 | moreover have "restrict (h \<circ> g) (carrier ?FT) x = (h \<circ> g) x" if "x \<in> carrier ?FT" for x | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 444 | using g that by (simp add: hom_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 445 | ultimately show "restrict (h \<circ> g) (carrier ?FT) \<in> hom ?FT (integer_mod_group 2)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 446 | using T.hom_restrict by fastforce | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 447 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 448 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 449 | next | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 450 | show "?HT \<lesssim> ?HS" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 451 | unfolding lepoll_def | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 452 | proof (intro exI conjI) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 453 | show "inj_on ?g ?HT" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 454 | apply (rule inj_on_inverseI [where g = ?f]) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 455 | using hom_in_carrier [OF g] | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 456 | by (auto simp: fg fun_eq_iff carrier_integer_mod_group Ball_def Pi_def extensional_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 457 | show "?g ` ?HT \<subseteq> ?HS" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 458 | proof clarsimp | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 459 | fix k | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 460 | assume k: "k \<in> hom ?FT (integer_mod_group 2)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 461 | have "k \<circ> f \<in> hom ?FS (integer_mod_group 2)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 462 | by (rule hom_compose [OF f k]) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 463 | moreover have "restrict (k \<circ> f) (carrier ?FS) x = (k \<circ> f) x" if "x \<in> carrier ?FS" for x | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 464 | using f that by (simp add: hom_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 465 | ultimately show "restrict (k \<circ> f) (carrier ?FS) \<in> hom ?FS (integer_mod_group 2)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 466 | using S.hom_restrict by fastforce | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 467 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 468 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 469 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 470 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 471 | also have "\<dots> \<approx> T \<rightarrow>\<^sub>E ?Two" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 472 | using G2.eqpoll_homomorphisms_from_free_Abelian_group by (simp add: carrier_integer_mod_group) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 473 | finally have *: "S \<rightarrow>\<^sub>E ?Two \<approx> T \<rightarrow>\<^sub>E ?Two" . | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 474 | then have "finite (S \<rightarrow>\<^sub>E ?Two) \<longleftrightarrow> finite (T \<rightarrow>\<^sub>E ?Two)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 475 | by (rule eqpoll_finite_iff) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 476 | then have "finite S \<longleftrightarrow> finite T" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 477 | by (auto simp: finite_funcset_iff) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 478 | then consider "finite S" "finite T" | "~ finite S" "~ finite T" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 479 | by blast | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 480 | then show ?rhs | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 481 | proof cases | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 482 | case 1 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 483 | with * have "2 ^ card S = (2::nat) ^ card T" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 484 | by (simp add: card_PiE finite_PiE eqpoll_iff_card) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 485 | then have "card S = card T" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 486 | by auto | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 487 | then show ?thesis | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 488 | using eqpoll_iff_card 1 by blast | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 489 | next | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 490 | case 2 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 491 | have "carrier (free_Abelian_group S) \<approx> carrier (free_Abelian_group T)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 492 | using L by (simp add: iso_imp_eqpoll_carrier) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 493 | then show ?thesis | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 494 | using 2 eqpoll_free_Abelian_group_infinite eqpoll_sym eqpoll_trans by metis | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 495 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 496 | next | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 497 | assume ?rhs | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 498 | then obtain f g where f: "\<And>x. x \<in> S \<Longrightarrow> f x \<in> T \<and> g(f x) = x" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 499 | and g: "\<And>y. y \<in> T \<Longrightarrow> g y \<in> S \<and> f(g y) = y" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 500 | using eqpoll_iff_bijections by metis | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 501 | interpret S: comm_group "?FS" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 502 | by (simp add: abelian_free_Abelian_group) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 503 | interpret T: comm_group "?FT" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 504 | by (simp add: abelian_free_Abelian_group) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 505 | have "(frag_of \<circ> f) ` S \<subseteq> carrier (free_Abelian_group T)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 506 | using f by auto | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 507 | then obtain h where h: "h \<in> hom (free_Abelian_group S) (free_Abelian_group T)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 508 | and h_frag: "\<And>x. x \<in> S \<Longrightarrow> h (frag_of x) = (frag_of \<circ> f) x" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 509 | using T.free_Abelian_group_universal [of "frag_of \<circ> f" S] by blast | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 510 | interpret hhom: group_hom "free_Abelian_group S" "free_Abelian_group T" h | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 511 | by (simp add: h group_hom_axioms_def group_hom_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 512 | have "(frag_of \<circ> g) ` T \<subseteq> carrier (free_Abelian_group S)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 513 | using g by auto | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 514 | then obtain k where k: "k \<in> hom (free_Abelian_group T) (free_Abelian_group S)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 515 | and k_frag: "\<And>x. x \<in> T \<Longrightarrow> k (frag_of x) = (frag_of \<circ> g) x" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 516 | using S.free_Abelian_group_universal [of "frag_of \<circ> g" T] by blast | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 517 | interpret khom: group_hom "free_Abelian_group T" "free_Abelian_group S" k | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 518 | by (simp add: k group_hom_axioms_def group_hom_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 519 | have kh: "Poly_Mapping.keys x \<subseteq> S \<Longrightarrow> Poly_Mapping.keys x \<subseteq> S \<and> k (h x) = x" for x | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 520 | proof (induction rule: frag_induction) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 521 | case zero | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 522 | then show ?case | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 523 | apply auto | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 524 | by (metis group_free_Abelian_group h hom_one k one_free_Abelian_group) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 525 | next | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 526 | case (one x) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 527 | then show ?case | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 528 | by (auto simp: h_frag k_frag f) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 529 | next | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 530 | case (diff a b) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 531 | with keys_diff have "Poly_Mapping.keys (a - b) \<subseteq> S" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 532 | by (metis Un_least order_trans) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 533 | with diff hhom.hom_closed show ?case | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 534 | by (simp add: hom_frag_diff [OF h] hom_frag_diff [OF k]) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 535 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 536 | have hk: "Poly_Mapping.keys y \<subseteq> T \<Longrightarrow> Poly_Mapping.keys y \<subseteq> T \<and> h (k y) = y" for y | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 537 | proof (induction rule: frag_induction) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 538 | case zero | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 539 | then show ?case | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 540 | apply auto | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 541 | by (metis group_free_Abelian_group h hom_one k one_free_Abelian_group) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 542 | next | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 543 | case (one y) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 544 | then show ?case | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 545 | by (auto simp: h_frag k_frag g) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 546 | next | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 547 | case (diff a b) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 548 | with keys_diff have "Poly_Mapping.keys (a - b) \<subseteq> T" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 549 | by (metis Un_least order_trans) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 550 | with diff khom.hom_closed show ?case | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 551 | by (simp add: hom_frag_diff [OF h] hom_frag_diff [OF k]) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 552 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 553 | have "h \<in> iso ?FS ?FT" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 554 | unfolding iso_def bij_betw_iff_bijections mem_Collect_eq | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 555 | proof (intro conjI exI ballI h) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 556 | fix x | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 557 | assume x: "x \<in> carrier (free_Abelian_group S)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 558 | show "h x \<in> carrier (free_Abelian_group T)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 559 | by (meson x h hom_in_carrier) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 560 | show "k (h x) = x" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 561 | using x by (simp add: kh) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 562 | next | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 563 | fix y | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 564 | assume y: "y \<in> carrier (free_Abelian_group T)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 565 | show "k y \<in> carrier (free_Abelian_group S)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 566 | by (meson y k hom_in_carrier) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 567 | show "h (k y) = y" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 568 | using y by (simp add: hk) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 569 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 570 | then show "?FS \<cong> ?FT" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 571 | by (auto simp: is_iso_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 572 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 573 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 574 | lemma isomorphic_group_integer_free_Abelian_group_singleton: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 575 |   "integer_group \<cong> free_Abelian_group {x}"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 576 | proof - | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 577 |   have "(\<lambda>n. frag_cmul n (frag_of x)) \<in> iso integer_group (free_Abelian_group {x})"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 578 | proof (rule isoI [OF homI]) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 579 |     show "bij_betw (\<lambda>n. frag_cmul n (frag_of x)) (carrier integer_group) (carrier (free_Abelian_group {x}))"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 580 | apply (rule bij_betwI [where g = "\<lambda>y. Poly_Mapping.lookup y x"]) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 581 | by (auto simp: integer_group_def in_keys_iff intro!: poly_mapping_eqI) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 582 | qed (auto simp: frag_cmul_distrib) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 583 | then show ?thesis | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 584 | unfolding is_iso_def | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 585 | by blast | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 586 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 587 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 588 | lemma group_hom_free_Abelian_groups_id: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 589 | "id \<in> hom (free_Abelian_group S) (free_Abelian_group T) \<longleftrightarrow> S \<subseteq> T" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 590 | proof - | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 591 | have "x \<in> T" if ST: "\<And>c:: 'a \<Rightarrow>\<^sub>0 int. Poly_Mapping.keys c \<subseteq> S \<longrightarrow> Poly_Mapping.keys c \<subseteq> T" and "x \<in> S" for x | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 592 | using ST [of "frag_of x"] \<open>x \<in> S\<close> by simp | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 593 | then show ?thesis | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 594 | by (auto simp: hom_def free_Abelian_group_def Pi_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 595 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 596 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 597 | proposition iso_free_Abelian_group_sum: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 598 | assumes "pairwise (\<lambda>i j. disjnt (S i) (S j)) I" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 599 | shows "(\<lambda>f. sum' f I) \<in> iso (sum_group I (\<lambda>i. free_Abelian_group(S i))) (free_Abelian_group (\<Union>(S ` I)))" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 600 | (is "?h \<in> iso ?G ?H") | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 601 | proof (rule isoI) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 602 | show hom: "?h \<in> hom ?G ?H" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 603 | proof (rule homI) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 604 | show "?h c \<in> carrier ?H" if "c \<in> carrier ?G" for c | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 605 | using that | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 606 | apply (simp add: sum.G_def carrier_sum_group) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 607 | apply (rule order_trans [OF keys_sum]) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 608 | apply (auto simp: free_Abelian_group_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 609 | done | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 610 | show "?h (x \<otimes>\<^bsub>?G\<^esub> y) = ?h x \<otimes>\<^bsub>?H\<^esub> ?h y" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 611 | if "x \<in> carrier ?G" "y \<in> carrier ?G" for x y | 
| 70063 | 612 | using that by (simp add: sum.finite_Collect_op carrier_sum_group sum.distrib') | 
| 70045 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 613 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 614 | interpret GH: group_hom "?G" "?H" "?h" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 615 | using hom by (simp add: group_hom_def group_hom_axioms_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 616 | show "bij_betw ?h (carrier ?G) (carrier ?H)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 617 | unfolding bij_betw_def | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 618 | proof (intro conjI subset_antisym) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 619 | show "?h ` carrier ?G \<subseteq> carrier ?H" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 620 | apply (clarsimp simp: sum.G_def carrier_sum_group simp del: carrier_free_Abelian_group_iff) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 621 | by (force simp: PiE_def Pi_iff intro!: sum_closed_free_Abelian_group) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 622 | have *: "poly_mapping.lookup (Abs_poly_mapping (\<lambda>j. if j \<in> S i then poly_mapping.lookup x j else 0)) k | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 623 | = (if k \<in> S i then poly_mapping.lookup x k else 0)" if "i \<in> I" for i k and x :: "'b \<Rightarrow>\<^sub>0 int" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 624 | using that by (auto simp: conj_commute cong: conj_cong) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 625 | have eq: "Abs_poly_mapping (\<lambda>j. if j \<in> S i then poly_mapping.lookup x j else 0) = 0 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 626 | \<longleftrightarrow> (\<forall>c \<in> S i. poly_mapping.lookup x c = 0)" if "i \<in> I" for i and x :: "'b \<Rightarrow>\<^sub>0 int" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 627 | apply (auto simp: poly_mapping_eq_iff fun_eq_iff) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 628 | apply (simp add: * Abs_poly_mapping_inverse conj_commute cong: conj_cong) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 629 | apply (force dest!: spec split: if_split_asm) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 630 | done | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 631 |     have "x \<in> ?h ` {x \<in> \<Pi>\<^sub>E i\<in>I. {c. Poly_Mapping.keys c \<subseteq> S i}. finite {i \<in> I. x i \<noteq> 0}}"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 632 | if x: "Poly_Mapping.keys x \<subseteq> \<Union> (S ` I)" for x :: "'b \<Rightarrow>\<^sub>0 int" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 633 | proof - | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 634 | let ?f = "(\<lambda>i c. if c \<in> S i then Poly_Mapping.lookup x c else 0)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 635 |       define J where "J \<equiv> {i \<in> I. \<exists>c\<in>S i. c \<in> Poly_Mapping.keys x}"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 636 | have "J \<subseteq> (\<lambda>c. THE i. i \<in> I \<and> c \<in> S i) ` Poly_Mapping.keys x" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 637 | proof (clarsimp simp: J_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 638 | show "i \<in> (\<lambda>c. THE i. i \<in> I \<and> c \<in> S i) ` Poly_Mapping.keys x" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 639 | if "i \<in> I" "c \<in> S i" "c \<in> Poly_Mapping.keys x" for i c | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 640 | proof | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 641 | show "i = (THE i. i \<in> I \<and> c \<in> S i)" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 642 | using assms that by (auto simp: pairwise_def disjnt_def intro: the_equality [symmetric]) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 643 | qed (simp add: that) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 644 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 645 | then have fin: "finite J" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 646 | using finite_subset finite_keys by blast | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 647 |       have [simp]: "Poly_Mapping.keys (Abs_poly_mapping (?f i)) = {k. ?f i k \<noteq> 0}" if "i \<in> I" for i
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 648 | by (simp add: eq_onp_def keys.abs_eq conj_commute cong: conj_cong) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 649 | have [simp]: "Poly_Mapping.lookup (Abs_poly_mapping (?f i)) c = ?f i c" if "i \<in> I" for i c | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 650 | by (auto simp: Abs_poly_mapping_inverse conj_commute cong: conj_cong) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 651 | show ?thesis | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 652 | proof | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 653 | have "poly_mapping.lookup x c = poly_mapping.lookup (?h (\<lambda>i\<in>I. Abs_poly_mapping (?f i))) c" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 654 | for c | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 655 | proof (cases "c \<in> Poly_Mapping.keys x") | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 656 | case True | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 657 | then obtain i where "i \<in> I" "c \<in> S i" "?f i c \<noteq> 0" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 658 | using x by (auto simp: in_keys_iff) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 659 |           then have 1: "poly_mapping.lookup (sum' (\<lambda>j. Abs_poly_mapping (?f j)) (I - {i})) c = 0"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 660 | using assms | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 661 | apply (simp add: sum.G_def Poly_Mapping.lookup_sum pairwise_def disjnt_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 662 | apply (force simp: eq split: if_split_asm intro!: comm_monoid_add_class.sum.neutral) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 663 | done | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 664 | have 2: "poly_mapping.lookup x c = poly_mapping.lookup (Abs_poly_mapping (?f i)) c" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 665 | by (auto simp: \<open>c \<in> S i\<close> Abs_poly_mapping_inverse conj_commute cong: conj_cong) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 666 |           have "finite {i \<in> I. Abs_poly_mapping (?f i) \<noteq> 0}"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 667 | by (rule finite_subset [OF _ fin]) (use \<open>i \<in> I\<close> J_def eq in \<open>auto simp: in_keys_iff\<close>) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 668 |           with \<open>i \<in> I\<close> have "?h (\<lambda>j\<in>I. Abs_poly_mapping (?f j)) = Abs_poly_mapping (?f i) + sum' (\<lambda>j. Abs_poly_mapping (?f j)) (I - {i})"
 | 
| 70065 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70063diff
changeset | 669 | by (simp add: sum_diff1') | 
| 70045 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 670 | then show ?thesis | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 671 | by (simp add: 1 2 Poly_Mapping.lookup_add) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 672 | next | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 673 | case False | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 674 | then have "poly_mapping.lookup x c = 0" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 675 | using keys.rep_eq by force | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 676 | then show ?thesis | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 677 | unfolding sum.G_def by (simp add: lookup_sum * comm_monoid_add_class.sum.neutral) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 678 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 679 | then show "x = ?h (\<lambda>i\<in>I. Abs_poly_mapping (?f i))" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 680 | by (rule poly_mapping_eqI) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 681 |         have "(\<lambda>i. Abs_poly_mapping (?f i)) \<in> (\<Pi> i\<in>I. {c. Poly_Mapping.keys c \<subseteq> S i})"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 682 | by (auto simp: PiE_def Pi_def in_keys_iff) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 683 | then show "(\<lambda>i\<in>I. Abs_poly_mapping (?f i)) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 684 |                  \<in> {x \<in> \<Pi>\<^sub>E i\<in>I. {c. Poly_Mapping.keys c \<subseteq> S i}. finite {i \<in> I. x i \<noteq> 0}}"
 | 
| 73348 
65c45cba3f54
reverted simprule status on a new lemma
 paulson <lp15@cam.ac.uk> parents: 
72630diff
changeset | 685 | using fin unfolding J_def by (force simp add: eq in_keys_iff cong: conj_cong) | 
| 70045 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 686 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 687 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 688 | then show "carrier ?H \<subseteq> ?h ` carrier ?G" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 689 | by (simp add: carrier_sum_group) (auto simp: free_Abelian_group_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 690 | show "inj_on ?h (carrier (sum_group I (\<lambda>i. free_Abelian_group (S i))))" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 691 | unfolding GH.inj_on_one_iff | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 692 | proof clarify | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 693 | fix x | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 694 | assume "x \<in> carrier ?G" "?h x = \<one>\<^bsub>?H\<^esub>" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 695 | then have eq0: "sum' x I = 0" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 696 | and xs: "\<And>i. i \<in> I \<Longrightarrow> Poly_Mapping.keys (x i) \<subseteq> S i" and xext: "x \<in> extensional I" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 697 |         and fin: "finite {i \<in> I. x i \<noteq> 0}"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 698 | by (simp_all add: carrier_sum_group PiE_def Pi_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 699 | have "x i = 0" if "i \<in> I" for i | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 700 | proof - | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 701 |         have "sum' x (insert i (I - {i})) = 0"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 702 | using eq0 that by (simp add: insert_absorb) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 703 |         moreover have "Poly_Mapping.keys (sum' x (I - {i})) = {}"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 704 | proof - | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 705 |           have "x i = - sum' x (I - {i})"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 706 | by (metis (mono_tags, lifting) diff_zero eq0 fin sum_diff1' minus_diff_eq that) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 707 |           then have "Poly_Mapping.keys (x i) = Poly_Mapping.keys (sum' x (I - {i}))"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 708 | by simp | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 709 |           then have "Poly_Mapping.keys (sum' x (I - {i})) \<subseteq> S i"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 710 | using that xs by metis | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 711 | moreover | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 712 |           have "Poly_Mapping.keys (sum' x (I - {i})) \<subseteq> (\<Union>j \<in> I - {i}. S j)"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 713 | proof - | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 714 |             have "Poly_Mapping.keys (sum' x (I - {i})) \<subseteq> (\<Union>i\<in>{j \<in> I. j \<noteq> i \<and> x j \<noteq> 0}. Poly_Mapping.keys (x i))"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 715 |               using keys_sum [of x "{j \<in> I. j \<noteq> i \<and> x j \<noteq> 0}"] by (simp add: sum.G_def)
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 716 |             also have "\<dots> \<subseteq> \<Union> (S ` (I - {i}))"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 717 | using xs by force | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 718 | finally show ?thesis . | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 719 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 720 |           moreover have "A = {}" if "A \<subseteq> S i" "A \<subseteq> \<Union> (S ` (I - {i}))" for A
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 721 | using assms that \<open>i \<in> I\<close> | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 722 | by (force simp: pairwise_def disjnt_def image_def subset_iff) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 723 | ultimately show ?thesis | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 724 | by metis | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 725 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 726 |         then have [simp]: "sum' x (I - {i}) = 0"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 727 | by (auto simp: sum.G_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 728 |         have "sum' x (insert i (I - {i})) = x i"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 729 | by (subst sum.insert' [OF finite_subset [OF _ fin]]) auto | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 730 | ultimately show ?thesis | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 731 | by metis | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 732 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 733 | with xext [unfolded extensional_def] | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 734 | show "x = \<one>\<^bsub>sum_group I (\<lambda>i. free_Abelian_group (S i))\<^esub>" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 735 | by (force simp: free_Abelian_group_def) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 736 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 737 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 738 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 739 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 740 | lemma isomorphic_free_Abelian_group_Union: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 741 | "pairwise disjnt I \<Longrightarrow> free_Abelian_group(\<Union> I) \<cong> sum_group I free_Abelian_group" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 742 | using iso_free_Abelian_group_sum [of "\<lambda>X. X" I] | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 743 | by (metis SUP_identity_eq empty_iff group.iso_sym group_free_Abelian_group is_iso_def sum_group) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 744 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 745 | lemma isomorphic_sum_integer_group: | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 746 | "sum_group I (\<lambda>i. integer_group) \<cong> free_Abelian_group I" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 747 | proof - | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 748 |   have "sum_group I (\<lambda>i. integer_group) \<cong> sum_group I (\<lambda>i. free_Abelian_group {i})"
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 749 | by (rule iso_sum_groupI) (auto simp: isomorphic_group_integer_free_Abelian_group_singleton) | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 750 | also have "\<dots> \<cong> free_Abelian_group I" | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 751 |     using iso_free_Abelian_group_sum [of "\<lambda>x. {x}" I] by (auto simp: is_iso_def)
 | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 752 | finally show ?thesis . | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 753 | qed | 
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 754 | |
| 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 755 | end |