| author | wenzelm | 
| Wed, 15 Jul 2020 11:56:43 +0200 | |
| changeset 72034 | 452073b64f28 | 
| parent 70041 | 2b23dd163c7f | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 68582 | 1 | (* Title: HOL/Algebra/Exact_Sequence.thy | 
| 70041 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 2 | Author: Martin Baillon (first part) and LC Paulson (material ported from HOL Light) | 
| 68582 | 3 | *) | 
| 68578 | 4 | |
| 70041 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 5 | section \<open>Exact Sequences\<close> | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 6 | |
| 68578 | 7 | theory Exact_Sequence | 
| 70041 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 8 | imports Elementary_Groups Solvable_Groups | 
| 68578 | 9 | begin | 
| 10 | ||
| 11 | ||
| 12 | ||
| 13 | subsection \<open>Definitions\<close> | |
| 14 | ||
| 15 | inductive exact_seq :: "'a monoid list \<times> ('a \<Rightarrow> 'a) list \<Rightarrow> bool"  where
 | |
| 16 | unity: " group_hom G1 G2 f \<Longrightarrow> exact_seq ([G2, G1], [f])" | | |
| 17 | extension: "\<lbrakk> exact_seq ((G # K # l), (g # q)); group H ; h \<in> hom G H ; | |
| 18 | kernel G H h = image g (carrier K) \<rbrakk> \<Longrightarrow> exact_seq (H # G # K # l, h # g # q)" | |
| 19 | ||
| 70041 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 20 | inductive_simps exact_seq_end_iff [simp]: "exact_seq ([G,H], (g # q))" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 21 | inductive_simps exact_seq_cons_iff [simp]: "exact_seq ((G # K # H # l), (g # h # q))" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 22 | |
| 68578 | 23 | abbreviation exact_seq_arrow :: | 
| 24 |   "('a \<Rightarrow> 'a) \<Rightarrow> 'a monoid list \<times> ('a \<Rightarrow> 'a) list \<Rightarrow>  'a monoid \<Rightarrow> 'a monoid list \<times> ('a \<Rightarrow> 'a) list"
 | |
| 25 |   ("(3_ / \<longlongrightarrow>\<index> _)" [1000, 60])
 | |
| 26 | where "exact_seq_arrow f t G \<equiv> (G # (fst t), f # (snd t))" | |
| 27 | ||
| 28 | ||
| 29 | subsection \<open>Basic Properties\<close> | |
| 30 | ||
| 31 | lemma exact_seq_length1: "exact_seq t \<Longrightarrow> length (fst t) = Suc (length (snd t))" | |
| 32 | by (induct t rule: exact_seq.induct) auto | |
| 33 | ||
| 34 | lemma exact_seq_length2: "exact_seq t \<Longrightarrow> length (snd t) \<ge> Suc 0" | |
| 35 | by (induct t rule: exact_seq.induct) auto | |
| 36 | ||
| 37 | lemma dropped_seq_is_exact_seq: | |
| 38 | assumes "exact_seq (G, F)" and "(i :: nat) < length F" | |
| 39 | shows "exact_seq (drop i G, drop i F)" | |
| 40 | proof- | |
| 41 |   { fix t i assume "exact_seq t" "i < length (snd t)"
 | |
| 42 | hence "exact_seq (drop i (fst t), drop i (snd t))" | |
| 43 | proof (induction arbitrary: i) | |
| 44 | case (unity G1 G2 f) thus ?case | |
| 45 | by (simp add: exact_seq.unity) | |
| 46 | next | |
| 47 | case (extension G K l g q H h) show ?case | |
| 48 | proof (cases) | |
| 49 | assume "i = 0" thus ?case | |
| 50 | using exact_seq.extension[OF extension.hyps] by simp | |
| 51 | next | |
| 52 | assume "i \<noteq> 0" hence "i \<ge> Suc 0" by simp | |
| 53 | then obtain k where "k < length (snd (G # K # l, g # q))" "i = Suc k" | |
| 54 | using Suc_le_D extension.prems by auto | |
| 55 | thus ?thesis using extension.IH by simp | |
| 56 | qed | |
| 57 | qed } | |
| 58 | ||
| 59 | thus ?thesis using assms by auto | |
| 60 | qed | |
| 61 | ||
| 62 | lemma truncated_seq_is_exact_seq: | |
| 63 | assumes "exact_seq (l, q)" and "length l \<ge> 3" | |
| 64 | shows "exact_seq (tl l, tl q)" | |
| 65 | using exact_seq_length1[OF assms(1)] dropped_seq_is_exact_seq[OF assms(1), of "Suc 0"] | |
| 66 | exact_seq_length2[OF assms(1)] assms(2) by (simp add: drop_Suc) | |
| 67 | ||
| 68 | lemma exact_seq_imp_exact_hom: | |
| 69 | assumes "exact_seq (G1 # l,q) \<longlongrightarrow>\<^bsub>g1\<^esub> G2 \<longlongrightarrow>\<^bsub>g2\<^esub> G3" | |
| 70 | shows "g1 ` (carrier G1) = kernel G2 G3 g2" | |
| 71 | proof- | |
| 72 |   { fix t assume "exact_seq t" and "length (fst t) \<ge> 3 \<and> length (snd t) \<ge> 2"
 | |
| 73 | hence "(hd (tl (snd t))) ` (carrier (hd (tl (tl (fst t))))) = | |
| 74 | kernel (hd (tl (fst t))) (hd (fst t)) (hd (snd t))" | |
| 75 | proof (induction) | |
| 76 | case (unity G1 G2 f) | |
| 77 | then show ?case by auto | |
| 78 | next | |
| 79 | case (extension G l g q H h) | |
| 80 | then show ?case by auto | |
| 81 | qed } | |
| 82 | thus ?thesis using assms by fastforce | |
| 83 | qed | |
| 84 | ||
| 85 | lemma exact_seq_imp_exact_hom_arbitrary: | |
| 86 | assumes "exact_seq (G, F)" | |
| 87 | and "Suc i < length F" | |
| 88 | shows "(F ! (Suc i)) ` (carrier (G ! (Suc (Suc i)))) = kernel (G ! (Suc i)) (G ! i) (F ! i)" | |
| 89 | proof - | |
| 90 | have "length (drop i F) \<ge> 2" "length (drop i G) \<ge> 3" | |
| 91 | using assms(2) exact_seq_length1[OF assms(1)] by auto | |
| 92 | then obtain l q | |
| 93 | where "drop i G = (G ! i) # (G ! (Suc i)) # (G ! (Suc (Suc i))) # l" | |
| 94 | and "drop i F = (F ! i) # (F ! (Suc i)) # q" | |
| 95 | by (metis Cons_nth_drop_Suc Suc_less_eq assms exact_seq_length1 fst_conv | |
| 96 | le_eq_less_or_eq le_imp_less_Suc prod.sel(2)) | |
| 97 | thus ?thesis | |
| 98 | using dropped_seq_is_exact_seq[OF assms(1), of i] assms(2) | |
| 99 | exact_seq_imp_exact_hom[of "G ! i" "G ! (Suc i)" "G ! (Suc (Suc i))" l q] by auto | |
| 100 | qed | |
| 101 | ||
| 102 | lemma exact_seq_imp_group_hom : | |
| 103 | assumes "exact_seq ((G # l, q)) \<longlongrightarrow>\<^bsub>g\<^esub> H" | |
| 104 | shows "group_hom G H g" | |
| 105 | proof- | |
| 106 |   { fix t assume "exact_seq t"
 | |
| 107 | hence "group_hom (hd (tl (fst t))) (hd (fst t)) (hd(snd t))" | |
| 108 | proof (induction) | |
| 109 | case (unity G1 G2 f) | |
| 110 | then show ?case by auto | |
| 111 | next | |
| 112 | case (extension G l g q H h) | |
| 113 | then show ?case unfolding group_hom_def group_hom_axioms_def by auto | |
| 114 | qed } | |
| 115 | note aux_lemma = this | |
| 116 | show ?thesis using aux_lemma[OF assms] | |
| 117 | by simp | |
| 118 | qed | |
| 119 | ||
| 120 | lemma exact_seq_imp_group_hom_arbitrary: | |
| 121 | assumes "exact_seq (G, F)" and "(i :: nat) < length F" | |
| 122 | shows "group_hom (G ! (Suc i)) (G ! i) (F ! i)" | |
| 123 | proof - | |
| 124 | have "length (drop i F) \<ge> 1" "length (drop i G) \<ge> 2" | |
| 125 | using assms(2) exact_seq_length1[OF assms(1)] by auto | |
| 126 | then obtain l q | |
| 127 | where "drop i G = (G ! i) # (G ! (Suc i)) # l" | |
| 128 | and "drop i F = (F ! i) # q" | |
| 129 | by (metis Cons_nth_drop_Suc Suc_leI assms exact_seq_length1 fst_conv | |
| 130 | le_eq_less_or_eq le_imp_less_Suc prod.sel(2)) | |
| 131 | thus ?thesis | |
| 132 | using dropped_seq_is_exact_seq[OF assms(1), of i] assms(2) | |
| 133 | exact_seq_imp_group_hom[of "G ! i" "G ! (Suc i)" l q "F ! i"] by simp | |
| 134 | qed | |
| 135 | ||
| 136 | ||
| 137 | subsection \<open>Link Between Exact Sequences and Solvable Conditions\<close> | |
| 138 | ||
| 139 | lemma exact_seq_solvable_imp : | |
| 140 | assumes "exact_seq ([G1],[]) \<longlongrightarrow>\<^bsub>g1\<^esub> G2 \<longlongrightarrow>\<^bsub>g2\<^esub> G3" | |
| 141 | and "inj_on g1 (carrier G1)" | |
| 142 | and "g2 ` (carrier G2) = carrier G3" | |
| 143 | shows "solvable G2 \<Longrightarrow> (solvable G1) \<and> (solvable G3)" | |
| 144 | proof - | |
| 145 | assume G2: "solvable G2" | |
| 146 | have "group_hom G1 G2 g1" | |
| 147 | using exact_seq_imp_group_hom_arbitrary[OF assms(1), of "Suc 0"] by simp | |
| 148 | hence "solvable G1" | |
| 149 | using group_hom.inj_hom_imp_solvable[of G1 G2 g1] assms(2) G2 by simp | |
| 150 | moreover have "group_hom G2 G3 g2" | |
| 151 | using exact_seq_imp_group_hom_arbitrary[OF assms(1), of 0] by simp | |
| 152 | hence "solvable G3" | |
| 153 | using group_hom.surj_hom_imp_solvable[of G2 G3 g2] assms(3) G2 by simp | |
| 154 | ultimately show ?thesis by simp | |
| 155 | qed | |
| 156 | ||
| 157 | lemma exact_seq_solvable_recip : | |
| 158 | assumes "exact_seq ([G1],[]) \<longlongrightarrow>\<^bsub>g1\<^esub> G2 \<longlongrightarrow>\<^bsub>g2\<^esub> G3" | |
| 159 | and "inj_on g1 (carrier G1)" | |
| 160 | and "g2 ` (carrier G2) = carrier G3" | |
| 161 | shows "(solvable G1) \<and> (solvable G3) \<Longrightarrow> solvable G2" | |
| 162 | proof - | |
| 163 | assume "(solvable G1) \<and> (solvable G3)" | |
| 164 | hence G1: "solvable G1" and G3: "solvable G3" by auto | |
| 165 | have g1: "group_hom G1 G2 g1" and g2: "group_hom G2 G3 g2" | |
| 166 | using exact_seq_imp_group_hom_arbitrary[OF assms(1), of "Suc 0"] | |
| 167 | exact_seq_imp_group_hom_arbitrary[OF assms(1), of 0] by auto | |
| 168 | show ?thesis | |
| 169 | using solvable_condition[OF g1 g2 assms(3)] | |
| 170 | exact_seq_imp_exact_hom[OF assms(1)] G1 G3 by auto | |
| 171 | qed | |
| 172 | ||
| 173 | proposition exact_seq_solvable_iff : | |
| 174 | assumes "exact_seq ([G1],[]) \<longlongrightarrow>\<^bsub>g1\<^esub> G2 \<longlongrightarrow>\<^bsub>g2\<^esub> G3" | |
| 175 | and "inj_on g1 (carrier G1)" | |
| 176 | and "g2 ` (carrier G2) = carrier G3" | |
| 177 | shows "(solvable G1) \<and> (solvable G3) \<longleftrightarrow> solvable G2" | |
| 178 | using exact_seq_solvable_recip exact_seq_solvable_imp assms by blast | |
| 179 | ||
| 70041 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 180 | |
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 181 | lemma exact_seq_eq_triviality: | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 182 | assumes "exact_seq ([E,D,C,B,A], [k,h,g,f])" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 183 | shows "trivial_group C \<longleftrightarrow> f ` carrier A = carrier B \<and> inj_on k (carrier D)" (is "_ = ?rhs") | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 184 | proof | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 185 | assume C: "trivial_group C" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 186 | with assms have "inj_on k (carrier D)" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 187 | apply (auto simp: group_hom.image_from_trivial_group trivial_group_def hom_one) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 188 | apply (simp add: group_hom_def group_hom_axioms_def group_hom.inj_iff_trivial_ker) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 189 | done | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 190 | with assms C show ?rhs | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 191 | apply (auto simp: group_hom.image_from_trivial_group trivial_group_def hom_one) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 192 | apply (auto simp: group_hom_def group_hom_axioms_def hom_def kernel_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 193 | done | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 194 | next | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 195 | assume ?rhs | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 196 | with assms show "trivial_group C" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 197 | apply (simp add: trivial_group_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 198 | by (metis group_hom.inj_iff_trivial_ker group_hom.trivial_hom_iff group_hom_axioms.intro group_hom_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 199 | qed | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 200 | |
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 201 | lemma exact_seq_imp_triviality: | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 202 | "\<lbrakk>exact_seq ([E,D,C,B,A], [k,h,g,f]); f \<in> iso A B; k \<in> iso D E\<rbrakk> \<Longrightarrow> trivial_group C" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 203 | by (metis (no_types, lifting) Group.iso_def bij_betw_def exact_seq_eq_triviality mem_Collect_eq) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 204 | |
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 205 | lemma exact_seq_epi_eq_triviality: | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 206 | "exact_seq ([D,C,B,A], [h,g,f]) \<Longrightarrow> (f ` carrier A = carrier B) \<longleftrightarrow> trivial_homomorphism B C g" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 207 | by (auto simp: trivial_homomorphism_def kernel_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 208 | |
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 209 | lemma exact_seq_mon_eq_triviality: | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 210 | "exact_seq ([D,C,B,A], [h,g,f]) \<Longrightarrow> inj_on h (carrier C) \<longleftrightarrow> trivial_homomorphism B C g" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 211 | by (auto simp: trivial_homomorphism_def kernel_def group.is_monoid inj_on_one_iff' image_def) blast | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 212 | |
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 213 | lemma exact_sequence_sum_lemma: | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 214 | assumes "comm_group G" and h: "h \<in> iso A C" and k: "k \<in> iso B D" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 215 | and ex: "exact_seq ([D,G,A], [g,i])" "exact_seq ([C,G,B], [f,j])" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 216 | and fih: "\<And>x. x \<in> carrier A \<Longrightarrow> f(i x) = h x" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 217 | and gjk: "\<And>x. x \<in> carrier B \<Longrightarrow> g(j x) = k x" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 218 | shows "(\<lambda>(x, y). i x \<otimes>\<^bsub>G\<^esub> j y) \<in> Group.iso (A \<times>\<times> B) G \<and> (\<lambda>z. (f z, g z)) \<in> Group.iso G (C \<times>\<times> D)" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 219 | (is "?ij \<in> _ \<and> ?gf \<in> _") | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 220 | proof (rule epi_iso_compose_rev) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 221 | interpret comm_group G | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 222 | by (rule assms) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 223 | interpret f: group_hom G C f | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 224 | using ex by (simp add: group_hom_def group_hom_axioms_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 225 | interpret g: group_hom G D g | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 226 | using ex by (simp add: group_hom_def group_hom_axioms_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 227 | interpret i: group_hom A G i | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 228 | using ex by (simp add: group_hom_def group_hom_axioms_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 229 | interpret j: group_hom B G j | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 230 | using ex by (simp add: group_hom_def group_hom_axioms_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 231 | have kerf: "kernel G C f = j ` carrier B" and "group A" "group B" "i \<in> hom A G" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 232 | using ex by (auto simp: group_hom_def group_hom_axioms_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 233 | then obtain h' where "h' \<in> hom C A" "(\<forall>x \<in> carrier A. h'(h x) = x)" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 234 | and hh': "(\<forall>y \<in> carrier C. h(h' y) = y)" and "group_isomorphisms A C h h'" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 235 | using h by (auto simp: group.iso_iff_group_isomorphisms group_isomorphisms_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 236 | have homij: "?ij \<in> hom (A \<times>\<times> B) G" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 237 | unfolding case_prod_unfold | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 238 | apply (rule hom_group_mult) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 239 | using ex by (simp_all add: group_hom_def hom_of_fst [unfolded o_def] hom_of_snd [unfolded o_def]) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 240 | show homgf: "?gf \<in> hom G (C \<times>\<times> D)" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 241 | using ex by (simp add: hom_paired) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 242 | show "?ij \<in> epi (A \<times>\<times> B) G" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 243 | proof (clarsimp simp add: epi_iff_subset homij) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 244 | fix x | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 245 | assume x: "x \<in> carrier G" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 246 | with \<open>i \<in> hom A G\<close> \<open>h' \<in> hom C A\<close> have "x \<otimes>\<^bsub>G\<^esub> inv\<^bsub>G\<^esub>(i(h'(f x))) \<in> kernel G C f" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 247 | by (simp add: kernel_def hom_in_carrier hh' fih) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 248 | with kerf obtain y where y: "y \<in> carrier B" "j y = x \<otimes>\<^bsub>G\<^esub> inv\<^bsub>G\<^esub>(i(h'(f x)))" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 249 | by auto | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 250 | have "i (h' (f x)) \<otimes>\<^bsub>G\<^esub> (x \<otimes>\<^bsub>G\<^esub> inv\<^bsub>G\<^esub> i (h' (f x))) = x \<otimes>\<^bsub>G\<^esub> (i (h' (f x)) \<otimes>\<^bsub>G\<^esub> inv\<^bsub>G\<^esub> i (h' (f x)))" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 251 | by (meson \<open>h' \<in> hom C A\<close> x f.hom_closed hom_in_carrier i.hom_closed inv_closed m_lcomm) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 252 | also have "\<dots> = x" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 253 | using \<open>h' \<in> hom C A\<close> hom_in_carrier x by fastforce | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 254 | finally show "x \<in> (\<lambda>(x, y). i x \<otimes>\<^bsub>G\<^esub> j y) ` (carrier A \<times> carrier B)" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 255 | using x y apply (clarsimp simp: image_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 256 | apply (rule_tac x="h'(f x)" in bexI) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 257 | apply (rule_tac x=y in bexI, auto) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 258 | by (meson \<open>h' \<in> hom C A\<close> f.hom_closed hom_in_carrier) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 259 | qed | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 260 | show "(\<lambda>z. (f z, g z)) \<circ> (\<lambda>(x, y). i x \<otimes>\<^bsub>G\<^esub> j y) \<in> Group.iso (A \<times>\<times> B) (C \<times>\<times> D)" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 261 | apply (rule group.iso_eq [where f = "\<lambda>(x,y). (h x,k y)"]) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 262 | using ex | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 263 | apply (auto simp: group_hom_def group_hom_axioms_def DirProd_group iso_paired2 h k fih gjk kernel_def set_eq_iff) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 264 | apply (metis f.hom_closed f.r_one fih imageI) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 265 | apply (metis g.hom_closed g.l_one gjk imageI) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 266 | done | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 267 | qed | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 268 | |
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 269 | subsection \<open>Splitting lemmas and Short exact sequences\<close> | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 270 | text\<open>Ported from HOL Light by LCP\<close> | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 271 | |
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 272 | definition short_exact_sequence | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 273 | where "short_exact_sequence A B C f g \<equiv> \<exists>T1 T2 e1 e2. exact_seq ([T1,A,B,C,T2], [e1,f,g,e2]) \<and> trivial_group T1 \<and> trivial_group T2" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 274 | |
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 275 | lemma short_exact_sequenceD: | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 276 | assumes "short_exact_sequence A B C f g" shows "exact_seq ([A,B,C], [f,g]) \<and> f \<in> epi B A \<and> g \<in> mon C B" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 277 | using assms | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 278 | apply (auto simp: short_exact_sequence_def group_hom_def group_hom_axioms_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 279 | apply (simp add: epi_iff_subset group_hom.intro group_hom.kernel_to_trivial_group group_hom_axioms.intro) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 280 | by (metis (no_types, lifting) group_hom.inj_iff_trivial_ker group_hom.intro group_hom_axioms.intro | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 281 | hom_one image_empty image_insert mem_Collect_eq mon_def trivial_group_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 282 | |
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 283 | lemma short_exact_sequence_iff: | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 284 | "short_exact_sequence A B C f g \<longleftrightarrow> exact_seq ([A,B,C], [f,g]) \<and> f \<in> epi B A \<and> g \<in> mon C B" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 285 | proof - | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 286 | have "short_exact_sequence A B C f g" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 287 | if "exact_seq ([A, B, C], [f, g])" and "f \<in> epi B A" and "g \<in> mon C B" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 288 | proof - | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 289 | show ?thesis | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 290 | unfolding short_exact_sequence_def | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 291 | proof (intro exI conjI) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 292 | have "kernel A (singleton_group \<one>\<^bsub>A\<^esub>) (\<lambda>x. \<one>\<^bsub>A\<^esub>) = f ` carrier B" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 293 | using that by (simp add: kernel_def singleton_group_def epi_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 294 |       moreover have "kernel C B g = {\<one>\<^bsub>C\<^esub>}"
 | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 295 | using that group_hom.inj_iff_trivial_ker mon_def by fastforce | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 296 | ultimately show "exact_seq ([singleton_group (one A), A, B, C, singleton_group (one C)], [\<lambda>x. \<one>\<^bsub>A\<^esub>, f, g, id])" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 297 | using that | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 298 | by (simp add: group_hom_def group_hom_axioms_def group.id_hom_singleton) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 299 | qed auto | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 300 | qed | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 301 | then show ?thesis | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 302 | using short_exact_sequenceD by blast | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 303 | qed | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 304 | |
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 305 | lemma very_short_exact_sequence: | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 306 | assumes "exact_seq ([D,C,B,A], [h,g,f])" "trivial_group A" "trivial_group D" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 307 | shows "g \<in> iso B C" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 308 | using assms | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 309 | apply simp | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 310 | by (metis (no_types, lifting) group_hom.image_from_trivial_group group_hom.iso_iff | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 311 | group_hom.kernel_to_trivial_group group_hom.trivial_ker_imp_inj group_hom_axioms.intro group_hom_def hom_carrier inj_on_one_iff') | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 312 | |
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 313 | lemma splitting_sublemma_gen: | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 314 | assumes ex: "exact_seq ([C,B,A], [g,f])" and fim: "f ` carrier A = H" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 315 |       and "subgroup K B" and 1: "H \<inter> K \<subseteq> {one B}" and eq: "set_mult B H K = carrier B"
 | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 316 | shows "g \<in> iso (subgroup_generated B K) (subgroup_generated C(g ` carrier B))" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 317 | proof - | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 318 | interpret KB: subgroup K B | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 319 | by (rule assms) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 320 | interpret fAB: group_hom A B f | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 321 | using ex by simp | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 322 | interpret gBC: group_hom B C g | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 323 | using ex by (simp add: group_hom_def group_hom_axioms_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 324 | have "group A" "group B" "group C" and kerg: "kernel B C g = f ` carrier A" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 325 | using ex by (auto simp: group_hom_def group_hom_axioms_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 326 | have ker_eq: "kernel B C g = H" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 327 | using ex by (simp add: fim) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 328 | then have "subgroup H B" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 329 | using ex by (simp add: group_hom.img_is_subgroup) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 330 | show ?thesis | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 331 | unfolding iso_iff | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 332 | proof (intro conjI) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 333 | show "g \<in> hom (subgroup_generated B K) (subgroup_generated C(g ` carrier B))" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 334 | by (metis ker_eq \<open>subgroup K B\<close> eq gBC.hom_between_subgroups gBC.set_mult_ker_hom(2) order_refl subgroup.subset) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 335 | show "g ` carrier (subgroup_generated B K) = carrier (subgroup_generated C(g ` carrier B))" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 336 | by (metis assms(3) eq fAB.H.subgroupE(1) gBC.img_is_subgroup gBC.set_mult_ker_hom(2) ker_eq subgroup.carrier_subgroup_generated_subgroup) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 337 | interpret gKBC: group_hom "subgroup_generated B K" C g | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 338 | apply (auto simp: group_hom_def group_hom_axioms_def \<open>group C\<close>) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 339 | by (simp add: fAB.H.hom_from_subgroup_generated gBC.homh) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 340 | have *: "x = \<one>\<^bsub>B\<^esub>" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 341 | if x: "x \<in> carrier (subgroup_generated B K)" and "g x = \<one>\<^bsub>C\<^esub>" for x | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 342 | proof - | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 343 | have x': "x \<in> carrier B" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 344 | using that fAB.H.carrier_subgroup_generated_subset by blast | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 345 | moreover have "x \<in> H" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 346 | using kerg fim x' that by (auto simp: kernel_def set_eq_iff) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 347 | ultimately show ?thesis | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 348 | by (metis "1" x Int_iff singletonD KB.carrier_subgroup_generated_subgroup subsetCE) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 349 | qed | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 350 | show "inj_on g (carrier (subgroup_generated B K))" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 351 | using "*" gKBC.inj_on_one_iff by auto | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 352 | qed | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 353 | qed | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 354 | |
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 355 | lemma splitting_sublemma: | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 356 | assumes ex: "short_exact_sequence C B A g f" and fim: "f ` carrier A = H" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 357 |       and "subgroup K B" and 1: "H \<inter> K \<subseteq> {one B}" and eq: "set_mult B H K = carrier B"
 | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 358 | shows "f \<in> iso A (subgroup_generated B H)" (is ?f) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 359 | "g \<in> iso (subgroup_generated B K) C" (is ?g) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 360 | proof - | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 361 | show ?f | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 362 | using short_exact_sequenceD [OF ex] | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 363 | apply (clarsimp simp add: group_hom_def group.iso_onto_image) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 364 | using fim group.iso_onto_image by blast | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 365 | have "C = subgroup_generated C(g ` carrier B)" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 366 | using short_exact_sequenceD [OF ex] | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 367 | apply simp | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 368 | by (metis epi_iff_subset group.subgroup_generated_group_carrier hom_carrier subset_antisym) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 369 | then show ?g | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 370 | using short_exact_sequenceD [OF ex] | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 371 | by (metis "1" \<open>subgroup K B\<close> eq fim splitting_sublemma_gen) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 372 | qed | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 373 | |
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 374 | lemma splitting_lemma_left_gen: | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 375 | assumes ex: "exact_seq ([C,B,A], [g,f])" and f': "f' \<in> hom B A" and iso: "(f' \<circ> f) \<in> iso A A" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 376 | and injf: "inj_on f (carrier A)" and surj: "g ` carrier B = carrier C" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 377 |  obtains H K where "H \<lhd> B" "K \<lhd> B" "H \<inter> K \<subseteq> {one B}" "set_mult B H K = carrier B"
 | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 378 | "f \<in> iso A (subgroup_generated B H)" "g \<in> iso (subgroup_generated B K) C" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 379 | proof - | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 380 | interpret gBC: group_hom B C g | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 381 | using ex by (simp add: group_hom_def group_hom_axioms_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 382 | have "group A" "group B" "group C" and kerg: "kernel B C g = f ` carrier A" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 383 | using ex by (auto simp: group_hom_def group_hom_axioms_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 384 |   then have *: "f ` carrier A \<inter> kernel B A f' = {\<one>\<^bsub>B\<^esub>} \<and> f ` carrier A <#>\<^bsub>B\<^esub> kernel B A f' = carrier B"
 | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 385 | using group_semidirect_sum_image_ker [of f A B f' A] assms by auto | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 386 | interpret f'AB: group_hom B A f' | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 387 | using assms by (auto simp: group_hom_def group_hom_axioms_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 388 | let ?H = "f ` carrier A" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 389 | let ?K = "kernel B A f'" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 390 | show thesis | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 391 | proof | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 392 | show "?H \<lhd> B" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 393 | by (simp add: gBC.normal_kernel flip: kerg) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 394 | show "?K \<lhd> B" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 395 | by (rule f'AB.normal_kernel) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 396 |     show "?H \<inter> ?K \<subseteq> {\<one>\<^bsub>B\<^esub>}" "?H <#>\<^bsub>B\<^esub> ?K = carrier B"
 | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 397 | using * by auto | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 398 | show "f \<in> Group.iso A (subgroup_generated B ?H)" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 399 | using ex by (simp add: injf iso_onto_image group_hom_def group_hom_axioms_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 400 | have C: "C = subgroup_generated C(g ` carrier B)" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 401 | using surj by (simp add: gBC.subgroup_generated_group_carrier) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 402 | show "g \<in> Group.iso (subgroup_generated B ?K) C" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 403 | apply (subst C) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 404 | apply (rule splitting_sublemma_gen [OF ex refl]) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 405 | using * by (auto simp: f'AB.subgroup_kernel) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 406 | qed | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 407 | qed | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 408 | |
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 409 | lemma splitting_lemma_left: | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 410 | assumes ex: "exact_seq ([C,B,A], [g,f])" and f': "f' \<in> hom B A" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 411 | and inv: "(\<And>x. x \<in> carrier A \<Longrightarrow> f'(f x) = x)" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 412 | and injf: "inj_on f (carrier A)" and surj: "g ` carrier B = carrier C" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 413 |  obtains H K where "H \<lhd> B" "K \<lhd> B" "H \<inter> K \<subseteq> {one B}" "set_mult B H K = carrier B"
 | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 414 | "f \<in> iso A (subgroup_generated B H)" "g \<in> iso (subgroup_generated B K) C" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 415 | proof - | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 416 | interpret fAB: group_hom A B f | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 417 | using ex by simp | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 418 | interpret gBC: group_hom B C g | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 419 | using ex by (simp add: group_hom_def group_hom_axioms_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 420 | have "group A" "group B" "group C" and kerg: "kernel B C g = f ` carrier A" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 421 | using ex by (auto simp: group_hom_def group_hom_axioms_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 422 | have iso: "f' \<circ> f \<in> Group.iso A A" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 423 | using ex by (auto simp: inv intro: group.iso_eq [OF \<open>group A\<close> id_iso]) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 424 | show thesis | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 425 | by (metis that splitting_lemma_left_gen [OF ex f' iso injf surj]) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 426 | qed | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 427 | |
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 428 | lemma splitting_lemma_right_gen: | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 429 | assumes ex: "short_exact_sequence C B A g f" and g': "g' \<in> hom C B" and iso: "(g \<circ> g') \<in> iso C C" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 430 |  obtains H K where "H \<lhd> B" "subgroup K B" "H \<inter> K \<subseteq> {one B}" "set_mult B H K = carrier B"
 | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 431 | "f \<in> iso A (subgroup_generated B H)" "g \<in> iso (subgroup_generated B K) C" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 432 | proof | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 433 | interpret fAB: group_hom A B f | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 434 | using short_exact_sequenceD [OF ex] by (simp add: group_hom_def group_hom_axioms_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 435 | interpret gBC: group_hom B C g | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 436 | using short_exact_sequenceD [OF ex] by (simp add: group_hom_def group_hom_axioms_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 437 |   have *: "f ` carrier A \<inter> g' ` carrier C = {\<one>\<^bsub>B\<^esub>}"
 | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 438 | "f ` carrier A <#>\<^bsub>B\<^esub> g' ` carrier C = carrier B" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 439 | "group A" "group B" "group C" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 440 | "kernel B C g = f ` carrier A" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 441 | using group_semidirect_sum_ker_image [of g g' C C B] short_exact_sequenceD [OF ex] | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 442 | by (simp_all add: g' iso group_hom_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 443 | show "kernel B C g \<lhd> B" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 444 | by (simp add: gBC.normal_kernel) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 445 |   show "(kernel B C g) \<inter> (g' ` carrier C) \<subseteq> {\<one>\<^bsub>B\<^esub>}" "(kernel B C g) <#>\<^bsub>B\<^esub> (g' ` carrier C) = carrier B"
 | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 446 | by (auto simp: *) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 447 | show "f \<in> Group.iso A (subgroup_generated B (kernel B C g))" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 448 | by (metis "*"(6) fAB.group_hom_axioms group.iso_onto_image group_hom_def short_exact_sequenceD [OF ex]) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 449 | show "subgroup (g' ` carrier C) B" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 450 | using splitting_sublemma | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 451 | by (simp add: fAB.H.is_group g' gBC.is_group group_hom.img_is_subgroup group_hom_axioms_def group_hom_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 452 | then show "g \<in> Group.iso (subgroup_generated B (g' ` carrier C)) C" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 453 | by (metis (no_types, lifting) iso_iff fAB.H.hom_from_subgroup_generated gBC.homh image_comp inj_on_imageI iso subgroup.carrier_subgroup_generated_subgroup) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 454 | qed | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 455 | |
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 456 | lemma splitting_lemma_right: | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 457 | assumes ex: "short_exact_sequence C B A g f" and g': "g' \<in> hom C B" and gg': "\<And>z. z \<in> carrier C \<Longrightarrow> g(g' z) = z" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 458 |  obtains H K where "H \<lhd> B" "subgroup K B" "H \<inter> K \<subseteq> {one B}" "set_mult B H K = carrier B"
 | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 459 | "f \<in> iso A (subgroup_generated B H)" "g \<in> iso (subgroup_generated B K) C" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 460 | proof - | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 461 | have *: "group A" "group B" "group C" | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 462 | using group_semidirect_sum_ker_image [of g g' C C B] short_exact_sequenceD [OF ex] | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 463 | by (simp_all add: g' group_hom_def) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 464 | show thesis | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 465 | apply (rule splitting_lemma_right_gen [OF ex g' group.iso_eq [OF _ id_iso]]) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 466 | using * apply (auto simp: gg' intro: that) | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 467 | done | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 468 | qed | 
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 469 | |
| 
2b23dd163c7f
Material concerning exact sequences of groups
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 470 | |
| 68578 | 471 | end |