author | wenzelm |
Tue, 17 Jan 2017 14:56:15 +0100 | |
changeset 64912 | 68f0465d956b |
parent 63537 | 831816778409 |
child 64914 | 51f015bd4565 |
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:
61382
diff
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:
61382
diff
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:
61382
diff
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:
61382
diff
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:
61382
diff
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:
61382
diff
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:
61382
diff
changeset
|
49 |
end |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
16663
diff
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:
61382
diff
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:
61382
diff
changeset
|
61 |
lemma M_subset_calM: "M \<subseteq> calM" |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
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:
61382
diff
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:
61382
diff
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:
61382
diff
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:
61382
diff
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:
61382
diff
changeset
|
75 |
proof - |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
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:
61382
diff
changeset
|
79 |
show ?thesis |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
80 |
proof |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
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) |
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
83 |
show "restrict (op \<otimes> m1) H \<in> H \<rightarrow> M1" |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
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:
61382
diff
changeset
|
87 |
show "m1 \<otimes> z \<in> M1" |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
88 |
proof - |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
89 |
from zH |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
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:
61382
diff
changeset
|
91 |
by (auto simp add: H_def) |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
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:
61382
diff
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:
61382
diff
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:
61382
diff
changeset
|
104 |
context sylow |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
105 |
begin |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
106 |
|
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
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:
61382
diff
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:
61382
diff
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:
61382
diff
changeset
|
116 |
lemma zero_less_m: "m > 0" |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
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:
61382
diff
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:
63167
diff
changeset
|
126 |
proof |
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
63167
diff
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:
63167
diff
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:
63167
diff
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:
63167
diff
changeset
|
132 |
also have "multiplicity p m = multiplicity p (card calM)" |
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
63167
diff
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:
63167
diff
changeset
|
134 |
finally show False by simp |
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
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:
61382
diff
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:
61382
diff
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:
61382
diff
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 |
|
64912 | 148 |
lemma (in sylow_central) H_I: "\<lbrakk>g \<in> carrier G; M1 #> g = M1\<rbrakk> \<Longrightarrow> g \<in> H" |
149 |
by (simp add: H_def) |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
150 |
|
64912 | 151 |
lemma (in sylow_central) H_into_carrier_G: "x \<in> H \<Longrightarrow> x \<in> carrier G" |
152 |
by (simp add: H_def) |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
153 |
|
64912 | 154 |
lemma (in sylow_central) in_H_imp_eq: "g \<in> H \<Longrightarrow> M1 #> g = M1" |
155 |
by (simp add: H_def) |
|
156 |
||
157 |
lemma (in sylow_central) H_m_closed: "\<lbrakk>x \<in> H; y \<in> H\<rbrakk> \<Longrightarrow> x \<otimes> y \<in> H" |
|
158 |
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
|
159 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
160 |
lemma (in sylow_central) H_not_empty: "H \<noteq> {}" |
64912 | 161 |
apply (simp add: H_def) |
162 |
apply (rule exI [of _ \<one>]) |
|
163 |
apply simp |
|
164 |
done |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
165 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
166 |
lemma (in sylow_central) H_is_subgroup: "subgroup H G" |
64912 | 167 |
apply (rule subgroupI) |
168 |
apply (rule subsetI) |
|
169 |
apply (erule H_into_carrier_G) |
|
170 |
apply (rule H_not_empty) |
|
171 |
apply (simp add: H_def) |
|
172 |
apply clarify |
|
173 |
apply (erule_tac P = "\<lambda>z. lhs z = M1" for lhs in subst) |
|
174 |
apply (simp add: coset_mult_assoc ) |
|
175 |
apply (blast intro: H_m_closed) |
|
176 |
done |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
177 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
178 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
179 |
lemma (in sylow_central) rcosetGM1g_subset_G: |
64912 | 180 |
"\<lbrakk>g \<in> carrier G; x \<in> M1 #> g\<rbrakk> \<Longrightarrow> x \<in> carrier G" |
181 |
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
|
182 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
183 |
lemma (in sylow_central) finite_M1: "finite M1" |
64912 | 184 |
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
|
185 |
|
64912 | 186 |
lemma (in sylow_central) 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
|
187 |
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
|
188 |
|
64912 | 189 |
lemma (in sylow_central) M1_cardeq_rcosetGM1g: "g \<in> carrier G \<Longrightarrow> card (M1 #> g) = card M1" |
190 |
by (simp add: card_cosets_equal rcosetsI) |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
191 |
|
64912 | 192 |
lemma (in sylow_central) M1_RelM_rcosetGM1g: "g \<in> carrier G \<Longrightarrow> (M1, M1 #> g) \<in> RelM" |
193 |
apply (simp add: RelM_def calM_def card_M1) |
|
194 |
apply (rule conjI) |
|
195 |
apply (blast intro: rcosetGM1g_subset_G) |
|
196 |
apply (simp add: card_M1 M1_cardeq_rcosetGM1g) |
|
197 |
apply (metis M1_subset_G coset_mult_assoc coset_mult_one r_inv_ex) |
|
198 |
done |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
199 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
200 |
|
64912 | 201 |
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
|
202 |
|
64912 | 203 |
text \<open>Injections between @{term M} and @{term "rcosets\<^bsub>G\<^esub> H"} show that |
61382 | 204 |
their cardinalities are equal.\<close> |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
205 |
|
64912 | 206 |
lemma ElemClassEquiv: "\<lbrakk>equiv A r; C \<in> A // r\<rbrakk> \<Longrightarrow> \<forall>x \<in> C. \<forall>y \<in> C. (x, y) \<in> r" |
207 |
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
|
208 |
|
64912 | 209 |
lemma (in sylow_central) M_elem_map: "M2 \<in> M \<Longrightarrow> \<exists>g. g \<in> carrier G \<and> M1 #> g = M2" |
210 |
using M1_in_M M_in_quot [THEN RelM_equiv [THEN ElemClassEquiv]] |
|
211 |
by (simp add: RelM_def) (blast dest!: bspec) |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
212 |
|
14666 | 213 |
lemmas (in sylow_central) M_elem_map_carrier = |
64912 | 214 |
M_elem_map [THEN someI_ex, THEN conjunct1] |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
215 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
216 |
lemmas (in sylow_central) M_elem_map_eq = |
64912 | 217 |
M_elem_map [THEN someI_ex, THEN conjunct2] |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
218 |
|
14963 | 219 |
lemma (in sylow_central) M_funcset_rcosets_H: |
64912 | 220 |
"(\<lambda>x\<in>M. H #> (SOME g. g \<in> carrier G \<and> M1 #> g = x)) \<in> M \<rightarrow> rcosets H" |
55157 | 221 |
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
|
222 |
|
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
223 |
lemma (in sylow_central) inj_M_GmodH: "\<exists>f \<in> M \<rightarrow> rcosets H. inj_on f M" |
64912 | 224 |
apply (rule bexI) |
225 |
apply (rule_tac [2] M_funcset_rcosets_H) |
|
226 |
apply (rule inj_onI, simp) |
|
227 |
apply (rule trans [OF _ M_elem_map_eq]) |
|
228 |
prefer 2 apply assumption |
|
229 |
apply (rule M_elem_map_eq [symmetric, THEN trans], assumption) |
|
230 |
apply (rule coset_mult_inv1) |
|
231 |
apply (erule_tac [2] M_elem_map_carrier)+ |
|
232 |
apply (rule_tac [2] M1_subset_G) |
|
233 |
apply (rule coset_join1 [THEN in_H_imp_eq]) |
|
234 |
apply (rule_tac [3] H_is_subgroup) |
|
235 |
prefer 2 apply (blast intro: M_elem_map_carrier) |
|
236 |
apply (simp add: coset_mult_inv2 H_def M_elem_map_carrier subset_eq) |
|
237 |
done |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
238 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
239 |
|
61382 | 240 |
subsubsection\<open>The Opposite Injection\<close> |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
241 |
|
64912 | 242 |
lemma (in sylow_central) H_elem_map: "H1 \<in> rcosets H \<Longrightarrow> \<exists>g. g \<in> carrier G \<and> H #> g = H1" |
243 |
by (auto simp: RCOSETS_def) |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
244 |
|
14666 | 245 |
lemmas (in sylow_central) H_elem_map_carrier = |
64912 | 246 |
H_elem_map [THEN someI_ex, THEN conjunct1] |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
247 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
248 |
lemmas (in sylow_central) H_elem_map_eq = |
64912 | 249 |
H_elem_map [THEN someI_ex, THEN conjunct2] |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
250 |
|
14963 | 251 |
lemma (in sylow_central) rcosets_H_funcset_M: |
252 |
"(\<lambda>C \<in> rcosets H. M1 #> (@g. g \<in> carrier G \<and> H #> g = C)) \<in> rcosets H \<rightarrow> M" |
|
64912 | 253 |
apply (simp add: RCOSETS_def) |
254 |
apply (fast intro: someI2 |
|
255 |
intro!: M1_in_M in_quotient_imp_closed [OF RelM_equiv M_in_quot _ M1_RelM_rcosetGM1g]) |
|
256 |
done |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
257 |
|
64912 | 258 |
text \<open>Close to a duplicate of \<open>inj_M_GmodH\<close>.\<close> |
259 |
lemma (in sylow_central) inj_GmodH_M: "\<exists>g \<in> rcosets H\<rightarrow>M. inj_on g (rcosets H)" |
|
260 |
apply (rule bexI) |
|
261 |
apply (rule_tac [2] rcosets_H_funcset_M) |
|
262 |
apply (rule inj_onI) |
|
263 |
apply (simp) |
|
264 |
apply (rule trans [OF _ H_elem_map_eq]) |
|
265 |
prefer 2 apply assumption |
|
266 |
apply (rule H_elem_map_eq [symmetric, THEN trans], assumption) |
|
267 |
apply (rule coset_mult_inv1) |
|
268 |
apply (erule_tac [2] H_elem_map_carrier)+ |
|
269 |
apply (rule_tac [2] H_is_subgroup [THEN subgroup.subset]) |
|
270 |
apply (rule coset_join2) |
|
271 |
apply (blast intro: H_elem_map_carrier) |
|
272 |
apply (rule H_is_subgroup) |
|
273 |
apply (simp add: H_I coset_mult_inv2 H_elem_map_carrier) |
|
274 |
done |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
275 |
|
64912 | 276 |
lemma (in sylow_central) calM_subset_PowG: "calM \<subseteq> Pow (carrier G)" |
277 |
by (auto simp: calM_def) |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
278 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
279 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
280 |
lemma (in sylow_central) finite_M: "finite M" |
64912 | 281 |
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
|
282 |
|
64912 | 283 |
lemma (in sylow_central) cardMeqIndexH: "card M = card (rcosets H)" |
284 |
apply (insert inj_M_GmodH inj_GmodH_M) |
|
285 |
apply (blast intro: card_bij finite_M H_is_subgroup |
|
286 |
rcosets_subset_PowG [THEN finite_subset] |
|
287 |
finite_Pow_iff [THEN iffD2]) |
|
288 |
done |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
289 |
|
64912 | 290 |
lemma (in sylow_central) index_lem: "card M * card H = order G" |
291 |
by (simp add: cardMeqIndexH lagrange H_is_subgroup) |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
292 |
|
64912 | 293 |
lemma (in sylow_central) lemma_leq1: "p^a \<le> card H" |
294 |
apply (rule dvd_imp_le) |
|
295 |
apply (rule div_combine [OF prime_imp_prime_elem[OF prime_p] not_dvd_M]) |
|
296 |
prefer 2 apply (blast intro: subgroup.finite_imp_card_positive H_is_subgroup) |
|
297 |
apply (simp add: index_lem order_G power_add mult_dvd_mono multiplicity_dvd zero_less_m) |
|
298 |
done |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
299 |
|
64912 | 300 |
lemma (in sylow_central) lemma_leq2: "card H \<le> p^a" |
301 |
apply (subst card_M1 [symmetric]) |
|
302 |
apply (cut_tac M1_inj_H) |
|
303 |
apply (blast intro!: M1_subset_G intro: card_inj H_into_carrier_G finite_subset [OF _ finite_G]) |
|
304 |
done |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
305 |
|
64912 | 306 |
lemma (in sylow_central) card_H_eq: "card H = p^a" |
307 |
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
|
308 |
|
64912 | 309 |
lemma (in sylow) sylow_thm: "\<exists>H. subgroup H G \<and> card H = p^a" |
310 |
using lemma_A1 |
|
311 |
apply clarify |
|
312 |
apply (frule existsM1inM, clarify) |
|
313 |
apply (subgoal_tac "sylow_central G p a m M1 M") |
|
314 |
apply (blast dest: sylow_central.H_is_subgroup sylow_central.card_H_eq) |
|
315 |
apply (simp add: sylow_central_def sylow_central_axioms_def sylow_axioms calM_def RelM_def) |
|
316 |
done |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
317 |
|
64912 | 318 |
text \<open>Needed because the locale's automatic definition refers to |
319 |
@{term "semigroup G"} and @{term "group_axioms G"} rather than |
|
61382 | 320 |
simply to @{term "group G"}.\<close> |
64912 | 321 |
lemma sylow_eq: "sylow G p a m \<longleftrightarrow> group G \<and> sylow_axioms G p a m" |
322 |
by (simp add: sylow_def group_def) |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
323 |
|
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
16663
diff
changeset
|
324 |
|
61382 | 325 |
subsection \<open>Sylow's Theorem\<close> |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
16663
diff
changeset
|
326 |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
327 |
theorem sylow_thm: |
64912 | 328 |
"\<lbrakk>prime p; group G; order G = (p^a) * m; finite (carrier G)\<rbrakk> |
329 |
\<Longrightarrow> \<exists>H. subgroup H G \<and> card H = p^a" |
|
330 |
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
|
331 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
332 |
end |