| author | paulson | 
| Wed, 17 Apr 2024 22:07:21 +0100 | |
| changeset 80130 | 8262d4f63b58 | 
| parent 69700 | 7a92cbec7030 | 
| permissions | -rw-r--r-- | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1 | (* Title: HOL/Algebra/Weak_Morphisms.thy | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2 | Author: Paulo EmÃlio de Vilhena | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3 | *) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5 | theory Weak_Morphisms | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 6 | imports QuotRing | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 7 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 8 | begin | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 9 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 10 | section \<open>Weak Morphisms\<close> | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 11 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 12 | text \<open>The definition of ring isomorphism, as well as the definition of group isomorphism, doesn't | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 13 | enforce any algebraic constraint to the structure of the schemes involved. This seems | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 14 | unnatural, but it turns out to be very useful: in order to prove that a scheme B satisfy | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 15 | certain algebraic constraints, it's sufficient to prove those for a scheme A and show | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 16 | the existence of an isomorphism between A and B. In this section, we explore this idea | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 17 | in a different way: given a scheme A and a function f, we build a scheme B with an | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 18 | algebraic structure of same strength as A where f is an homomorphism from A to B.\<close> | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 19 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 20 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 21 | subsection \<open>Definitions\<close> | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 22 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 23 | locale weak_group_morphism = normal H G for f and H and G (structure) + | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 24 | assumes inj_mod_subgroup: "\<lbrakk> a \<in> carrier G; b \<in> carrier G \<rbrakk> \<Longrightarrow> f a = f b \<longleftrightarrow> a \<otimes> (inv b) \<in> H" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 25 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 26 | locale weak_ring_morphism = ideal I R for f and I and R (structure) + | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 27 | assumes inj_mod_ideal: "\<lbrakk> a \<in> carrier R; b \<in> carrier R \<rbrakk> \<Longrightarrow> f a = f b \<longleftrightarrow> a \<ominus> b \<in> I" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 28 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 29 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 30 | definition image_group :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a, 'c) monoid_scheme \<Rightarrow> 'b monoid"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 31 | where "image_group f G \<equiv> | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 32 | \<lparr> carrier = f ` (carrier G), | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 33 | mult = (\<lambda>a b. f ((inv_into (carrier G) f a) \<otimes>\<^bsub>G\<^esub> (inv_into (carrier G) f b))), | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 34 | one = f \<one>\<^bsub>G\<^esub> \<rparr>" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 35 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 36 | definition image_ring :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a, 'c) ring_scheme \<Rightarrow> 'b ring"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 37 | where "image_ring f R \<equiv> monoid.extend (image_group f R) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 38 | \<lparr> zero = f \<zero>\<^bsub>R\<^esub>, | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 39 | add = (\<lambda>a b. f ((inv_into (carrier R) f a) \<oplus>\<^bsub>R\<^esub> (inv_into (carrier R) f b))) \<rparr>" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 40 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 41 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 42 | subsection \<open>Weak Group Morphisms\<close> | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 43 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 44 | lemma image_group_carrier: "carrier (image_group f G) = f ` (carrier G)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 45 | unfolding image_group_def by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 46 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 47 | lemma image_group_one: "one (image_group f G) = f \<one>\<^bsub>G\<^esub>" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 48 | unfolding image_group_def by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 49 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 50 | lemma weak_group_morphismsI: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 51 | assumes "H \<lhd> G" and "\<And>a b. \<lbrakk> a \<in> carrier G; b \<in> carrier G \<rbrakk> \<Longrightarrow> f a = f b \<longleftrightarrow> a \<otimes>\<^bsub>G\<^esub> (inv\<^bsub>G\<^esub> b) \<in> H" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 52 | shows "weak_group_morphism f H G" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 53 | using assms unfolding weak_group_morphism_def weak_group_morphism_axioms_def by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 54 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 55 | lemma image_group_truncate: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 56 |   fixes R :: "('a, 'b) monoid_scheme"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 57 | shows "monoid.truncate (image_group f R) = image_group f (monoid.truncate R)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 58 | by (simp add: image_group_def monoid.defs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 59 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 60 | lemma image_ring_truncate: "monoid.truncate (image_ring f R) = image_group f R" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 61 | by (simp add: image_ring_def monoid.defs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 62 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 63 | lemma (in ring) weak_add_group_morphism: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 64 | assumes "weak_ring_morphism f I R" shows "weak_group_morphism f I (add_monoid R)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 65 | proof - | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 66 | have is_normal: "I \<lhd> (add_monoid R)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 67 | using ideal_is_normal[OF weak_ring_morphism.axioms(1)[OF assms]] . | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 68 | show ?thesis | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 69 | using weak_group_morphism.intro[OF is_normal] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 70 | weak_ring_morphism.inj_mod_ideal[OF assms] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 71 | unfolding weak_group_morphism_axioms_def a_minus_def a_inv_def | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 72 | by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 73 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 74 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 75 | lemma (in group) weak_group_morphism_range: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 76 |   assumes "weak_group_morphism f H G" and "a \<in> carrier G" shows "f ` (H #> a) = { f a }"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 77 | proof - | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 78 | interpret H: subgroup H G | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 79 | using weak_group_morphism.axioms(1)[OF assms(1)] unfolding normal_def by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 80 | show ?thesis | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 81 | proof | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 82 |     show "{ f a } \<subseteq> f ` (H #> a)"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 83 | using H.one_closed assms(2) unfolding r_coset_def by force | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 84 | next | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 85 |     show "f ` (H #> a) \<subseteq> { f a }"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 86 | proof | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 87 | fix b assume "b \<in> f ` (H #> a)" then obtain h where "h \<in> H" "h \<in> carrier G" "b = f (h \<otimes> a)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 88 | unfolding r_coset_def using H.subset by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 89 |       thus "b \<in> { f a }"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 90 | using weak_group_morphism.inj_mod_subgroup[OF assms(1)] assms(2) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 91 | by (metis inv_solve_right m_closed singleton_iff) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 92 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 93 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 94 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 95 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 96 | lemma (in group) vimage_eq_rcoset: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 97 | assumes "weak_group_morphism f H G" and "a \<in> carrier G" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 98 |   shows "{ b \<in> carrier G. f b = f a } = H #> a" and "{ b \<in> carrier G. f b = f a } = a <# H"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 99 | proof - | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 100 | interpret H: normal H G | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 101 | using weak_group_morphism.axioms(1)[OF assms(1)] by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 102 |   show "{ b \<in> carrier G. f b = f a } = H #> a"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 103 | proof | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 104 |     show "H #> a \<subseteq> { b \<in> carrier G. f b = f a }"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 105 | using r_coset_subset_G[OF H.subset assms(2)] weak_group_morphism_range[OF assms] by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 106 | next | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 107 |     show "{ b \<in> carrier G. f b = f a } \<subseteq> H #> a"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 108 | proof | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 109 |       fix b assume b: "b \<in> { b \<in> carrier G. f b = f a }" then obtain h where "h \<in> H" "b \<otimes> (inv a) = h"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 110 | using weak_group_morphism.inj_mod_subgroup[OF assms(1)] assms(2) by fastforce | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 111 | thus "b \<in> H #> a" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 112 | using H.rcos_module[OF is_group] b assms(2) by blast | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 113 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 114 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 115 |   thus "{ b \<in> carrier G. f b = f a } = a <# H"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 116 | by (simp add: assms(2) H.coset_eq) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 117 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 118 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 119 | lemma (in group) weak_group_morphism_ker: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 120 | assumes "weak_group_morphism f H G" shows "kernel G (image_group f G) f = H" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 121 | using vimage_eq_rcoset(1)[OF assms one_closed] weak_group_morphism.axioms(1)[OF assms(1)] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 122 | by (simp add: image_group_def kernel_def normal_def subgroup.subset) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 123 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 124 | lemma (in group) weak_group_morphism_inv_into: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 125 | assumes "weak_group_morphism f H G" and "a \<in> carrier G" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 126 | obtains h h' where "h \<in> H" "inv_into (carrier G) f (f a) = h \<otimes> a" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 127 | and "h' \<in> H" "inv_into (carrier G) f (f a) = a \<otimes> h'" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 128 | proof - | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 129 |   have "inv_into (carrier G) f (f a) \<in> { b \<in> carrier G. f b = f a }"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 130 | using assms(2) by (auto simp add: inv_into_into f_inv_into_f) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 131 | thus thesis | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 132 | using that vimage_eq_rcoset[OF assms] unfolding r_coset_def l_coset_def by blast | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 133 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 134 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 135 | proposition (in group) weak_group_morphism_is_iso: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 136 | assumes "weak_group_morphism f H G" shows "(\<lambda>x. the_elem (f ` x)) \<in> iso (G Mod H) (image_group f G)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 137 | proof (auto simp add: iso_def hom_def image_group_def) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 138 | interpret H: normal H G | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 139 | using weak_group_morphism.axioms(1)[OF assms] . | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 140 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 141 | show "\<And>x. x \<in> carrier (G Mod H) \<Longrightarrow> the_elem (f ` x) \<in> f ` carrier G" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 142 | unfolding FactGroup_def RCOSETS_def using weak_group_morphism_range[OF assms] by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 143 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 144 | thus "bij_betw (\<lambda>x. the_elem (f ` x)) (carrier (G Mod H)) (f ` carrier G)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 145 | unfolding bij_betw_def | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 146 | proof (auto) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 147 | fix a assume "a \<in> carrier G" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 148 | hence "the_elem (f ` (H #> a)) = f a" and "H #> a \<in> carrier (G Mod H)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 149 | using weak_group_morphism_range[OF assms] unfolding FactGroup_def RCOSETS_def by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 150 | thus "f a \<in> (\<lambda>x. the_elem (f ` x)) ` carrier (G Mod H)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 151 | using image_iff by fastforce | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 152 | next | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 153 | show "inj_on (\<lambda>x. the_elem (f ` x)) (carrier (G Mod H))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 154 | proof (rule inj_onI) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 155 | fix x y assume "x \<in> (carrier (G Mod H))" and "y \<in> (carrier (G Mod H))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 156 | then obtain a b where a: "a \<in> carrier G" "x = H #> a" and b: "b \<in> carrier G" "y = H #> b" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 157 | unfolding FactGroup_def RCOSETS_def by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 158 | assume "the_elem (f ` x) = the_elem (f ` y)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 159 | hence "a \<otimes> (inv b) \<in> H" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 160 | using weak_group_morphism.inj_mod_subgroup[OF assms] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 161 | weak_group_morphism_range[OF assms] a b by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 162 | thus "x = y" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 163 | using a(1) b(1) unfolding a b | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 164 | by (meson H.rcos_const H.subset group.coset_mult_inv1 is_group) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 165 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 166 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 167 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 168 | fix x y assume "x \<in> carrier (G Mod H)" "y \<in> carrier (G Mod H)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 169 | then obtain a b where a: "a \<in> carrier G" "x = H #> a" and b: "b \<in> carrier G" "y = H #> b" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 170 | unfolding FactGroup_def RCOSETS_def by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 171 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 172 | show "the_elem (f ` (x <#> y)) = f (inv_into (carrier G) f (the_elem (f ` x)) \<otimes> | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 173 | inv_into (carrier G) f (the_elem (f ` y)))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 174 | proof (simp add: weak_group_morphism_range[OF assms] a b) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 175 | obtain h1 h2 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 176 | where h1: "h1 \<in> H" "inv_into (carrier G) f (f a) = a \<otimes> h1" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 177 | and h2: "h2 \<in> H" "inv_into (carrier G) f (f b) = h2 \<otimes> b" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 178 | using weak_group_morphism_inv_into[OF assms] a(1) b(1) by metis | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 179 | have "the_elem (f ` ((H #> a) <#> (H #> b))) = the_elem (f ` (H #> (a \<otimes> b)))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 180 | by (simp add: a b H.rcos_sum) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 181 | hence "the_elem (f ` ((H #> a) <#> (H #> b))) = f (a \<otimes> b)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 182 | using weak_group_morphism_range[OF assms] a(1) b(1) by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 183 | moreover | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 184 | have "(a \<otimes> h1) \<otimes> (h2 \<otimes> b) = a \<otimes> (h1 \<otimes> h2 \<otimes> b)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 185 | by (simp add: a(1) b(1) h1(1) h2(1) H.subset m_assoc) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 186 | hence "(a \<otimes> h1) \<otimes> (h2 \<otimes> b) \<in> a <# (H #> b)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 187 | using h1(1) h2(1) unfolding l_coset_def r_coset_def by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 188 | hence "(a \<otimes> h1) \<otimes> (h2 \<otimes> b) \<in> (a \<otimes> b) <# H" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 189 | by (simp add: H.subset H.coset_eq a(1) b(1) lcos_m_assoc) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 190 | hence "f (inv_into (carrier G) f (f a) \<otimes> inv_into (carrier G) f (f b)) = f (a \<otimes> b)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 191 | using vimage_eq_rcoset(2)[OF assms] a(1) b(1) unfolding h1 h2 by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 192 | ultimately | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 193 | show "the_elem (f ` ((H #> a) <#> (H #> b))) = f (inv_into (carrier G) f (f a) \<otimes> | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 194 | inv_into (carrier G) f (f b))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 195 | by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 196 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 197 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 198 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 199 | corollary (in group) image_group_is_group: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 200 | assumes "weak_group_morphism f H G" shows "group (image_group f G)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 201 | proof - | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 202 | interpret H: normal H G | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 203 | using weak_group_morphism.axioms(1)[OF assms] . | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 204 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 205 | have "group ((image_group f G) \<lparr> one := the_elem (f ` \<one>\<^bsub>G Mod H\<^esub>) \<rparr>)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 206 | using group.iso_imp_img_group[OF H.factorgroup_is_group weak_group_morphism_is_iso[OF assms]] . | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 207 | moreover have "\<one>\<^bsub>G Mod H\<^esub> = H #> \<one>" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 208 | unfolding FactGroup_def using H.subset by force | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 209 | hence "the_elem (f ` \<one>\<^bsub>G Mod H\<^esub>) = f \<one>" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 210 | using weak_group_morphism_range[OF assms one_closed] by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 211 | ultimately show ?thesis by (simp add: image_group_def) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 212 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 213 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 214 | corollary (in group) weak_group_morphism_is_hom: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 215 | assumes "weak_group_morphism f H G" shows "f \<in> hom G (image_group f G)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 216 | proof - | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 217 | interpret H: normal H G | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 218 | using weak_group_morphism.axioms(1)[OF assms] . | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 219 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 220 | have the_elem_hom: "(\<lambda>x. the_elem (f ` x)) \<in> hom (G Mod H) (image_group f G)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 221 | using weak_group_morphism_is_iso[OF assms] by (simp add: iso_def) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 222 | have hom: "(\<lambda>x. the_elem (f ` x)) \<circ> (#>) H \<in> hom G (image_group f G)" | 
| 69700 
7a92cbec7030
new material about summations and powers, along with some tweaks
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 223 | using hom_compose[OF H.r_coset_hom_Mod the_elem_hom] | 
| 
7a92cbec7030
new material about summations and powers, along with some tweaks
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 224 | using Group.hom_compose H.r_coset_hom_Mod the_elem_hom by blast | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 225 | have restrict: "\<And>a. a \<in> carrier G \<Longrightarrow> ((\<lambda>x. the_elem (f ` x)) \<circ> (#>) H) a = f a" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 226 | using weak_group_morphism_range[OF assms] by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 227 | show ?thesis | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 228 | using hom_restrict[OF hom restrict] by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 229 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 230 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 231 | corollary (in group) weak_group_morphism_group_hom: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 232 | assumes "weak_group_morphism f H G" shows "group_hom G (image_group f G) f" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 233 | using image_group_is_group[OF assms] weak_group_morphism_is_hom[OF assms] group_axioms | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 234 | unfolding group_hom_def group_hom_axioms_def by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 235 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 236 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 237 | subsection \<open>Weak Ring Morphisms\<close> | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 238 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 239 | lemma image_ring_carrier: "carrier (image_ring f R) = f ` (carrier R)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 240 | unfolding image_ring_def image_group_def by (simp add: monoid.defs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 241 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 242 | lemma image_ring_one: "one (image_ring f R) = f \<one>\<^bsub>R\<^esub>" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 243 | unfolding image_ring_def image_group_def by (simp add: monoid.defs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 244 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 245 | lemma image_ring_zero: "zero (image_ring f R) = f \<zero>\<^bsub>R\<^esub>" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 246 | unfolding image_ring_def image_group_def by (simp add: monoid.defs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 247 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 248 | lemma weak_ring_morphismI: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 249 | assumes "ideal I R" and "\<And>a b. \<lbrakk> a \<in> carrier R; b \<in> carrier R \<rbrakk> \<Longrightarrow> f a = f b \<longleftrightarrow> a \<ominus>\<^bsub>R\<^esub> b \<in> I" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 250 | shows "weak_ring_morphism f I R" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 251 | using assms unfolding weak_ring_morphism_def weak_ring_morphism_axioms_def by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 252 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 253 | lemma (in ring) weak_ring_morphism_range: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 254 |   assumes "weak_ring_morphism f I R" and "a \<in> carrier R" shows "f ` (I +> a) = { f a }"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 255 | using add.weak_group_morphism_range[OF weak_add_group_morphism[OF assms(1)] assms(2)] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 256 | unfolding a_r_coset_def . | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 257 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 258 | lemma (in ring) vimage_eq_a_rcoset: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 259 |   assumes "weak_ring_morphism f I R" and "a \<in> carrier R" shows "{ b \<in> carrier R. f b = f a } = I +> a"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 260 | using add.vimage_eq_rcoset[OF weak_add_group_morphism[OF assms(1)] assms(2)] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 261 | unfolding a_r_coset_def by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 262 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 263 | lemma (in ring) weak_ring_morphism_ker: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 264 | assumes "weak_ring_morphism f I R" shows "a_kernel R (image_ring f R) f = I" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 265 | using add.weak_group_morphism_ker[OF weak_add_group_morphism[OF assms]] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 266 | unfolding kernel_def a_kernel_def' image_ring_def image_group_def by (simp add: monoid.defs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 267 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 268 | lemma (in ring) weak_ring_morphism_inv_into: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 269 | assumes "weak_ring_morphism f I R" and "a \<in> carrier R" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 270 | obtains i where "i \<in> I" "inv_into (carrier R) f (f a) = i \<oplus> a" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 271 | using add.weak_group_morphism_inv_into(1)[OF weak_add_group_morphism[OF assms(1)] assms(2)] by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 272 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 273 | proposition (in ring) weak_ring_morphism_is_iso: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 274 | assumes "weak_ring_morphism f I R" shows "(\<lambda>x. the_elem (f ` x)) \<in> ring_iso (R Quot I) (image_ring f R)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 275 | proof (rule ring_iso_memI) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 276 | show "bij_betw (\<lambda>x. the_elem (f ` x)) (carrier (R Quot I)) (carrier (image_ring f R))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 277 | and add_hom: "\<And>x y. \<lbrakk> x \<in> carrier (R Quot I); y \<in> carrier (R Quot I) \<rbrakk> \<Longrightarrow> | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 278 | the_elem (f ` (x \<oplus>\<^bsub>R Quot I\<^esub> y)) = the_elem (f ` x) \<oplus>\<^bsub>image_ring f R\<^esub> the_elem (f ` y)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 279 | using add.weak_group_morphism_is_iso[OF weak_add_group_morphism[OF assms]] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 280 | unfolding iso_def hom_def FactGroup_def FactRing_def A_RCOSETS_def set_add_def | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 281 | by (auto simp add: image_ring_def image_group_def monoid.defs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 282 | next | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 283 | interpret I: ideal I R | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 284 | using weak_ring_morphism.axioms(1)[OF assms] . | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 285 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 286 | show "the_elem (f ` \<one>\<^bsub>R Quot I\<^esub>) = \<one>\<^bsub>image_ring f R\<^esub>" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 287 | and "\<And>x. x \<in> carrier (R Quot I) \<Longrightarrow> the_elem (f ` x) \<in> carrier (image_ring f R)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 288 | using weak_ring_morphism_range[OF assms] one_closed I.Icarr | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 289 | by (auto simp add: image_ring_def image_group_def monoid.defs FactRing_def A_RCOSETS_def') | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 290 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 291 | fix x y assume "x \<in> carrier (R Quot I)" "y \<in> carrier (R Quot I)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 292 | then obtain a b where a: "a \<in> carrier R" "x = I +> a" and b: "b \<in> carrier R" "y = I +> b" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 293 | unfolding FactRing_def A_RCOSETS_def' by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 294 | hence prod: "x \<otimes>\<^bsub>R Quot I\<^esub> y = I +> (a \<otimes> b)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 295 | unfolding FactRing_def by (simp add: I.rcoset_mult_add) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 296 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 297 | show "the_elem (f ` (x \<otimes>\<^bsub>R Quot I\<^esub> y)) = the_elem (f ` x) \<otimes>\<^bsub>image_ring f R\<^esub> the_elem (f ` y)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 298 | unfolding prod | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 299 | proof (simp add: weak_ring_morphism_range[OF assms] a b image_ring_def image_group_def monoid.defs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 300 | obtain i j | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 301 | where i: "i \<in> I" "inv_into (carrier R) f (f a) = i \<oplus> a" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 302 | and j: "j \<in> I" "inv_into (carrier R) f (f b) = j \<oplus> b" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 303 | using weak_ring_morphism_inv_into[OF assms] a(1) b(1) by metis | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 304 | have "i \<in> carrier R" and "j \<in> carrier R" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 305 | using I.Icarr i(1) j(1) by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 306 | hence "(i \<oplus> a) \<otimes> (j \<oplus> b) = (i \<oplus> a) \<otimes> j \<oplus> (i \<otimes> b) \<oplus> (a \<otimes> b)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 307 | using a(1) b(1) by algebra | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 308 | hence "(i \<oplus> a) \<otimes> (j \<oplus> b) \<in> I +> (a \<otimes> b)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 309 | using i(1) j(1) a(1) b(1) unfolding a_r_coset_def' | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 310 | by (simp add: I.I_l_closed I.I_r_closed) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 311 | thus "f (a \<otimes> b) = f (inv_into (carrier R) f (f a) \<otimes> inv_into (carrier R) f (f b))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 312 | unfolding i j using weak_ring_morphism_range[OF assms m_closed[OF a(1) b(1)]] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 313 | by (metis imageI singletonD) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 314 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 315 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 316 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 317 | corollary (in ring) image_ring_zero': | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 318 | assumes "weak_ring_morphism f I R" shows "the_elem (f ` \<zero>\<^bsub>R Quot I\<^esub>) = \<zero>\<^bsub>image_ring f R\<^esub>" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 319 | proof - | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 320 | interpret I: ideal I R | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 321 | using weak_ring_morphism.axioms(1)[OF assms] . | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 322 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 323 | have "\<zero>\<^bsub>R Quot I\<^esub> = I +> \<zero>" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 324 | unfolding FactRing_def a_r_coset_def' by force | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 325 | thus ?thesis | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 326 | using weak_ring_morphism_range[OF assms zero_closed] unfolding image_ring_zero by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 327 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 328 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 329 | corollary (in ring) image_ring_is_ring: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 330 | assumes "weak_ring_morphism f I R" shows "ring (image_ring f R)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 331 | proof - | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 332 | interpret I: ideal I R | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 333 | using weak_ring_morphism.axioms(1)[OF assms] . | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 334 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 335 | have "ring ((image_ring f R) \<lparr> zero := the_elem (f ` \<zero>\<^bsub>R Quot I\<^esub>) \<rparr>)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 336 | using ring.ring_iso_imp_img_ring[OF I.quotient_is_ring weak_ring_morphism_is_iso[OF assms]] by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 337 | thus ?thesis | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 338 | unfolding image_ring_zero'[OF assms] by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 339 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 340 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 341 | corollary (in ring) image_ring_is_field: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 342 | assumes "weak_ring_morphism f I R" and "field (R Quot I)" shows "field (image_ring f R)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 343 | using field.ring_iso_imp_img_field[OF assms(2) weak_ring_morphism_is_iso[OF assms(1)]] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 344 | unfolding image_ring_zero'[OF assms(1)] by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 345 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 346 | corollary (in ring) weak_ring_morphism_is_hom: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 347 | assumes "weak_ring_morphism f I R" shows "f \<in> ring_hom R (image_ring f R)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 348 | proof - | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 349 | interpret I: ideal I R | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 350 | using weak_ring_morphism.axioms(1)[OF assms] . | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 351 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 352 | have the_elem_hom: "(\<lambda>x. the_elem (f ` x)) \<in> ring_hom (R Quot I) (image_ring f R)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 353 | using weak_ring_morphism_is_iso[OF assms] by (simp add: ring_iso_def) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 354 | have ring_hom: "(\<lambda>x. the_elem (f ` x)) \<circ> (+>) I \<in> ring_hom R (image_ring f R)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 355 | using ring_hom_trans[OF I.rcos_ring_hom the_elem_hom] . | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 356 | have restrict: "\<And>a. a \<in> carrier R \<Longrightarrow> ((\<lambda>x. the_elem (f ` x)) \<circ> (+>) I) a = f a" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 357 | using weak_ring_morphism_range[OF assms] by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 358 | show ?thesis | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 359 | using ring_hom_restrict[OF ring_hom restrict] by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 360 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 361 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 362 | corollary (in ring) weak_ring_morphism_ring_hom: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 363 | assumes "weak_ring_morphism f I R" shows "ring_hom_ring R (image_ring f R) f" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 364 | using ring_hom_ringI2[OF ring_axioms image_ring_is_ring[OF assms] weak_ring_morphism_is_hom[OF assms]] . | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 365 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 366 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 367 | subsection \<open>Injective Functions\<close> | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 368 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 369 | text \<open>If the fuction is injective, we don't need to impose any algebraic restriction to the input | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 370 | scheme in order to state an isomorphism.\<close> | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 371 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 372 | lemma inj_imp_image_group_iso: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 373 | assumes "inj_on f (carrier G)" shows "f \<in> iso G (image_group f G)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 374 | using assms by (auto simp add: image_group_def iso_def bij_betw_def hom_def) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 375 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 376 | lemma inj_imp_image_group_inv_iso: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 377 | assumes "inj f" shows "Hilbert_Choice.inv f \<in> iso (image_group f G) G" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 378 | using assms by (auto simp add: image_group_def iso_def bij_betw_def hom_def inj_on_def) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 379 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 380 | lemma inj_imp_image_ring_iso: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 381 | assumes "inj_on f (carrier R)" shows "f \<in> ring_iso R (image_ring f R)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 382 | using assms by (auto simp add: image_ring_def image_group_def ring_iso_def | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 383 | bij_betw_def ring_hom_def monoid.defs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 384 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 385 | lemma inj_imp_image_ring_inv_iso: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 386 | assumes "inj f" shows "Hilbert_Choice.inv f \<in> ring_iso (image_ring f R) R" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 387 | using assms by (auto simp add: image_ring_def image_group_def ring_iso_def | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 388 | bij_betw_def ring_hom_def inj_on_def monoid.defs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 389 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 390 | lemma (in group) inj_imp_image_group_is_group: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 391 | assumes "inj_on f (carrier G)" shows "group (image_group f G)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 392 | using iso_imp_img_group[OF inj_imp_image_group_iso[OF assms]] by (simp add: image_group_def) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 393 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 394 | lemma (in ring) inj_imp_image_ring_is_ring: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 395 | assumes "inj_on f (carrier R)" shows "ring (image_ring f R)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 396 | using ring_iso_imp_img_ring[OF inj_imp_image_ring_iso[OF assms]] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 397 | by (simp add: image_ring_def image_group_def monoid.defs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 398 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 399 | lemma (in domain) inj_imp_image_ring_is_domain: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 400 | assumes "inj_on f (carrier R)" shows "domain (image_ring f R)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 401 | using ring_iso_imp_img_domain[OF inj_imp_image_ring_iso[OF assms]] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 402 | by (simp add: image_ring_def image_group_def monoid.defs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 403 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 404 | lemma (in field) inj_imp_image_ring_is_field: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 405 | assumes "inj_on f (carrier R)" shows "field (image_ring f R)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 406 | using ring_iso_imp_img_field[OF inj_imp_image_ring_iso[OF assms]] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 407 | by (simp add: image_ring_def image_group_def monoid.defs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 408 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 409 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 410 | section \<open>Examples\<close> | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 411 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 412 | text \<open>In a lot of different contexts, the lack of dependent types make some definitions quite | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 413 | complicated. The tools developed in this theory give us a way to change the type of a | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 414 | scheme and preserve all of its algebraic properties. We show, in this section, how to | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 415 | make use of this feature in order to solve the problem mentioned above. \<close> | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 416 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 417 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 418 | subsection \<open>Direct Product\<close> | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 419 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 420 | abbreviation nil_monoid :: "('a list) monoid"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 421 |   where "nil_monoid \<equiv> \<lparr> carrier = { [] }, mult = (\<lambda>a b. []), one = [] \<rparr>"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 422 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 423 | definition DirProd_list :: "(('a, 'b) monoid_scheme) list \<Rightarrow> ('a list) monoid"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 424 | where "DirProd_list Gs = foldr (\<lambda>G H. image_group (\<lambda>(x, xs). x # xs) (G \<times>\<times> H)) Gs nil_monoid" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 425 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 426 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 427 | subsubsection \<open>Basic Properties\<close> | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 428 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 429 | lemma DirProd_list_carrier: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 430 | shows "carrier (DirProd_list (G # Gs)) = (\<lambda>(x, xs). x # xs) ` (carrier G \<times> carrier (DirProd_list Gs))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 431 | unfolding DirProd_list_def image_group_def by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 432 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 433 | lemma DirProd_list_one: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 434 | shows "one (DirProd_list Gs) = foldr (\<lambda>G tl. (one G) # tl) Gs []" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 435 | unfolding DirProd_list_def DirProd_def image_group_def by (induct Gs) (auto) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 436 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 437 | lemma DirProd_list_carrier_mem: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 438 | assumes "gs \<in> carrier (DirProd_list Gs)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 439 | shows "length gs = length Gs" and "\<And>i. i < length Gs \<Longrightarrow> (gs ! i) \<in> carrier (Gs ! i)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 440 | proof - | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 441 | let ?same_length = "\<lambda>xs ys. length xs = length ys" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 442 | let ?in_carrier = "\<lambda>i gs Gs. (gs ! i) \<in> carrier (Gs ! i)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 443 | from assms have "?same_length gs Gs \<and> (\<forall>i < length Gs. ?in_carrier i gs Gs)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 444 | proof (induct Gs arbitrary: gs, simp add: DirProd_list_def) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 445 | case (Cons G Gs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 446 | then obtain g' gs' | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 447 | where g': "g' \<in> carrier G" and gs': "gs' \<in> carrier (DirProd_list Gs)" and gs: "gs = g' # gs'" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 448 | unfolding DirProd_list_carrier by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 449 |     hence "?same_length gs (G # Gs)" and "\<And>i. i \<in> {(Suc 0)..< length (G # Gs)} \<Longrightarrow> ?in_carrier i gs (G # Gs)"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 450 | using Cons(1) by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 451 | moreover have "?in_carrier 0 gs (G # Gs)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 452 | unfolding gs using g' by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 453 | ultimately show ?case | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 454 | by (metis atLeastLessThan_iff eq_imp_le less_Suc0 linorder_neqE_nat nat_less_le) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 455 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 456 | thus "?same_length gs Gs" and "\<And>i. i < length Gs \<Longrightarrow> ?in_carrier i gs Gs" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 457 | by simp+ | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 458 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 459 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 460 | lemma DirProd_list_carrier_memI: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 461 | assumes "length gs = length Gs" and "\<And>i. i < length Gs \<Longrightarrow> (gs ! i) \<in> carrier (Gs ! i)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 462 | shows "gs \<in> carrier (DirProd_list Gs)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 463 | using assms | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 464 | proof (induct Gs arbitrary: gs, simp add: DirProd_list_def) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 465 | case (Cons G Gs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 466 | then obtain g' gs' where gs: "gs = g' # gs'" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 467 | by (metis length_Suc_conv) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 468 | have "g' \<in> carrier G" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 469 | using Cons(3)[of 0] unfolding gs by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 470 | moreover have "gs' \<in> carrier (DirProd_list Gs)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 471 | using Cons unfolding gs by force | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 472 | ultimately show ?case | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 473 | unfolding DirProd_list_carrier gs by blast | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 474 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 475 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 476 | lemma inj_on_DirProd_carrier: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 477 | shows "inj_on (\<lambda>(g, gs). g # gs) (carrier (G \<times>\<times> (DirProd_list Gs)))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 478 | unfolding DirProd_def inj_on_def by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 479 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 480 | lemma DirProd_list_is_group: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 481 | assumes "\<And>i. i < length Gs \<Longrightarrow> group (Gs ! i)" shows "group (DirProd_list Gs)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 482 | using assms | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 483 | proof (induct Gs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 484 | case Nil thus ?case | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 485 | unfolding DirProd_list_def by (unfold_locales, auto simp add: Units_def) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 486 | next | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 487 | case (Cons G Gs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 488 | hence is_group: "group (G \<times>\<times> (DirProd_list Gs))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 489 | using DirProd_group[of G "DirProd_list Gs"] by force | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 490 | show ?case | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 491 | using group.inj_imp_image_group_is_group[OF is_group inj_on_DirProd_carrier] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 492 | unfolding DirProd_list_def by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 493 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 494 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 495 | lemma DirProd_list_iso: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 496 | "(\<lambda>(g, gs). g # gs) \<in> iso (G \<times>\<times> (DirProd_list Gs)) (DirProd_list (G # Gs))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 497 | using inj_imp_image_group_iso[OF inj_on_DirProd_carrier] unfolding DirProd_list_def by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 498 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 499 | end |