1 (* Title: HOL/Algebra/Sylow.thy
2 Author: Florian Kammueller, with new proofs by L C Paulson
10 See also \cite{Kammueller-Paulson:1999}.
13 text{*The combinatorial argument is in theory Exponent*}
15 locale sylow = group +
16 fixes p and a and m and calM and RelM
17 assumes prime_p: "prime p"
18 and order_G: "order(G) = (p^a) * m"
19 and finite_G [iff]: "finite (carrier G)"
20 defines "calM == {s. s \<subseteq> carrier(G) & card(s) = p^a}"
21 and "RelM == {(N1,N2). N1 \<in> calM & N2 \<in> calM &
22 (\<exists>g \<in> carrier(G). N1 = (N2 #> g) )}"
24 lemma (in sylow) RelM_refl_on: "refl_on calM RelM"
25 apply (auto simp add: refl_on_def RelM_def calM_def)
26 apply (blast intro!: coset_mult_one [symmetric])
29 lemma (in sylow) RelM_sym: "sym RelM"
30 proof (unfold sym_def RelM_def, clarify)
33 and g: "g \<in> carrier G"
34 hence "y = y #> g #> (inv g)" by (simp add: coset_mult_assoc calM_def)
35 thus "\<exists>g'\<in>carrier G. y = y #> g #> g'" by (blast intro: g)
38 lemma (in sylow) RelM_trans: "trans RelM"
39 by (auto simp add: trans_def RelM_def calM_def coset_mult_assoc)
41 lemma (in sylow) RelM_equiv: "equiv calM RelM"
42 apply (unfold equiv_def)
43 apply (blast intro: RelM_refl_on RelM_sym RelM_trans)
46 lemma (in sylow) M_subset_calM_prep: "M' \<in> calM // RelM ==> M' \<subseteq> calM"
47 apply (unfold RelM_def)
48 apply (blast elim!: quotientE)
52 subsection{*Main Part of the Proof*}
54 locale sylow_central = sylow +
56 assumes M_in_quot: "M \<in> calM // RelM"
57 and not_dvd_M: "~(p ^ Suc(exponent p m) dvd card(M))"
58 and M1_in_M: "M1 \<in> M"
59 defines "H == {g. g\<in>carrier G & M1 #> g = M1}"
61 lemma (in sylow_central) M_subset_calM: "M \<subseteq> calM"
62 by (rule M_in_quot [THEN M_subset_calM_prep])
64 lemma (in sylow_central) card_M1: "card(M1) = p^a"
65 apply (cut_tac M_subset_calM M1_in_M)
66 apply (simp add: calM_def, blast)
69 lemma card_nonempty: "0 < card(S) ==> S \<noteq> {}"
72 lemma (in sylow_central) exists_x_in_M1: "\<exists>x. x\<in>M1"
73 apply (subgoal_tac "0 < card M1")
74 apply (blast dest: card_nonempty)
75 apply (cut_tac prime_p [THEN prime_imp_one_less])
76 apply (simp (no_asm_simp) add: card_M1)
79 lemma (in sylow_central) M1_subset_G [simp]: "M1 \<subseteq> carrier G"
80 apply (rule subsetD [THEN PowD])
81 apply (rule_tac [2] M1_in_M)
82 apply (rule M_subset_calM [THEN subset_trans])
83 apply (auto simp add: calM_def)
86 lemma (in sylow_central) M1_inj_H: "\<exists>f \<in> H\<rightarrow>M1. inj_on f H"
88 from exists_x_in_M1 obtain m1 where m1M: "m1 \<in> M1"..
89 have m1G: "m1 \<in> carrier G" by (simp add: m1M M1_subset_G [THEN subsetD])
92 show "inj_on (\<lambda>z\<in>H. m1 \<otimes> z) H"
93 by (simp add: inj_on_def l_cancel [of m1 x y, THEN iffD1] H_def m1G)
94 show "restrict (op \<otimes> m1) H \<in> H \<rightarrow> M1"
95 proof (rule restrictI)
96 fix z assume zH: "z \<in> H"
97 show "m1 \<otimes> z \<in> M1"
100 have zG: "z \<in> carrier G" and M1zeq: "M1 #> z = M1"
101 by (auto simp add: H_def)
103 by (rule subst [OF M1zeq], simp add: m1M zG rcosI)
110 subsection{*Discharging the Assumptions of @{text sylow_central}*}
112 lemma (in sylow) EmptyNotInEquivSet: "{} \<notin> calM // RelM"
113 by (blast elim!: quotientE dest: RelM_equiv [THEN equiv_class_self])
115 lemma (in sylow) existsM1inM: "M \<in> calM // RelM ==> \<exists>M1. M1 \<in> M"
116 apply (subgoal_tac "M \<noteq> {}")
118 apply (cut_tac EmptyNotInEquivSet, blast)
121 lemma (in sylow) zero_less_o_G: "0 < order(G)"
122 apply (unfold order_def)
123 apply (blast intro: zero_less_card_empty)
126 lemma (in sylow) zero_less_m: "m > 0"
127 apply (cut_tac zero_less_o_G)
128 apply (simp add: order_G)
131 lemma (in sylow) card_calM: "card(calM) = (p^a) * m choose p^a"
132 by (simp add: calM_def n_subsets order_G [symmetric] order_def)
134 lemma (in sylow) zero_less_card_calM: "card calM > 0"
135 by (simp add: card_calM zero_less_binomial le_extend_mult zero_less_m)
137 lemma (in sylow) max_p_div_calM:
138 "~ (p ^ Suc(exponent p m) dvd card(calM))"
139 apply (subgoal_tac "exponent p m = exponent p (card calM) ")
140 apply (cut_tac zero_less_card_calM prime_p)
141 apply (force dest: power_Suc_exponent_Not_dvd)
142 apply (simp add: card_calM zero_less_m [THEN const_p_fac])
145 lemma (in sylow) finite_calM: "finite calM"
146 apply (unfold calM_def)
147 apply (rule_tac B = "Pow (carrier G) " in finite_subset)
151 lemma (in sylow) lemma_A1:
152 "\<exists>M \<in> calM // RelM. ~ (p ^ Suc(exponent p m) dvd card(M))"
153 apply (rule max_p_div_calM [THEN contrapos_np])
154 apply (simp add: finite_calM equiv_imp_dvd_card [OF _ RelM_equiv])
158 subsubsection{*Introduction and Destruct Rules for @{term H}*}
160 lemma (in sylow_central) H_I: "[|g \<in> carrier G; M1 #> g = M1|] ==> g \<in> H"
163 lemma (in sylow_central) H_into_carrier_G: "x \<in> H ==> x \<in> carrier G"
166 lemma (in sylow_central) in_H_imp_eq: "g : H ==> M1 #> g = M1"
169 lemma (in sylow_central) H_m_closed: "[| x\<in>H; y\<in>H|] ==> x \<otimes> y \<in> H"
171 apply (simp add: coset_mult_assoc [symmetric])
174 lemma (in sylow_central) H_not_empty: "H \<noteq> {}"
175 apply (simp add: H_def)
176 apply (rule exI [of _ \<one>], simp)
179 lemma (in sylow_central) H_is_subgroup: "subgroup H G"
180 apply (rule subgroupI)
182 apply (erule H_into_carrier_G)
183 apply (rule H_not_empty)
184 apply (simp add: H_def, clarify)
185 apply (erule_tac P = "%z. ?lhs(z) = M1" in subst)
186 apply (simp add: coset_mult_assoc )
187 apply (blast intro: H_m_closed)
191 lemma (in sylow_central) rcosetGM1g_subset_G:
192 "[| g \<in> carrier G; x \<in> M1 #> g |] ==> x \<in> carrier G"
193 by (blast intro: M1_subset_G [THEN r_coset_subset_G, THEN subsetD])
195 lemma (in sylow_central) finite_M1: "finite M1"
196 by (rule finite_subset [OF M1_subset_G finite_G])
198 lemma (in sylow_central) finite_rcosetGM1g: "g\<in>carrier G ==> finite (M1 #> g)"
199 apply (rule finite_subset)
201 apply (erule rcosetGM1g_subset_G, assumption)
202 apply (rule finite_G)
205 lemma (in sylow_central) M1_cardeq_rcosetGM1g:
206 "g \<in> carrier G ==> card(M1 #> g) = card(M1)"
207 by (simp (no_asm_simp) add: card_cosets_equal rcosetsI)
209 lemma (in sylow_central) M1_RelM_rcosetGM1g:
210 "g \<in> carrier G ==> (M1, M1 #> g) \<in> RelM"
211 apply (simp (no_asm) add: RelM_def calM_def card_M1)
213 apply (blast intro: rcosetGM1g_subset_G)
214 apply (simp (no_asm_simp) add: card_M1 M1_cardeq_rcosetGM1g)
215 apply (rule bexI [of _ "inv g"])
216 apply (simp_all add: coset_mult_assoc)
220 subsection{*Equal Cardinalities of @{term M} and the Set of Cosets*}
222 text{*Injections between @{term M} and @{term "rcosets\<^bsub>G\<^esub> H"} show that
223 their cardinalities are equal.*}
225 lemma ElemClassEquiv:
226 "[| equiv A r; C \<in> A // r |] ==> \<forall>x \<in> C. \<forall>y \<in> C. (x,y)\<in>r"
227 by (unfold equiv_def quotient_def sym_def trans_def, blast)
229 lemma (in sylow_central) M_elem_map:
230 "M2 \<in> M ==> \<exists>g. g \<in> carrier G & M1 #> g = M2"
231 apply (cut_tac M1_in_M M_in_quot [THEN RelM_equiv [THEN ElemClassEquiv]])
232 apply (simp add: RelM_def)
233 apply (blast dest!: bspec)
236 lemmas (in sylow_central) M_elem_map_carrier =
237 M_elem_map [THEN someI_ex, THEN conjunct1]
239 lemmas (in sylow_central) M_elem_map_eq =
240 M_elem_map [THEN someI_ex, THEN conjunct2]
242 lemma (in sylow_central) M_funcset_rcosets_H:
243 "(%x:M. H #> (SOME g. g \<in> carrier G & M1 #> g = x)) \<in> M \<rightarrow> rcosets H"
244 apply (rule rcosetsI [THEN restrictI])
245 apply (rule H_is_subgroup [THEN subgroup.subset])
246 apply (erule M_elem_map_carrier)
249 lemma (in sylow_central) inj_M_GmodH: "\<exists>f \<in> M\<rightarrow>rcosets H. inj_on f M"
251 apply (rule_tac [2] M_funcset_rcosets_H)
252 apply (rule inj_onI, simp)
253 apply (rule trans [OF _ M_elem_map_eq])
254 prefer 2 apply assumption
255 apply (rule M_elem_map_eq [symmetric, THEN trans], assumption)
256 apply (rule coset_mult_inv1)
257 apply (erule_tac [2] M_elem_map_carrier)+
258 apply (rule_tac [2] M1_subset_G)
259 apply (rule coset_join1 [THEN in_H_imp_eq])
260 apply (rule_tac [3] H_is_subgroup)
261 prefer 2 apply (blast intro: M_elem_map_carrier)
262 apply (simp add: coset_mult_inv2 H_def M_elem_map_carrier subset_eq)
266 subsubsection{*The Opposite Injection*}
268 lemma (in sylow_central) H_elem_map:
269 "H1 \<in> rcosets H ==> \<exists>g. g \<in> carrier G & H #> g = H1"
270 by (auto simp add: RCOSETS_def)
272 lemmas (in sylow_central) H_elem_map_carrier =
273 H_elem_map [THEN someI_ex, THEN conjunct1]
275 lemmas (in sylow_central) H_elem_map_eq =
276 H_elem_map [THEN someI_ex, THEN conjunct2]
279 lemma EquivElemClass:
280 "[|equiv A r; M \<in> A//r; M1\<in>M; (M1,M2) \<in> r |] ==> M2 \<in> M"
281 by (unfold equiv_def quotient_def sym_def trans_def, blast)
284 lemma (in sylow_central) rcosets_H_funcset_M:
285 "(\<lambda>C \<in> rcosets H. M1 #> (@g. g \<in> carrier G \<and> H #> g = C)) \<in> rcosets H \<rightarrow> M"
286 apply (simp add: RCOSETS_def)
287 apply (fast intro: someI2
288 intro!: M1_in_M EquivElemClass [OF RelM_equiv M_in_quot _ M1_RelM_rcosetGM1g])
291 text{*close to a duplicate of @{text inj_M_GmodH}*}
292 lemma (in sylow_central) inj_GmodH_M:
293 "\<exists>g \<in> rcosets H\<rightarrow>M. inj_on g (rcosets H)"
295 apply (rule_tac [2] rcosets_H_funcset_M)
298 apply (rule trans [OF _ H_elem_map_eq])
299 prefer 2 apply assumption
300 apply (rule H_elem_map_eq [symmetric, THEN trans], assumption)
301 apply (rule coset_mult_inv1)
302 apply (erule_tac [2] H_elem_map_carrier)+
303 apply (rule_tac [2] H_is_subgroup [THEN subgroup.subset])
304 apply (rule coset_join2)
305 apply (blast intro: H_elem_map_carrier)
306 apply (rule H_is_subgroup)
307 apply (simp add: H_I coset_mult_inv2 H_elem_map_carrier)
310 lemma (in sylow_central) calM_subset_PowG: "calM \<subseteq> Pow(carrier G)"
311 by (auto simp add: calM_def)
314 lemma (in sylow_central) finite_M: "finite M"
315 apply (rule finite_subset)
316 apply (rule M_subset_calM [THEN subset_trans])
317 apply (rule calM_subset_PowG, blast)
320 lemma (in sylow_central) cardMeqIndexH: "card(M) = card(rcosets H)"
321 apply (insert inj_M_GmodH inj_GmodH_M)
322 apply (blast intro: card_bij finite_M H_is_subgroup
323 rcosets_subset_PowG [THEN finite_subset]
324 finite_Pow_iff [THEN iffD2])
327 lemma (in sylow_central) index_lem: "card(M) * card(H) = order(G)"
328 by (simp add: cardMeqIndexH lagrange H_is_subgroup)
330 lemma (in sylow_central) lemma_leq1: "p^a \<le> card(H)"
331 apply (rule dvd_imp_le)
332 apply (rule div_combine [OF prime_p not_dvd_M])
333 prefer 2 apply (blast intro: subgroup.finite_imp_card_positive H_is_subgroup)
334 apply (simp add: index_lem order_G power_add mult_dvd_mono power_exponent_dvd
338 lemma (in sylow_central) lemma_leq2: "card(H) \<le> p^a"
339 apply (subst card_M1 [symmetric])
340 apply (cut_tac M1_inj_H)
341 apply (blast intro!: M1_subset_G intro:
342 card_inj H_into_carrier_G finite_subset [OF _ finite_G])
345 lemma (in sylow_central) card_H_eq: "card(H) = p^a"
346 by (blast intro: le_antisym lemma_leq1 lemma_leq2)
348 lemma (in sylow) sylow_thm: "\<exists>H. subgroup H G & card(H) = p^a"
349 apply (cut_tac lemma_A1, clarify)
350 apply (frule existsM1inM, clarify)
351 apply (subgoal_tac "sylow_central G p a m M1 M")
352 apply (blast dest: sylow_central.H_is_subgroup sylow_central.card_H_eq)
353 apply (simp add: sylow_central_def sylow_central_axioms_def sylow_axioms calM_def RelM_def)
356 text{*Needed because the locale's automatic definition refers to
357 @{term "semigroup G"} and @{term "group_axioms G"} rather than
358 simply to @{term "group G"}.*}
359 lemma sylow_eq: "sylow G p a m = (group G & sylow_axioms G p a m)"
360 by (simp add: sylow_def group_def)
363 subsection {* Sylow's Theorem *}
366 "[| prime p; group(G); order(G) = (p^a) * m; finite (carrier G)|]
367 ==> \<exists>H. subgroup H G & card(H) = p^a"
368 apply (rule sylow.sylow_thm [of G p a m])
369 apply (simp add: sylow_eq sylow_axioms_def)