author | wenzelm |
Sat, 21 Jan 2006 23:07:26 +0100 | |
changeset 18738 | b6925d782fae |
parent 16663 | 13e9c402308b |
child 20318 | 0e0ea63fe768 |
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 |
ID: $Id$ |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
3 |
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
|
4 |
*) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
5 |
|
14706 | 6 |
header {* Sylow's theorem *} |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
7 |
|
16417 | 8 |
theory Sylow imports Coset begin |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
9 |
|
14706 | 10 |
text {* |
11 |
See also \cite{Kammueller-Paulson:1999}. |
|
12 |
*} |
|
13 |
||
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
14 |
text{*The combinatorial argument is in theory Exponent*} |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
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 |
16663 | 18 |
assumes prime_p: "prime p" |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
19 |
and order_G: "order(G) = (p^a) * m" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
20 |
and finite_G [iff]: "finite (carrier G)" |
14747 | 21 |
defines "calM == {s. s \<subseteq> carrier(G) & card(s) = p^a}" |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
22 |
and "RelM == {(N1,N2). N1 \<in> calM & N2 \<in> calM & |
14666 | 23 |
(\<exists>g \<in> carrier(G). N1 = (N2 #> g) )}" |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
24 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
25 |
lemma (in sylow) RelM_refl: "refl calM RelM" |
14666 | 26 |
apply (auto simp add: refl_def RelM_def calM_def) |
27 |
apply (blast intro!: coset_mult_one [symmetric]) |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
28 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
29 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
30 |
lemma (in sylow) RelM_sym: "sym RelM" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
31 |
proof (unfold sym_def RelM_def, clarify) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
32 |
fix y g |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
33 |
assume "y \<in> calM" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
34 |
and g: "g \<in> carrier G" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
35 |
hence "y = y #> g #> (inv g)" by (simp add: coset_mult_assoc calM_def) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
36 |
thus "\<exists>g'\<in>carrier G. y = y #> g #> g'" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
37 |
by (blast intro: g inv_closed) |
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 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
40 |
lemma (in sylow) RelM_trans: "trans RelM" |
14666 | 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 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
43 |
lemma (in sylow) RelM_equiv: "equiv calM RelM" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
44 |
apply (unfold equiv_def) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
45 |
apply (blast intro: RelM_refl RelM_sym RelM_trans) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
46 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
47 |
|
14747 | 48 |
lemma (in sylow) M_subset_calM_prep: "M' \<in> calM // RelM ==> M' \<subseteq> calM" |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
49 |
apply (unfold RelM_def) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
50 |
apply (blast elim!: quotientE) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
51 |
done |
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 |
subsection{*Main Part of the Proof*} |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
54 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
55 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
56 |
locale sylow_central = sylow + |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
57 |
fixes H and M1 and M |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
58 |
assumes M_in_quot: "M \<in> calM // RelM" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
59 |
and not_dvd_M: "~(p ^ Suc(exponent p m) dvd card(M))" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
60 |
and M1_in_M: "M1 \<in> M" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
61 |
defines "H == {g. g\<in>carrier G & M1 #> g = M1}" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
62 |
|
14747 | 63 |
lemma (in sylow_central) M_subset_calM: "M \<subseteq> calM" |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
64 |
by (rule M_in_quot [THEN M_subset_calM_prep]) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
65 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
66 |
lemma (in sylow_central) card_M1: "card(M1) = p^a" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
67 |
apply (cut_tac M_subset_calM M1_in_M) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
68 |
apply (simp add: calM_def, blast) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
69 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
70 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
71 |
lemma card_nonempty: "0 < card(S) ==> S \<noteq> {}" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
72 |
by force |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
73 |
|
14666 | 74 |
lemma (in sylow_central) exists_x_in_M1: "\<exists>x. x\<in>M1" |
75 |
apply (subgoal_tac "0 < card M1") |
|
76 |
apply (blast dest: card_nonempty) |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
77 |
apply (cut_tac prime_p [THEN prime_imp_one_less]) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
78 |
apply (simp (no_asm_simp) add: card_M1) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
79 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
80 |
|
14747 | 81 |
lemma (in sylow_central) M1_subset_G [simp]: "M1 \<subseteq> carrier G" |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
82 |
apply (rule subsetD [THEN PowD]) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
83 |
apply (rule_tac [2] M1_in_M) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
84 |
apply (rule M_subset_calM [THEN subset_trans]) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
85 |
apply (auto simp add: calM_def) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
86 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
87 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
88 |
lemma (in sylow_central) M1_inj_H: "\<exists>f \<in> H\<rightarrow>M1. inj_on f H" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
89 |
proof - |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
90 |
from exists_x_in_M1 obtain m1 where m1M: "m1 \<in> M1".. |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
91 |
have m1G: "m1 \<in> carrier G" by (simp add: m1M M1_subset_G [THEN subsetD]) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
92 |
show ?thesis |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
93 |
proof |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
94 |
show "inj_on (\<lambda>z\<in>H. m1 \<otimes> z) H" |
14666 | 95 |
by (simp add: inj_on_def l_cancel [of m1 x y, THEN iffD1] H_def m1G) |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
96 |
show "restrict (op \<otimes> m1) H \<in> H \<rightarrow> M1" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
97 |
proof (rule restrictI) |
14666 | 98 |
fix z assume zH: "z \<in> H" |
99 |
show "m1 \<otimes> z \<in> M1" |
|
100 |
proof - |
|
101 |
from zH |
|
102 |
have zG: "z \<in> carrier G" and M1zeq: "M1 #> z = M1" |
|
103 |
by (auto simp add: H_def) |
|
104 |
show ?thesis |
|
105 |
by (rule subst [OF M1zeq], simp add: m1M zG rcosI) |
|
106 |
qed |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
107 |
qed |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
108 |
qed |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
109 |
qed |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
110 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
111 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
112 |
subsection{*Discharging the Assumptions of @{text sylow_central}*} |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
113 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
114 |
lemma (in sylow) EmptyNotInEquivSet: "{} \<notin> calM // RelM" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
115 |
by (blast elim!: quotientE dest: RelM_equiv [THEN equiv_class_self]) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
116 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
117 |
lemma (in sylow) existsM1inM: "M \<in> calM // RelM ==> \<exists>M1. M1 \<in> M" |
14666 | 118 |
apply (subgoal_tac "M \<noteq> {}") |
119 |
apply blast |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
120 |
apply (cut_tac EmptyNotInEquivSet, blast) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
121 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
122 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
123 |
lemma (in sylow) zero_less_o_G: "0 < order(G)" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
124 |
apply (unfold order_def) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
125 |
apply (blast intro: one_closed zero_less_card_empty) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
126 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
127 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
128 |
lemma (in sylow) zero_less_m: "0 < m" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
129 |
apply (cut_tac zero_less_o_G) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
130 |
apply (simp add: order_G) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
131 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
132 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
133 |
lemma (in sylow) card_calM: "card(calM) = (p^a) * m choose p^a" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
134 |
by (simp add: calM_def n_subsets order_G [symmetric] order_def) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
135 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
136 |
lemma (in sylow) zero_less_card_calM: "0 < card calM" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
137 |
by (simp add: card_calM zero_less_binomial le_extend_mult zero_less_m) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
138 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
139 |
lemma (in sylow) max_p_div_calM: |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
140 |
"~ (p ^ Suc(exponent p m) dvd card(calM))" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
141 |
apply (subgoal_tac "exponent p m = exponent p (card calM) ") |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
142 |
apply (cut_tac zero_less_card_calM prime_p) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
143 |
apply (force dest: power_Suc_exponent_Not_dvd) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
144 |
apply (simp add: card_calM zero_less_m [THEN const_p_fac]) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
145 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
146 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
147 |
lemma (in sylow) finite_calM: "finite calM" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
148 |
apply (unfold calM_def) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
149 |
apply (rule_tac B = "Pow (carrier G) " in finite_subset) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
150 |
apply auto |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
151 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
152 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
153 |
lemma (in sylow) lemma_A1: |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
154 |
"\<exists>M \<in> calM // RelM. ~ (p ^ Suc(exponent p m) dvd card(M))" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
155 |
apply (rule max_p_div_calM [THEN contrapos_np]) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
156 |
apply (simp add: finite_calM equiv_imp_dvd_card [OF _ RelM_equiv]) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
157 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
158 |
|
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 |
subsubsection{*Introduction and Destruct Rules for @{term H}*} |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
161 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
162 |
lemma (in sylow_central) H_I: "[|g \<in> carrier G; M1 #> g = M1|] ==> g \<in> H" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
163 |
by (simp add: H_def) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
164 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
165 |
lemma (in sylow_central) H_into_carrier_G: "x \<in> H ==> x \<in> carrier G" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
166 |
by (simp add: H_def) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
167 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
168 |
lemma (in sylow_central) in_H_imp_eq: "g : H ==> M1 #> g = M1" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
169 |
by (simp add: H_def) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
170 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
171 |
lemma (in sylow_central) H_m_closed: "[| x\<in>H; y\<in>H|] ==> x \<otimes> y \<in> H" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
172 |
apply (unfold H_def) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
173 |
apply (simp add: coset_mult_assoc [symmetric] m_closed) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
174 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
175 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
176 |
lemma (in sylow_central) H_not_empty: "H \<noteq> {}" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
177 |
apply (simp add: H_def) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
178 |
apply (rule exI [of _ \<one>], simp) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
179 |
done |
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 |
lemma (in sylow_central) H_is_subgroup: "subgroup H G" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
182 |
apply (rule subgroupI) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
183 |
apply (rule subsetI) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
184 |
apply (erule H_into_carrier_G) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
185 |
apply (rule H_not_empty) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
186 |
apply (simp add: H_def, clarify) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
187 |
apply (erule_tac P = "%z. ?lhs(z) = M1" in subst) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
188 |
apply (simp add: coset_mult_assoc ) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
189 |
apply (blast intro: H_m_closed) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
190 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
191 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
192 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
193 |
lemma (in sylow_central) rcosetGM1g_subset_G: |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
194 |
"[| g \<in> carrier G; x \<in> M1 #> g |] ==> x \<in> carrier G" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
195 |
by (blast intro: M1_subset_G [THEN r_coset_subset_G, THEN subsetD]) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
196 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
197 |
lemma (in sylow_central) finite_M1: "finite M1" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
198 |
by (rule finite_subset [OF M1_subset_G finite_G]) |
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 |
lemma (in sylow_central) finite_rcosetGM1g: "g\<in>carrier G ==> finite (M1 #> g)" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
201 |
apply (rule finite_subset) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
202 |
apply (rule subsetI) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
203 |
apply (erule rcosetGM1g_subset_G, assumption) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
204 |
apply (rule finite_G) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
205 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
206 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
207 |
lemma (in sylow_central) M1_cardeq_rcosetGM1g: |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
208 |
"g \<in> carrier G ==> card(M1 #> g) = card(M1)" |
14963 | 209 |
by (simp (no_asm_simp) add: M1_subset_G card_cosets_equal rcosetsI) |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
210 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
211 |
lemma (in sylow_central) M1_RelM_rcosetGM1g: |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
212 |
"g \<in> carrier G ==> (M1, M1 #> g) \<in> RelM" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
213 |
apply (simp (no_asm) add: RelM_def calM_def card_M1 M1_subset_G) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
214 |
apply (rule conjI) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
215 |
apply (blast intro: rcosetGM1g_subset_G) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
216 |
apply (simp (no_asm_simp) add: card_M1 M1_cardeq_rcosetGM1g) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
217 |
apply (rule bexI [of _ "inv g"]) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
218 |
apply (simp_all add: coset_mult_assoc M1_subset_G) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
219 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
220 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
221 |
|
14963 | 222 |
subsection{*Equal Cardinalities of @{term M} and the Set of Cosets*} |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
223 |
|
14963 | 224 |
text{*Injections between @{term M} and @{term "rcosets\<^bsub>G\<^esub> H"} show that |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
225 |
their cardinalities are equal.*} |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
226 |
|
14666 | 227 |
lemma ElemClassEquiv: |
14963 | 228 |
"[| equiv A r; C \<in> A // r |] ==> \<forall>x \<in> C. \<forall>y \<in> C. (x,y)\<in>r" |
229 |
by (unfold equiv_def quotient_def sym_def trans_def, blast) |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
230 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
231 |
lemma (in sylow_central) M_elem_map: |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
232 |
"M2 \<in> M ==> \<exists>g. g \<in> carrier G & M1 #> g = M2" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
233 |
apply (cut_tac M1_in_M M_in_quot [THEN RelM_equiv [THEN ElemClassEquiv]]) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
234 |
apply (simp add: RelM_def) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
235 |
apply (blast dest!: bspec) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
236 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
237 |
|
14666 | 238 |
lemmas (in sylow_central) M_elem_map_carrier = |
239 |
M_elem_map [THEN someI_ex, THEN conjunct1] |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
240 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
241 |
lemmas (in sylow_central) M_elem_map_eq = |
14666 | 242 |
M_elem_map [THEN someI_ex, THEN conjunct2] |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
243 |
|
14963 | 244 |
lemma (in sylow_central) M_funcset_rcosets_H: |
245 |
"(%x:M. H #> (SOME g. g \<in> carrier G & M1 #> g = x)) \<in> M \<rightarrow> rcosets H" |
|
246 |
apply (rule rcosetsI [THEN restrictI]) |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
247 |
apply (rule H_is_subgroup [THEN subgroup.subset]) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
248 |
apply (erule M_elem_map_carrier) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
249 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
250 |
|
14963 | 251 |
lemma (in sylow_central) inj_M_GmodH: "\<exists>f \<in> M\<rightarrow>rcosets H. inj_on f M" |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
252 |
apply (rule bexI) |
14963 | 253 |
apply (rule_tac [2] M_funcset_rcosets_H) |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
254 |
apply (rule inj_onI, simp) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
255 |
apply (rule trans [OF _ M_elem_map_eq]) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
256 |
prefer 2 apply assumption |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
257 |
apply (rule M_elem_map_eq [symmetric, THEN trans], assumption) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
258 |
apply (rule coset_mult_inv1) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
259 |
apply (erule_tac [2] M_elem_map_carrier)+ |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
260 |
apply (rule_tac [2] M1_subset_G) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
261 |
apply (rule coset_join1 [THEN in_H_imp_eq]) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
262 |
apply (rule_tac [3] H_is_subgroup) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
263 |
prefer 2 apply (blast intro: m_closed M_elem_map_carrier inv_closed) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
264 |
apply (simp add: coset_mult_inv2 H_def M_elem_map_carrier subset_def) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
265 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
266 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
267 |
|
14963 | 268 |
subsubsection{*The opposite injection*} |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
269 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
270 |
lemma (in sylow_central) H_elem_map: |
14963 | 271 |
"H1 \<in> rcosets H ==> \<exists>g. g \<in> carrier G & H #> g = H1" |
272 |
by (auto simp add: RCOSETS_def) |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
273 |
|
14666 | 274 |
lemmas (in sylow_central) H_elem_map_carrier = |
275 |
H_elem_map [THEN someI_ex, THEN conjunct1] |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
276 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
277 |
lemmas (in sylow_central) H_elem_map_eq = |
14666 | 278 |
H_elem_map [THEN someI_ex, THEN conjunct2] |
13870
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 |
|
14666 | 281 |
lemma EquivElemClass: |
14963 | 282 |
"[|equiv A r; M \<in> A//r; M1\<in>M; (M1,M2) \<in> r |] ==> M2 \<in> M" |
283 |
by (unfold equiv_def quotient_def sym_def trans_def, blast) |
|
284 |
||
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
285 |
|
14963 | 286 |
lemma (in sylow_central) rcosets_H_funcset_M: |
287 |
"(\<lambda>C \<in> rcosets H. M1 #> (@g. g \<in> carrier G \<and> H #> g = C)) \<in> rcosets H \<rightarrow> M" |
|
288 |
apply (simp add: RCOSETS_def) |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
289 |
apply (fast intro: someI2 |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
290 |
intro!: restrictI M1_in_M |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
291 |
EquivElemClass [OF RelM_equiv M_in_quot _ M1_RelM_rcosetGM1g]) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
292 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
293 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
294 |
text{*close to a duplicate of @{text inj_M_GmodH}*} |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
295 |
lemma (in sylow_central) inj_GmodH_M: |
14963 | 296 |
"\<exists>g \<in> rcosets H\<rightarrow>M. inj_on g (rcosets H)" |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
297 |
apply (rule bexI) |
14963 | 298 |
apply (rule_tac [2] rcosets_H_funcset_M) |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
299 |
apply (rule inj_onI) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
300 |
apply (simp) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
301 |
apply (rule trans [OF _ H_elem_map_eq]) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
302 |
prefer 2 apply assumption |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
303 |
apply (rule H_elem_map_eq [symmetric, THEN trans], assumption) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
304 |
apply (rule coset_mult_inv1) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
305 |
apply (erule_tac [2] H_elem_map_carrier)+ |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
306 |
apply (rule_tac [2] H_is_subgroup [THEN subgroup.subset]) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
307 |
apply (rule coset_join2) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
308 |
apply (blast intro: m_closed inv_closed H_elem_map_carrier) |
14666 | 309 |
apply (rule H_is_subgroup) |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
310 |
apply (simp add: H_I coset_mult_inv2 M1_subset_G H_elem_map_carrier) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
311 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
312 |
|
14747 | 313 |
lemma (in sylow_central) calM_subset_PowG: "calM \<subseteq> Pow(carrier G)" |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
314 |
by (auto simp add: calM_def) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
315 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
316 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
317 |
lemma (in sylow_central) finite_M: "finite M" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
318 |
apply (rule finite_subset) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
319 |
apply (rule M_subset_calM [THEN subset_trans]) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
320 |
apply (rule calM_subset_PowG, blast) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
321 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
322 |
|
14963 | 323 |
lemma (in sylow_central) cardMeqIndexH: "card(M) = card(rcosets H)" |
14666 | 324 |
apply (insert inj_M_GmodH inj_GmodH_M) |
325 |
apply (blast intro: card_bij finite_M H_is_subgroup |
|
14963 | 326 |
rcosets_subset_PowG [THEN finite_subset] |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
327 |
finite_Pow_iff [THEN iffD2]) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
328 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
329 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
330 |
lemma (in sylow_central) index_lem: "card(M) * card(H) = order(G)" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
331 |
by (simp add: cardMeqIndexH lagrange H_is_subgroup) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
332 |
|
14747 | 333 |
lemma (in sylow_central) lemma_leq1: "p^a \<le> card(H)" |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
334 |
apply (rule dvd_imp_le) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
335 |
apply (rule div_combine [OF prime_p not_dvd_M]) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
336 |
prefer 2 apply (blast intro: subgroup.finite_imp_card_positive H_is_subgroup) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
337 |
apply (simp add: index_lem order_G power_add mult_dvd_mono power_exponent_dvd |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
338 |
zero_less_m) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
339 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
340 |
|
14747 | 341 |
lemma (in sylow_central) lemma_leq2: "card(H) \<le> p^a" |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
342 |
apply (subst card_M1 [symmetric]) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
343 |
apply (cut_tac M1_inj_H) |
14666 | 344 |
apply (blast intro!: M1_subset_G intro: |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
345 |
card_inj H_into_carrier_G finite_subset [OF _ finite_G]) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
346 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
347 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
348 |
lemma (in sylow_central) card_H_eq: "card(H) = p^a" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
349 |
by (blast intro: le_anti_sym lemma_leq1 lemma_leq2) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
350 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
351 |
lemma (in sylow) sylow_thm: "\<exists>H. subgroup H G & card(H) = p^a" |
14666 | 352 |
apply (cut_tac lemma_A1, clarify) |
353 |
apply (frule existsM1inM, clarify) |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
354 |
apply (subgoal_tac "sylow_central G p a m M1 M") |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
355 |
apply (blast dest: sylow_central.H_is_subgroup sylow_central.card_H_eq) |
14666 | 356 |
apply (simp add: sylow_central_def sylow_central_axioms_def prems) |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
357 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
358 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
359 |
text{*Needed because the locale's automatic definition refers to |
14666 | 360 |
@{term "semigroup G"} and @{term "group_axioms G"} rather than |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
361 |
simply to @{term "group G"}.*} |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
362 |
lemma sylow_eq: "sylow G p a m = (group G & sylow_axioms G p a m)" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
363 |
by (simp add: sylow_def group_def) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
364 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
365 |
theorem sylow_thm: |
16663 | 366 |
"[| prime p; group(G); order(G) = (p^a) * m; finite (carrier G)|] |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
367 |
==> \<exists>H. subgroup H G & card(H) = p^a" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
368 |
apply (rule sylow.sylow_thm [of G p a m]) |
14666 | 369 |
apply (simp add: sylow_eq sylow_axioms_def) |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
370 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
371 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
372 |
end |
14963 | 373 |