| author | paulson | 
| Tue, 11 Jul 2023 20:22:08 +0100 | |
| changeset 78321 | 021fb1b01de5 | 
| parent 69122 | 1b5178abaf97 | 
| permissions | -rw-r--r-- | 
| 68582 | 1 | (* Title: HOL/Algebra/Solvable_Groups.thy | 
| 2 | Author: Paulo Emílio de Vilhena | |
| 3 | *) | |
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5 | theory Solvable_Groups | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 6 | imports Generated_Groups | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 7 | |
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 8 | begin | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 9 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 10 | section \<open>Solvable Groups\<close> | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 11 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 12 | subsection \<open>Definitions\<close> | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 13 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 14 | inductive solvable_seq :: "('a, 'b) monoid_scheme \<Rightarrow> 'a set \<Rightarrow> bool"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 15 | for G where | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 16 |     unity: "solvable_seq G { \<one>\<^bsub>G\<^esub> }"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 17 | | extension: "\<lbrakk> solvable_seq G K; K \<lhd> (G \<lparr> carrier := H \<rparr>); subgroup H G; | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 18 | comm_group ((G \<lparr> carrier := H \<rparr>) Mod K) \<rbrakk> \<Longrightarrow> solvable_seq G H" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 19 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 20 | definition solvable :: "('a, 'b) monoid_scheme \<Rightarrow> bool"
 | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 21 | where "solvable G \<longleftrightarrow> solvable_seq G (carrier G)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 22 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 23 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 24 | subsection \<open>Solvable Groups and Derived Subgroups\<close> | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 25 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 26 | text \<open>We show that a group G is solvable iff the subgroup (derived G ^^ n) (carrier G) | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 27 | is trivial for a sufficiently large n. \<close> | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 28 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 29 | lemma (in group) solvable_imp_subgroup: | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 30 | assumes "solvable_seq G H" shows "subgroup H G" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 31 | using assms normal.axioms(1)[OF one_is_normal] by (induction) (auto) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 32 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 33 | lemma (in group) augment_solvable_seq: | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 34 | assumes "subgroup H G" and "solvable_seq G (derived G H)" shows "solvable_seq G H" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 35 | using extension[OF _ derived_subgroup_is_normal _ derived_quot_of_subgroup_is_comm_group] assms by simp | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 36 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 37 | theorem (in group) trivial_derived_seq_imp_solvable: | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 38 |   assumes "subgroup H G" and "((derived G) ^^ n) H = { \<one> }" shows "solvable_seq G H"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 39 | using assms | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 40 | proof (induct n arbitrary: H, simp add: unity[of G]) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 41 | case (Suc n) thus ?case | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 42 | using augment_solvable_seq derived_is_subgroup[OF subgroup.subset] by (simp add: funpow_swap1) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 43 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 44 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 45 | theorem (in group) solvable_imp_trivial_derived_seq: | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 46 |   assumes "solvable_seq G H" shows "\<exists>n. (derived G ^^ n) H = { \<one> }"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 47 | using assms | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 48 | proof (induction) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 49 | case unity | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 50 |   have "(derived G ^^ 0) { \<one> } = { \<one> }"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 51 | by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 52 | thus ?case by blast | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 53 | next | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 54 | case (extension K H) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 55 |   obtain n where "(derived G ^^ n) K = { \<one> }"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 56 | using solvable_imp_subgroup extension(1,5) by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 57 |   hence "(derived G ^^ (Suc n)) H \<subseteq> { \<one> }"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 58 | using mono_exp_of_derived[OF derived_of_subgroup_minimal[OF extension(2-4)], of n] by (simp add: funpow_swap1) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 59 |   moreover have "{ \<one> } \<subseteq> (derived G ^^ (Suc n)) H"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 60 | using subgroup.one_closed[OF exp_of_derived_is_subgroup[OF extension(3)], of "Suc n"] by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 61 | ultimately show ?case | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 62 | by blast | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 63 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 64 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 65 | theorem (in group) solvable_iff_trivial_derived_seq: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 66 |   "solvable G \<longleftrightarrow> (\<exists>n. (derived G ^^ n) (carrier G) = { \<one> })"
 | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 67 | using solvable_imp_trivial_derived_seq subgroup_self trivial_derived_seq_imp_solvable | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 68 | by (auto simp add: solvable_def) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 69 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 70 | corollary (in group) solvable_subgroup: | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 71 | assumes "subgroup H G" and "solvable G" shows "solvable_seq G H" | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 72 | proof - | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 73 |   obtain n where n: "(derived G ^^ n) (carrier G) = { \<one> }"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 74 | using assms(2) solvable_imp_trivial_derived_seq by (auto simp add: solvable_def) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 75 | show ?thesis | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 76 | proof (rule trivial_derived_seq_imp_solvable[OF assms(1), of n]) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 77 |     show "(derived G ^^ n) H = { \<one> }"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 78 | using subgroup.one_closed[OF exp_of_derived_is_subgroup[OF assms(1)], of n] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 79 | mono_exp_of_derived[OF subgroup.subset[OF assms(1)], of n] n | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 80 | by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 81 | qed | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 82 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 83 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 84 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 85 | subsection \<open>Short Exact Sequences\<close> | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 86 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 87 | text \<open>Even if we don't talk about short exact sequences explicitly, we show that given an | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 88 | injective homomorphism from a group H to a group G, if H isn't solvable the group G | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 89 | isn't neither. \<close> | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 90 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 91 | theorem (in group_hom) solvable_img_imp_solvable: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 92 | assumes "subgroup K G" and "inj_on h K" and "solvable_seq H (h ` K)" shows "solvable_seq G K" | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 93 | proof - | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 94 |   obtain n where "(derived H ^^ n) (h ` K) = { \<one>\<^bsub>H\<^esub> }"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 95 | using solvable_imp_trivial_derived_seq assms(1,3) by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 96 |   hence "h ` ((derived G ^^ n) K) = { \<one>\<^bsub>H\<^esub> }"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 97 | unfolding exp_of_derived_img[OF subgroup.subset[OF assms(1)]] . | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 98 | moreover have "(derived G ^^ n) K \<subseteq> K" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 99 | using G.mono_derived[of _ K] G.derived_incl[OF _ assms(1)] by (induct n) (auto) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 100 | hence "inj_on h ((derived G ^^ n) K)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 101 | using inj_on_subset[OF assms(2)] by blast | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 102 |   moreover have "{ \<one> } \<subseteq> (derived G ^^ n) K"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 103 | using subgroup.one_closed[OF G.exp_of_derived_is_subgroup[OF assms(1)]] by blast | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 104 | ultimately show ?thesis | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 105 | using G.trivial_derived_seq_imp_solvable[OF assms(1), of n] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 106 | by (metis (no_types, lifting) hom_one image_empty image_insert inj_on_image_eq_iff order_refl) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 107 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 108 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 109 | corollary (in group_hom) inj_hom_imp_solvable: | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 110 | assumes "inj_on h (carrier G)" and "solvable H" shows "solvable G" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 111 | using solvable_img_imp_solvable[OF _ assms(1)] G.subgroup_self | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 112 | solvable_subgroup[OF subgroup_img_is_subgroup assms(2)] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 113 | unfolding solvable_def | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 114 | by simp | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 115 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 116 | theorem (in group_hom) solvable_imp_solvable_img: | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 117 | assumes "solvable_seq G K" shows "solvable_seq H (h ` K)" | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 118 | proof - | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 119 |   obtain n where "(derived G ^^ n) K = { \<one> }"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 120 | using G.solvable_imp_trivial_derived_seq[OF assms] by blast | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 121 | thus ?thesis | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 122 | using trivial_derived_seq_imp_solvable[OF subgroup_img_is_subgroup, of _ n] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 123 | exp_of_derived_img[OF subgroup.subset, of _ n] G.solvable_imp_subgroup[OF assms] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 124 | by auto | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 125 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 126 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 127 | corollary (in group_hom) surj_hom_imp_solvable: | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 128 | assumes "h ` carrier G = carrier H" and "solvable G" shows "solvable H" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 129 | using assms solvable_imp_solvable_img[of "carrier G"] unfolding solvable_def by simp | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 130 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 131 | lemma solvable_seq_condition: | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 132 | assumes "group_hom G H f" "group_hom H K g" and "f ` I \<subseteq> J" and "kernel H K g \<subseteq> f ` I" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 133 | and "subgroup J H" and "solvable_seq G I" "solvable_seq K (g ` J)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 134 | shows "solvable_seq H J" | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 135 | proof - | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 136 | interpret G: group G + H: group H + K: group K + J: subgroup J H + I: subgroup I G | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 137 | using assms(1-2,5) group.solvable_imp_subgroup[OF _ assms(6)] unfolding group_hom_def by auto | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 138 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 139 | obtain n m | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 140 |     where n: "(derived G ^^ n) I = { \<one>\<^bsub>G\<^esub> }" and m: "(derived K ^^ m) (g ` J) = { \<one>\<^bsub>K\<^esub> }"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 141 | using G.solvable_imp_trivial_derived_seq[OF assms(6)] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 142 | K.solvable_imp_trivial_derived_seq[OF assms(7)] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 143 | by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 144 | have "(derived H ^^ m) J \<subseteq> f ` I" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 145 | using m H.exp_of_derived_in_carrier[OF J.subset, of m] assms(4) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 146 | by (auto simp add: group_hom.exp_of_derived_img[OF assms(2) J.subset] kernel_def) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 147 | hence "(derived H ^^ n) ((derived H ^^ m) J) \<subseteq> f ` ((derived G ^^ n) I)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 148 | using n H.mono_exp_of_derived unfolding sym[OF group_hom.exp_of_derived_img[OF assms(1) I.subset, of n]] by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 149 |   hence "(derived H ^^ (n + m)) J \<subseteq> { \<one>\<^bsub>H\<^esub> }"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 150 | using group_hom.hom_one[OF assms(1)] unfolding n by (simp add: funpow_add) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 151 |   moreover have "{ \<one>\<^bsub>H\<^esub> } \<subseteq> (derived H ^^ (n + m)) J"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 152 | using subgroup.one_closed[OF H.exp_of_derived_is_subgroup[OF assms(5), of "n + m"]] by blast | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 153 | ultimately show ?thesis | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 154 | using H.trivial_derived_seq_imp_solvable[OF assms(5)] by simp | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 155 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 156 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 157 | lemma solvable_condition: | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 158 | assumes "group_hom G H f" "group_hom H K g" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 159 | and "g ` (carrier H) = carrier K" and "kernel H K g \<subseteq> f ` (carrier G)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 160 | and "solvable G" "solvable K" shows "solvable H" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 161 | using solvable_seq_condition[OF assms(1-2) _ assms(4) group.subgroup_self] assms(3,5-6) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 162 | subgroup.subset[OF group_hom.img_is_subgroup[OF assms(1)]] group_hom.axioms(2)[OF assms(1)] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 163 | by (simp add: solvable_def) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 164 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68605diff
changeset | 165 | end |