| author | wenzelm | 
| Tue, 27 Mar 2018 13:59:01 +0200 | |
| changeset 67953 | f646d1c826a1 | 
| parent 67613 | ce654b0e6d69 | 
| child 68399 | 0b71d08528f0 | 
| permissions | -rw-r--r-- | 
| 14706 | 1 | (* Title: HOL/Algebra/Sylow.thy | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 2 | Author: Florian Kammueller, with new proofs by L C Paulson | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 3 | *) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 4 | |
| 35849 | 5 | theory Sylow | 
| 64912 | 6 | imports Coset Exponent | 
| 35849 | 7 | begin | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 8 | |
| 64912 | 9 | text \<open>See also @{cite "Kammueller-Paulson:1999"}.\<close>
 | 
| 10 | ||
| 11 | text \<open>The combinatorial argument is in theory @{theory Exponent}.\<close>
 | |
| 14706 | 12 | |
| 64912 | 13 | lemma le_extend_mult: "\<lbrakk>0 < c; a \<le> b\<rbrakk> \<Longrightarrow> a \<le> b * c" | 
| 14 | for c :: nat | |
| 15 | by (metis divisors_zero dvd_triv_left leI less_le_trans nat_dvd_not_less zero_less_iff_neq_zero) | |
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 16 | |
| 14747 | 17 | locale sylow = group + | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 18 | fixes p and a and m and calM and RelM | 
| 64912 | 19 | assumes prime_p: "prime p" | 
| 20 | and order_G: "order G = (p^a) * m" | |
| 21 | and finite_G[iff]: "finite (carrier G)" | |
| 22 |   defines "calM \<equiv> {s. s \<subseteq> carrier G \<and> card s = p^a}"
 | |
| 23 |     and "RelM \<equiv> {(N1, N2). N1 \<in> calM \<and> N2 \<in> calM \<and> (\<exists>g \<in> carrier G. N1 = N2 #> g)}"
 | |
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 24 | begin | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 25 | |
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 26 | lemma RelM_refl_on: "refl_on calM RelM" | 
| 64912 | 27 | by (auto simp: refl_on_def RelM_def calM_def) (blast intro!: coset_mult_one [symmetric]) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 28 | |
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 29 | lemma RelM_sym: "sym RelM" | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 30 | proof (unfold sym_def RelM_def, clarify) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 31 | fix y g | 
| 64912 | 32 | assume "y \<in> calM" | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 33 | and g: "g \<in> carrier G" | 
| 64912 | 34 | then have "y = y #> g #> (inv g)" | 
| 35 | by (simp add: coset_mult_assoc calM_def) | |
| 36 | then show "\<exists>g'\<in>carrier G. y = y #> g #> g'" | |
| 37 | by (blast intro: g) | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 38 | qed | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 39 | |
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 40 | lemma RelM_trans: "trans RelM" | 
| 64912 | 41 | by (auto simp add: trans_def RelM_def calM_def coset_mult_assoc) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 42 | |
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 43 | lemma RelM_equiv: "equiv calM RelM" | 
| 64912 | 44 | unfolding equiv_def by (blast intro: RelM_refl_on RelM_sym RelM_trans) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 45 | |
| 64912 | 46 | lemma M_subset_calM_prep: "M' \<in> calM // RelM \<Longrightarrow> M' \<subseteq> calM" | 
| 47 | unfolding RelM_def by (blast elim!: quotientE) | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 48 | |
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 49 | end | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
16663diff
changeset | 50 | |
| 64912 | 51 | subsection \<open>Main Part of the Proof\<close> | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 52 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 53 | locale sylow_central = sylow + | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 54 | fixes H and M1 and M | 
| 64912 | 55 | assumes M_in_quot: "M \<in> calM // RelM" | 
| 56 | and not_dvd_M: "\<not> (p ^ Suc (multiplicity p m) dvd card M)" | |
| 57 | and M1_in_M: "M1 \<in> M" | |
| 58 |   defines "H \<equiv> {g. g \<in> carrier G \<and> M1 #> g = M1}"
 | |
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 59 | begin | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 60 | |
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 61 | lemma M_subset_calM: "M \<subseteq> calM" | 
| 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 62 | by (rule M_in_quot [THEN M_subset_calM_prep]) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 63 | |
| 64912 | 64 | lemma card_M1: "card M1 = p^a" | 
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 65 | using M1_in_M M_subset_calM calM_def by blast | 
| 64912 | 66 | |
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 67 | lemma exists_x_in_M1: "\<exists>x. x \<in> M1" | 
| 64912 | 68 | using prime_p [THEN prime_gt_Suc_0_nat] card_M1 | 
| 69 | by (metis Suc_lessD card_eq_0_iff empty_subsetI equalityI gr_implies_not0 nat_zero_less_power_iff subsetI) | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 70 | |
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 71 | lemma M1_subset_G [simp]: "M1 \<subseteq> carrier G" | 
| 64912 | 72 | using M1_in_M M_subset_calM calM_def mem_Collect_eq subsetCE by blast | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 73 | |
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 74 | lemma M1_inj_H: "\<exists>f \<in> H\<rightarrow>M1. inj_on f H" | 
| 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 75 | proof - | 
| 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 76 | from exists_x_in_M1 obtain m1 where m1M: "m1 \<in> M1".. | 
| 64912 | 77 | have m1: "m1 \<in> carrier G" | 
| 78 | by (simp add: m1M M1_subset_G [THEN subsetD]) | |
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 79 | show ?thesis | 
| 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 80 | proof | 
| 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 81 | show "inj_on (\<lambda>z\<in>H. m1 \<otimes> z) H" | 
| 64912 | 82 | by (simp add: inj_on_def l_cancel [of m1 x y, THEN iffD1] H_def m1) | 
| 67399 | 83 | show "restrict ((\<otimes>) m1) H \<in> H \<rightarrow> M1" | 
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 84 | proof (rule restrictI) | 
| 64912 | 85 | fix z | 
| 86 | assume zH: "z \<in> H" | |
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 87 | show "m1 \<otimes> z \<in> M1" | 
| 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 88 | proof - | 
| 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 89 | from zH | 
| 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 90 | have zG: "z \<in> carrier G" and M1zeq: "M1 #> z = M1" | 
| 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 91 | by (auto simp add: H_def) | 
| 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 92 | show ?thesis | 
| 64912 | 93 | by (rule subst [OF M1zeq]) (simp add: m1M zG rcosI) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 94 | qed | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 95 | qed | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 96 | qed | 
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 97 | qed | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 98 | |
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 99 | end | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 100 | |
| 64912 | 101 | |
| 102 | subsection \<open>Discharging the Assumptions of \<open>sylow_central\<close>\<close> | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 103 | |
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 104 | context sylow | 
| 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 105 | begin | 
| 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 106 | |
| 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 107 | lemma EmptyNotInEquivSet: "{} \<notin> calM // RelM"
 | 
| 64912 | 108 | by (blast elim!: quotientE dest: RelM_equiv [THEN equiv_class_self]) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 109 | |
| 64912 | 110 | lemma existsM1inM: "M \<in> calM // RelM \<Longrightarrow> \<exists>M1. M1 \<in> M" | 
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 111 | using RelM_equiv equiv_Eps_in by blast | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 112 | |
| 64912 | 113 | lemma zero_less_o_G: "0 < order G" | 
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 114 | by (simp add: order_def card_gt_0_iff carrier_not_empty) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 115 | |
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 116 | lemma zero_less_m: "m > 0" | 
| 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 117 | using zero_less_o_G by (simp add: order_G) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 118 | |
| 64912 | 119 | lemma card_calM: "card calM = (p^a) * m choose p^a" | 
| 120 | by (simp add: calM_def n_subsets order_G [symmetric] order_def) | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 121 | |
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 122 | lemma zero_less_card_calM: "card calM > 0" | 
| 64912 | 123 | by (simp add: card_calM zero_less_binomial le_extend_mult zero_less_m) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 124 | |
| 64912 | 125 | lemma max_p_div_calM: "\<not> (p ^ Suc (multiplicity p m) dvd card calM)" | 
| 63534 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 eberlm <eberlm@in.tum.de> parents: 
63167diff
changeset | 126 | proof | 
| 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 eberlm <eberlm@in.tum.de> parents: 
63167diff
changeset | 127 | assume "p ^ Suc (multiplicity p m) dvd card calM" | 
| 64912 | 128 | with zero_less_card_calM prime_p | 
| 63534 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 eberlm <eberlm@in.tum.de> parents: 
63167diff
changeset | 129 | have "Suc (multiplicity p m) \<le> multiplicity p (card calM)" | 
| 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 eberlm <eberlm@in.tum.de> parents: 
63167diff
changeset | 130 | by (intro multiplicity_geI) auto | 
| 64912 | 131 | then have "multiplicity p m < multiplicity p (card calM)" by simp | 
| 63534 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 eberlm <eberlm@in.tum.de> parents: 
63167diff
changeset | 132 | also have "multiplicity p m = multiplicity p (card calM)" | 
| 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 eberlm <eberlm@in.tum.de> parents: 
63167diff
changeset | 133 | by (simp add: const_p_fac prime_p zero_less_m card_calM) | 
| 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 eberlm <eberlm@in.tum.de> parents: 
63167diff
changeset | 134 | finally show False by simp | 
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 135 | qed | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 136 | |
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 137 | lemma finite_calM: "finite calM" | 
| 64912 | 138 | unfolding calM_def by (rule finite_subset [where B = "Pow (carrier G)"]) auto | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 139 | |
| 64912 | 140 | lemma lemma_A1: "\<exists>M \<in> calM // RelM. \<not> (p ^ Suc (multiplicity p m) dvd card M)" | 
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 141 | using RelM_equiv equiv_imp_dvd_card finite_calM max_p_div_calM by blast | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 142 | |
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 143 | end | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 144 | |
| 64912 | 145 | |
| 146 | subsubsection \<open>Introduction and Destruct Rules for \<open>H\<close>\<close> | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 147 | |
| 64914 | 148 | context sylow_central | 
| 149 | begin | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 150 | |
| 64914 | 151 | lemma H_I: "\<lbrakk>g \<in> carrier G; M1 #> g = M1\<rbrakk> \<Longrightarrow> g \<in> H" | 
| 64912 | 152 | by (simp add: H_def) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 153 | |
| 64914 | 154 | lemma H_into_carrier_G: "x \<in> H \<Longrightarrow> x \<in> carrier G" | 
| 64912 | 155 | by (simp add: H_def) | 
| 156 | ||
| 64914 | 157 | lemma in_H_imp_eq: "g \<in> H \<Longrightarrow> M1 #> g = M1" | 
| 158 | by (simp add: H_def) | |
| 159 | ||
| 160 | lemma H_m_closed: "\<lbrakk>x \<in> H; y \<in> H\<rbrakk> \<Longrightarrow> x \<otimes> y \<in> H" | |
| 64912 | 161 | by (simp add: H_def coset_mult_assoc [symmetric]) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 162 | |
| 64914 | 163 | lemma H_not_empty: "H \<noteq> {}"
 | 
| 64912 | 164 | apply (simp add: H_def) | 
| 165 | apply (rule exI [of _ \<one>]) | |
| 166 | apply simp | |
| 167 | done | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 168 | |
| 64914 | 169 | lemma H_is_subgroup: "subgroup H G" | 
| 64912 | 170 | apply (rule subgroupI) | 
| 171 | apply (rule subsetI) | |
| 172 | apply (erule H_into_carrier_G) | |
| 173 | apply (rule H_not_empty) | |
| 174 | apply (simp add: H_def) | |
| 175 | apply clarify | |
| 176 | apply (erule_tac P = "\<lambda>z. lhs z = M1" for lhs in subst) | |
| 177 | apply (simp add: coset_mult_assoc ) | |
| 178 | apply (blast intro: H_m_closed) | |
| 179 | done | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 180 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 181 | |
| 64914 | 182 | lemma rcosetGM1g_subset_G: "\<lbrakk>g \<in> carrier G; x \<in> M1 #> g\<rbrakk> \<Longrightarrow> x \<in> carrier G" | 
| 64912 | 183 | by (blast intro: M1_subset_G [THEN r_coset_subset_G, THEN subsetD]) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 184 | |
| 64914 | 185 | lemma finite_M1: "finite M1" | 
| 64912 | 186 | by (rule finite_subset [OF M1_subset_G finite_G]) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 187 | |
| 64914 | 188 | lemma finite_rcosetGM1g: "g \<in> carrier G \<Longrightarrow> finite (M1 #> g)" | 
| 62410 
2fc7a8d9c529
partial tidy-up of Sylow's theorem
 paulson <lp15@cam.ac.uk> parents: 
61382diff
changeset | 189 | using rcosetGM1g_subset_G finite_G M1_subset_G cosets_finite rcosetsI by blast | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 190 | |
| 64914 | 191 | lemma M1_cardeq_rcosetGM1g: "g \<in> carrier G \<Longrightarrow> card (M1 #> g) = card M1" | 
| 64912 | 192 | by (simp add: card_cosets_equal rcosetsI) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 193 | |
| 64914 | 194 | lemma M1_RelM_rcosetGM1g: "g \<in> carrier G \<Longrightarrow> (M1, M1 #> g) \<in> RelM" | 
| 64912 | 195 | apply (simp add: RelM_def calM_def card_M1) | 
| 196 | apply (rule conjI) | |
| 197 | apply (blast intro: rcosetGM1g_subset_G) | |
| 198 | apply (simp add: card_M1 M1_cardeq_rcosetGM1g) | |
| 199 | apply (metis M1_subset_G coset_mult_assoc coset_mult_one r_inv_ex) | |
| 200 | done | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 201 | |
| 64914 | 202 | end | 
| 203 | ||
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 204 | |
| 64912 | 205 | subsection \<open>Equal Cardinalities of \<open>M\<close> and the Set of Cosets\<close> | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 206 | |
| 64912 | 207 | text \<open>Injections between @{term M} and @{term "rcosets\<^bsub>G\<^esub> H"} show that
 | 
| 61382 | 208 | their cardinalities are equal.\<close> | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 209 | |
| 64912 | 210 | lemma ElemClassEquiv: "\<lbrakk>equiv A r; C \<in> A // r\<rbrakk> \<Longrightarrow> \<forall>x \<in> C. \<forall>y \<in> C. (x, y) \<in> r" | 
| 211 | unfolding equiv_def quotient_def sym_def trans_def by blast | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 212 | |
| 64914 | 213 | context sylow_central | 
| 214 | begin | |
| 215 | ||
| 216 | lemma M_elem_map: "M2 \<in> M \<Longrightarrow> \<exists>g. g \<in> carrier G \<and> M1 #> g = M2" | |
| 64912 | 217 | using M1_in_M M_in_quot [THEN RelM_equiv [THEN ElemClassEquiv]] | 
| 218 | by (simp add: RelM_def) (blast dest!: bspec) | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 219 | |
| 64914 | 220 | lemmas M_elem_map_carrier = M_elem_map [THEN someI_ex, THEN conjunct1] | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 221 | |
| 64914 | 222 | lemmas M_elem_map_eq = M_elem_map [THEN someI_ex, THEN conjunct2] | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 223 | |
| 64914 | 224 | lemma M_funcset_rcosets_H: | 
| 64912 | 225 | "(\<lambda>x\<in>M. H #> (SOME g. g \<in> carrier G \<and> M1 #> g = x)) \<in> M \<rightarrow> rcosets H" | 
| 55157 | 226 | by (metis (lifting) H_is_subgroup M_elem_map_carrier rcosetsI restrictI subgroup_imp_subset) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 227 | |
| 64914 | 228 | lemma inj_M_GmodH: "\<exists>f \<in> M \<rightarrow> rcosets H. inj_on f M" | 
| 64912 | 229 | apply (rule bexI) | 
| 230 | apply (rule_tac [2] M_funcset_rcosets_H) | |
| 231 | apply (rule inj_onI, simp) | |
| 232 | apply (rule trans [OF _ M_elem_map_eq]) | |
| 233 | prefer 2 apply assumption | |
| 234 | apply (rule M_elem_map_eq [symmetric, THEN trans], assumption) | |
| 235 | apply (rule coset_mult_inv1) | |
| 236 | apply (erule_tac [2] M_elem_map_carrier)+ | |
| 237 | apply (rule_tac [2] M1_subset_G) | |
| 238 | apply (rule coset_join1 [THEN in_H_imp_eq]) | |
| 239 | apply (rule_tac [3] H_is_subgroup) | |
| 240 | prefer 2 apply (blast intro: M_elem_map_carrier) | |
| 241 | apply (simp add: coset_mult_inv2 H_def M_elem_map_carrier subset_eq) | |
| 242 | done | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 243 | |
| 64914 | 244 | end | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 245 | |
| 64914 | 246 | |
| 247 | subsubsection \<open>The Opposite Injection\<close> | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 248 | |
| 64914 | 249 | context sylow_central | 
| 250 | begin | |
| 251 | ||
| 252 | lemma H_elem_map: "H1 \<in> rcosets H \<Longrightarrow> \<exists>g. g \<in> carrier G \<and> H #> g = H1" | |
| 64912 | 253 | by (auto simp: RCOSETS_def) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 254 | |
| 64914 | 255 | lemmas H_elem_map_carrier = H_elem_map [THEN someI_ex, THEN conjunct1] | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 256 | |
| 64914 | 257 | lemmas H_elem_map_eq = H_elem_map [THEN someI_ex, THEN conjunct2] | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 258 | |
| 64914 | 259 | lemma rcosets_H_funcset_M: | 
| 67613 | 260 | "(\<lambda>C \<in> rcosets H. M1 #> (SOME g. g \<in> carrier G \<and> H #> g = C)) \<in> rcosets H \<rightarrow> M" | 
| 64912 | 261 | apply (simp add: RCOSETS_def) | 
| 262 | apply (fast intro: someI2 | |
| 263 | intro!: M1_in_M in_quotient_imp_closed [OF RelM_equiv M_in_quot _ M1_RelM_rcosetGM1g]) | |
| 264 | done | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 265 | |
| 64912 | 266 | text \<open>Close to a duplicate of \<open>inj_M_GmodH\<close>.\<close> | 
| 64914 | 267 | lemma inj_GmodH_M: "\<exists>g \<in> rcosets H\<rightarrow>M. inj_on g (rcosets H)" | 
| 64912 | 268 | apply (rule bexI) | 
| 269 | apply (rule_tac [2] rcosets_H_funcset_M) | |
| 270 | apply (rule inj_onI) | |
| 271 | apply (simp) | |
| 272 | apply (rule trans [OF _ H_elem_map_eq]) | |
| 273 | prefer 2 apply assumption | |
| 274 | apply (rule H_elem_map_eq [symmetric, THEN trans], assumption) | |
| 275 | apply (rule coset_mult_inv1) | |
| 276 | apply (erule_tac [2] H_elem_map_carrier)+ | |
| 277 | apply (rule_tac [2] H_is_subgroup [THEN subgroup.subset]) | |
| 278 | apply (rule coset_join2) | |
| 279 | apply (blast intro: H_elem_map_carrier) | |
| 280 | apply (rule H_is_subgroup) | |
| 281 | apply (simp add: H_I coset_mult_inv2 H_elem_map_carrier) | |
| 282 | done | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 283 | |
| 64914 | 284 | lemma calM_subset_PowG: "calM \<subseteq> Pow (carrier G)" | 
| 64912 | 285 | by (auto simp: calM_def) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 286 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 287 | |
| 64914 | 288 | lemma finite_M: "finite M" | 
| 64912 | 289 | by (metis M_subset_calM finite_calM rev_finite_subset) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 290 | |
| 64914 | 291 | lemma cardMeqIndexH: "card M = card (rcosets H)" | 
| 64912 | 292 | apply (insert inj_M_GmodH inj_GmodH_M) | 
| 293 | apply (blast intro: card_bij finite_M H_is_subgroup | |
| 294 | rcosets_subset_PowG [THEN finite_subset] | |
| 295 | finite_Pow_iff [THEN iffD2]) | |
| 296 | done | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 297 | |
| 64914 | 298 | lemma index_lem: "card M * card H = order G" | 
| 64912 | 299 | by (simp add: cardMeqIndexH lagrange H_is_subgroup) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 300 | |
| 64914 | 301 | lemma lemma_leq1: "p^a \<le> card H" | 
| 64912 | 302 | apply (rule dvd_imp_le) | 
| 303 | apply (rule div_combine [OF prime_imp_prime_elem[OF prime_p] not_dvd_M]) | |
| 304 | prefer 2 apply (blast intro: subgroup.finite_imp_card_positive H_is_subgroup) | |
| 305 | apply (simp add: index_lem order_G power_add mult_dvd_mono multiplicity_dvd zero_less_m) | |
| 306 | done | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 307 | |
| 64914 | 308 | lemma lemma_leq2: "card H \<le> p^a" | 
| 64912 | 309 | apply (subst card_M1 [symmetric]) | 
| 310 | apply (cut_tac M1_inj_H) | |
| 311 | apply (blast intro!: M1_subset_G intro: card_inj H_into_carrier_G finite_subset [OF _ finite_G]) | |
| 312 | done | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 313 | |
| 64914 | 314 | lemma card_H_eq: "card H = p^a" | 
| 64912 | 315 | by (blast intro: le_antisym lemma_leq1 lemma_leq2) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 316 | |
| 64914 | 317 | end | 
| 318 | ||
| 64912 | 319 | lemma (in sylow) sylow_thm: "\<exists>H. subgroup H G \<and> card H = p^a" | 
| 320 | using lemma_A1 | |
| 321 | apply clarify | |
| 322 | apply (frule existsM1inM, clarify) | |
| 323 | apply (subgoal_tac "sylow_central G p a m M1 M") | |
| 324 | apply (blast dest: sylow_central.H_is_subgroup sylow_central.card_H_eq) | |
| 325 | apply (simp add: sylow_central_def sylow_central_axioms_def sylow_axioms calM_def RelM_def) | |
| 326 | done | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 327 | |
| 64912 | 328 | text \<open>Needed because the locale's automatic definition refers to | 
| 329 |   @{term "semigroup G"} and @{term "group_axioms G"} rather than
 | |
| 61382 | 330 |   simply to @{term "group G"}.\<close>
 | 
| 64912 | 331 | lemma sylow_eq: "sylow G p a m \<longleftrightarrow> group G \<and> sylow_axioms G p a m" | 
| 332 | by (simp add: sylow_def group_def) | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 333 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
16663diff
changeset | 334 | |
| 61382 | 335 | subsection \<open>Sylow's Theorem\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
16663diff
changeset | 336 | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 337 | theorem sylow_thm: | 
| 64912 | 338 | "\<lbrakk>prime p; group G; order G = (p^a) * m; finite (carrier G)\<rbrakk> | 
| 339 | \<Longrightarrow> \<exists>H. subgroup H G \<and> card H = p^a" | |
| 340 | by (rule sylow.sylow_thm [of G p a m]) (simp add: sylow_eq sylow_axioms_def) | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 341 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 342 | end |