author | wenzelm |
Tue, 24 Sep 2024 21:31:20 +0200 | |
changeset 80946 | b76f64d7d493 |
parent 80400 | 898034c8a799 |
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 |
|
76987 | 9 |
text \<open>See also \<^cite>\<open>"Kammueller-Paulson:1999"\<close>.\<close> |
64912 | 10 |
|
69272 | 11 |
text \<open>The combinatorial argument is in theory \<open>Exponent\<close>.\<close> |
14706 | 12 |
|
80400 | 13 |
lemma le_extend_mult: "\<lbrakk>0 < c; a \<le> b\<rbrakk> \<Longrightarrow> a \<le> b * c" for c :: nat |
14 |
using gr0_conv_Suc by fastforce |
|
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
15 |
|
14747 | 16 |
locale sylow = group + |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
17 |
fixes p and a and m and calM and RelM |
64912 | 18 |
assumes prime_p: "prime p" |
19 |
and order_G: "order G = (p^a) * m" |
|
20 |
and finite_G[iff]: "finite (carrier G)" |
|
21 |
defines "calM \<equiv> {s. s \<subseteq> carrier G \<and> card s = p^a}" |
|
22 |
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:
61382
diff
changeset
|
23 |
begin |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
24 |
|
80067 | 25 |
lemma RelM_subset: "RelM \<subseteq> calM \<times> calM" |
26 |
by (auto simp only: RelM_def) |
|
27 |
||
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
28 |
lemma RelM_refl_on: "refl_on calM RelM" |
64912 | 29 |
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
|
30 |
|
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
31 |
lemma RelM_sym: "sym RelM" |
80400 | 32 |
unfolding sym_def RelM_def calM_def |
33 |
using coset_mult_assoc coset_mult_one r_inv_ex |
|
34 |
by (smt (verit, best) case_prod_conv mem_Collect_eq) |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
35 |
|
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
36 |
lemma RelM_trans: "trans RelM" |
64912 | 37 |
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
|
38 |
|
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
39 |
lemma RelM_equiv: "equiv calM RelM" |
80067 | 40 |
using RelM_subset RelM_refl_on RelM_sym RelM_trans by (intro equivI) |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
41 |
|
64912 | 42 |
lemma M_subset_calM_prep: "M' \<in> calM // RelM \<Longrightarrow> M' \<subseteq> calM" |
43 |
unfolding RelM_def by (blast elim!: quotientE) |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
44 |
|
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
45 |
end |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
16663
diff
changeset
|
46 |
|
64912 | 47 |
subsection \<open>Main Part of the Proof\<close> |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
48 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
49 |
locale sylow_central = sylow + |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
50 |
fixes H and M1 and M |
64912 | 51 |
assumes M_in_quot: "M \<in> calM // RelM" |
52 |
and not_dvd_M: "\<not> (p ^ Suc (multiplicity p m) dvd card M)" |
|
53 |
and M1_in_M: "M1 \<in> M" |
|
54 |
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:
61382
diff
changeset
|
55 |
begin |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
56 |
|
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
57 |
lemma M_subset_calM: "M \<subseteq> calM" |
80400 | 58 |
by (simp add: M_in_quot M_subset_calM_prep) |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
59 |
|
64912 | 60 |
lemma card_M1: "card M1 = p^a" |
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
61 |
using M1_in_M M_subset_calM calM_def by blast |
64912 | 62 |
|
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
63 |
lemma exists_x_in_M1: "\<exists>x. x \<in> M1" |
80400 | 64 |
using prime_p [THEN prime_gt_Suc_0_nat] card_M1 one_in_subset by fastforce |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
65 |
|
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
66 |
lemma M1_subset_G [simp]: "M1 \<subseteq> carrier G" |
64912 | 67 |
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
|
68 |
|
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
69 |
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:
61382
diff
changeset
|
70 |
proof - |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
71 |
from exists_x_in_M1 obtain m1 where m1M: "m1 \<in> M1".. |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
72 |
show ?thesis |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
73 |
proof |
80400 | 74 |
have "m1 \<in> carrier G" |
75 |
by (simp add: m1M M1_subset_G [THEN subsetD]) |
|
76 |
then show "inj_on (\<lambda>z\<in>H. m1 \<otimes> z) H" |
|
77 |
by (simp add: H_def inj_on_def) |
|
67399 | 78 |
show "restrict ((\<otimes>) m1) H \<in> H \<rightarrow> M1" |
80400 | 79 |
using H_def m1M rcosI by auto |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
80 |
qed |
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
81 |
qed |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
82 |
|
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
83 |
end |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
84 |
|
64912 | 85 |
|
86 |
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
|
87 |
|
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
88 |
context sylow |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
89 |
begin |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
90 |
|
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
91 |
lemma EmptyNotInEquivSet: "{} \<notin> calM // RelM" |
80400 | 92 |
using RelM_equiv in_quotient_imp_non_empty by blast |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
93 |
|
64912 | 94 |
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:
61382
diff
changeset
|
95 |
using RelM_equiv equiv_Eps_in by blast |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
96 |
|
64912 | 97 |
lemma zero_less_o_G: "0 < order G" |
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
98 |
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
|
99 |
|
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
100 |
lemma zero_less_m: "m > 0" |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
101 |
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
|
102 |
|
64912 | 103 |
lemma card_calM: "card calM = (p^a) * m choose p^a" |
104 |
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
|
105 |
|
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
106 |
lemma zero_less_card_calM: "card calM > 0" |
64912 | 107 |
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
|
108 |
|
64912 | 109 |
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:
63167
diff
changeset
|
110 |
proof |
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
63167
diff
changeset
|
111 |
assume "p ^ Suc (multiplicity p m) dvd card calM" |
64912 | 112 |
with zero_less_card_calM prime_p |
63534
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
63167
diff
changeset
|
113 |
have "Suc (multiplicity p m) \<le> multiplicity p (card calM)" |
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
63167
diff
changeset
|
114 |
by (intro multiplicity_geI) auto |
80400 | 115 |
then show False |
116 |
by (simp add: card_calM const_p_fac prime_p zero_less_m) |
|
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
117 |
qed |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
118 |
|
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
119 |
lemma finite_calM: "finite calM" |
64912 | 120 |
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
|
121 |
|
64912 | 122 |
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:
61382
diff
changeset
|
123 |
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
|
124 |
|
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
125 |
end |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
126 |
|
64912 | 127 |
|
128 |
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
|
129 |
|
64914 | 130 |
context sylow_central |
131 |
begin |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
132 |
|
64914 | 133 |
lemma H_I: "\<lbrakk>g \<in> carrier G; M1 #> g = M1\<rbrakk> \<Longrightarrow> g \<in> H" |
64912 | 134 |
by (simp add: H_def) |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
135 |
|
64914 | 136 |
lemma H_into_carrier_G: "x \<in> H \<Longrightarrow> x \<in> carrier G" |
64912 | 137 |
by (simp add: H_def) |
138 |
||
64914 | 139 |
lemma in_H_imp_eq: "g \<in> H \<Longrightarrow> M1 #> g = M1" |
140 |
by (simp add: H_def) |
|
141 |
||
142 |
lemma H_m_closed: "\<lbrakk>x \<in> H; y \<in> H\<rbrakk> \<Longrightarrow> x \<otimes> y \<in> H" |
|
64912 | 143 |
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
|
144 |
|
64914 | 145 |
lemma H_not_empty: "H \<noteq> {}" |
68488
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
146 |
by (force simp add: H_def intro: exI [of _ \<one>]) |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
147 |
|
64914 | 148 |
lemma H_is_subgroup: "subgroup H G" |
68488
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
149 |
proof (rule subgroupI) |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
150 |
show "H \<subseteq> carrier G" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
151 |
using H_into_carrier_G by blast |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
152 |
show "\<And>a. a \<in> H \<Longrightarrow> inv a \<in> H" |
80400 | 153 |
by (metis H_I H_into_carrier_G M1_subset_G coset_mult_assoc coset_mult_one in_H_imp_eq inv_closed r_inv) |
68488
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
154 |
show "\<And>a b. \<lbrakk>a \<in> H; b \<in> H\<rbrakk> \<Longrightarrow> a \<otimes> b \<in> H" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
155 |
by (blast intro: H_m_closed) |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
156 |
qed (use H_not_empty in auto) |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
157 |
|
64914 | 158 |
lemma rcosetGM1g_subset_G: "\<lbrakk>g \<in> carrier G; x \<in> M1 #> g\<rbrakk> \<Longrightarrow> x \<in> carrier G" |
64912 | 159 |
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
|
160 |
|
64914 | 161 |
lemma finite_M1: "finite M1" |
64912 | 162 |
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
|
163 |
|
64914 | 164 |
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:
61382
diff
changeset
|
165 |
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
|
166 |
|
64914 | 167 |
lemma M1_cardeq_rcosetGM1g: "g \<in> carrier G \<Longrightarrow> card (M1 #> g) = card M1" |
68443
43055b016688
New material from Martin Baillon and Paulo EmÃlio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68399
diff
changeset
|
168 |
by (metis M1_subset_G card_rcosets_equal rcosetsI) |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
169 |
|
68488
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
170 |
lemma M1_RelM_rcosetGM1g: |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
171 |
assumes "g \<in> carrier G" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
172 |
shows "(M1, M1 #> g) \<in> RelM" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
173 |
proof - |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
174 |
have "M1 #> g \<subseteq> carrier G" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
175 |
by (simp add: assms r_coset_subset_G) |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
176 |
moreover have "card (M1 #> g) = p ^ a" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
177 |
using assms by (simp add: card_M1 M1_cardeq_rcosetGM1g) |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
178 |
moreover have "\<exists>h\<in>carrier G. M1 = M1 #> g #> h" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
179 |
by (metis assms M1_subset_G coset_mult_assoc coset_mult_one r_inv_ex) |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
180 |
ultimately show ?thesis |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
181 |
by (simp add: RelM_def calM_def card_M1) |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
182 |
qed |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
183 |
|
64914 | 184 |
end |
185 |
||
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
186 |
|
64912 | 187 |
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
|
188 |
|
69597 | 189 |
text \<open>Injections between \<^term>\<open>M\<close> and \<^term>\<open>rcosets\<^bsub>G\<^esub> H\<close> show that |
61382 | 190 |
their cardinalities are equal.\<close> |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
191 |
|
64912 | 192 |
lemma ElemClassEquiv: "\<lbrakk>equiv A r; C \<in> A // r\<rbrakk> \<Longrightarrow> \<forall>x \<in> C. \<forall>y \<in> C. (x, y) \<in> r" |
193 |
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
|
194 |
|
64914 | 195 |
context sylow_central |
196 |
begin |
|
197 |
||
198 |
lemma M_elem_map: "M2 \<in> M \<Longrightarrow> \<exists>g. g \<in> carrier G \<and> M1 #> g = M2" |
|
64912 | 199 |
using M1_in_M M_in_quot [THEN RelM_equiv [THEN ElemClassEquiv]] |
200 |
by (simp add: RelM_def) (blast dest!: bspec) |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
201 |
|
64914 | 202 |
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
|
203 |
|
64914 | 204 |
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
|
205 |
|
64914 | 206 |
lemma M_funcset_rcosets_H: |
64912 | 207 |
"(\<lambda>x\<in>M. H #> (SOME g. g \<in> carrier G \<and> M1 #> g = x)) \<in> M \<rightarrow> rcosets H" |
68445
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents:
68443
diff
changeset
|
208 |
by (metis (lifting) H_is_subgroup M_elem_map_carrier rcosetsI restrictI subgroup.subset) |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
209 |
|
64914 | 210 |
lemma inj_M_GmodH: "\<exists>f \<in> M \<rightarrow> rcosets H. inj_on f M" |
68488
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
211 |
proof |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
212 |
let ?inv = "\<lambda>x. SOME g. g \<in> carrier G \<and> M1 #> g = x" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
213 |
show "inj_on (\<lambda>x\<in>M. H #> ?inv x) M" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
214 |
proof (rule inj_onI, simp) |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
215 |
fix x y |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
216 |
assume eq: "H #> ?inv x = H #> ?inv y" and xy: "x \<in> M" "y \<in> M" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
217 |
have "x = M1 #> ?inv x" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
218 |
by (simp add: M_elem_map_eq \<open>x \<in> M\<close>) |
80400 | 219 |
also have "\<dots> = M1 #> ?inv y" |
68488
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
220 |
proof (rule coset_mult_inv1 [OF in_H_imp_eq [OF coset_join1]]) |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
221 |
show "H #> ?inv x \<otimes> inv (?inv y) = H" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
222 |
by (simp add: H_into_carrier_G M_elem_map_carrier xy coset_mult_inv2 eq subsetI) |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
223 |
qed (simp_all add: H_is_subgroup M_elem_map_carrier xy) |
80400 | 224 |
also have "\<dots> = y" |
68488
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
225 |
using M_elem_map_eq \<open>y \<in> M\<close> by simp |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
226 |
finally show "x=y" . |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
227 |
qed |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
228 |
show "(\<lambda>x\<in>M. H #> ?inv x) \<in> M \<rightarrow> rcosets H" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
229 |
by (rule M_funcset_rcosets_H) |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
230 |
qed |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
231 |
|
64914 | 232 |
end |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
233 |
|
64914 | 234 |
|
235 |
subsubsection \<open>The Opposite Injection\<close> |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
236 |
|
64914 | 237 |
context sylow_central |
238 |
begin |
|
239 |
||
240 |
lemma H_elem_map: "H1 \<in> rcosets H \<Longrightarrow> \<exists>g. g \<in> carrier G \<and> H #> g = H1" |
|
64912 | 241 |
by (auto simp: RCOSETS_def) |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
242 |
|
64914 | 243 |
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
|
244 |
|
64914 | 245 |
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
|
246 |
|
64914 | 247 |
lemma rcosets_H_funcset_M: |
67613 | 248 |
"(\<lambda>C \<in> rcosets H. M1 #> (SOME g. g \<in> carrier G \<and> H #> g = C)) \<in> rcosets H \<rightarrow> M" |
68488
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
249 |
using in_quotient_imp_closed [OF RelM_equiv M_in_quot _ M1_RelM_rcosetGM1g] |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
250 |
by (simp add: M1_in_M H_elem_map_carrier RCOSETS_def) |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
251 |
|
64914 | 252 |
lemma inj_GmodH_M: "\<exists>g \<in> rcosets H\<rightarrow>M. inj_on g (rcosets H)" |
68488
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
253 |
proof |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
254 |
let ?inv = "\<lambda>x. SOME g. g \<in> carrier G \<and> H #> g = x" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
255 |
show "inj_on (\<lambda>C\<in>rcosets H. M1 #> ?inv C) (rcosets H)" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
256 |
proof (rule inj_onI, simp) |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
257 |
fix x y |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
258 |
assume eq: "M1 #> ?inv x = M1 #> ?inv y" and xy: "x \<in> rcosets H" "y \<in> rcosets H" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
259 |
have "x = H #> ?inv x" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
260 |
by (simp add: H_elem_map_eq \<open>x \<in> rcosets H\<close>) |
80400 | 261 |
also have "\<dots> = H #> ?inv y" |
68488
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
262 |
proof (rule coset_mult_inv1 [OF coset_join2]) |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
263 |
show "?inv x \<otimes> inv (?inv y) \<in> carrier G" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
264 |
by (simp add: H_elem_map_carrier \<open>x \<in> rcosets H\<close> \<open>y \<in> rcosets H\<close>) |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
265 |
then show "(?inv x) \<otimes> inv (?inv y) \<in> H" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
266 |
by (simp add: H_I H_elem_map_carrier xy coset_mult_inv2 eq) |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
267 |
show "H \<subseteq> carrier G" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
268 |
by (simp add: H_is_subgroup subgroup.subset) |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
269 |
qed (simp_all add: H_is_subgroup H_elem_map_carrier xy) |
80400 | 270 |
also have "\<dots> = y" |
68488
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
271 |
by (simp add: H_elem_map_eq \<open>y \<in> rcosets H\<close>) |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
272 |
finally show "x=y" . |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
273 |
qed |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
274 |
show "(\<lambda>C\<in>rcosets H. M1 #> ?inv C) \<in> rcosets H \<rightarrow> M" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
275 |
using rcosets_H_funcset_M by blast |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
276 |
qed |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
277 |
|
64914 | 278 |
lemma calM_subset_PowG: "calM \<subseteq> Pow (carrier G)" |
64912 | 279 |
by (auto simp: calM_def) |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
280 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
281 |
|
64914 | 282 |
lemma finite_M: "finite M" |
64912 | 283 |
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
|
284 |
|
64914 | 285 |
lemma cardMeqIndexH: "card M = card (rcosets H)" |
68488
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
286 |
using inj_M_GmodH inj_GmodH_M |
80400 | 287 |
by (metis H_is_subgroup card_bij finite_G finite_M finite_UnionD rcosets_part_G) |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
288 |
|
64914 | 289 |
lemma index_lem: "card M * card H = order G" |
64912 | 290 |
by (simp add: cardMeqIndexH lagrange H_is_subgroup) |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
291 |
|
64914 | 292 |
lemma card_H_eq: "card H = p^a" |
68488
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
293 |
proof (rule antisym) |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
294 |
show "p^a \<le> card H" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
295 |
proof (rule dvd_imp_le) |
80400 | 296 |
have "p ^ (a + multiplicity p m) dvd card M * card H" |
68488
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
297 |
by (simp add: index_lem multiplicity_dvd order_G power_add) |
80400 | 298 |
then show "p ^ a dvd card H" |
299 |
using div_combine not_dvd_M prime_p by blast |
|
68488
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
300 |
show "0 < card H" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
301 |
by (blast intro: subgroup.finite_imp_card_positive H_is_subgroup) |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
302 |
qed |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
303 |
next |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
304 |
show "card H \<le> p^a" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
305 |
using M1_inj_H card_M1 card_inj finite_M1 by fastforce |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
306 |
qed |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
307 |
|
64914 | 308 |
end |
309 |
||
64912 | 310 |
lemma (in sylow) sylow_thm: "\<exists>H. subgroup H G \<and> card H = p^a" |
68488
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
311 |
proof - |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
312 |
obtain M where M: "M \<in> calM // RelM" "\<not> (p ^ Suc (multiplicity p m) dvd card M)" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
313 |
using lemma_A1 by blast |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
314 |
then obtain M1 where "M1 \<in> M" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
315 |
by (metis existsM1inM) |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
316 |
define H where "H \<equiv> {g. g \<in> carrier G \<and> M1 #> g = M1}" |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
317 |
with M \<open>M1 \<in> M\<close> |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
318 |
interpret sylow_central G p a m calM RelM H M1 M |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
319 |
by unfold_locales (auto simp add: H_def calM_def RelM_def) |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
320 |
show ?thesis |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
321 |
using H_is_subgroup card_H_eq by blast |
dfbd80c3d180
more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents:
68484
diff
changeset
|
322 |
qed |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
323 |
|
64912 | 324 |
text \<open>Needed because the locale's automatic definition refers to |
69597 | 325 |
\<^term>\<open>semigroup G\<close> and \<^term>\<open>group_axioms G\<close> rather than |
326 |
simply to \<^term>\<open>group G\<close>.\<close> |
|
64912 | 327 |
lemma sylow_eq: "sylow G p a m \<longleftrightarrow> group G \<and> sylow_axioms G p a m" |
328 |
by (simp add: sylow_def group_def) |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
329 |
|
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
16663
diff
changeset
|
330 |
|
61382 | 331 |
subsection \<open>Sylow's Theorem\<close> |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
16663
diff
changeset
|
332 |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
333 |
theorem sylow_thm: |
64912 | 334 |
"\<lbrakk>prime p; group G; order G = (p^a) * m; finite (carrier G)\<rbrakk> |
335 |
\<Longrightarrow> \<exists>H. subgroup H G \<and> card H = p^a" |
|
336 |
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
|
337 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
338 |
end |