| author | paulson | 
| Fri, 14 May 2004 16:52:53 +0200 | |
| changeset 14748 | 001323d6d75b | 
| parent 14747 | 2eaff69d3d10 | 
| child 14803 | f7557773cc87 | 
| 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 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 8 | theory Sylow = Coset: | 
| 
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 | subsection {*Order of a Group and Lagrange's Theorem*}
 | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 15 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 16 | constdefs | 
| 14666 | 17 |   order :: "('a, 'b) semigroup_scheme => nat"
 | 
| 14651 | 18 | "order S == card (carrier S)" | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 19 | |
| 14747 | 20 | theorem (in group) lagrange: | 
| 14666 | 21 | "[| finite(carrier G); subgroup H G |] | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 22 | ==> card(rcosets G H) * card(H) = order(G)" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 23 | apply (simp (no_asm_simp) add: order_def setrcos_part_G [symmetric]) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 24 | apply (subst mult_commute) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 25 | apply (rule card_partition) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 26 | apply (simp add: setrcos_subset_PowG [THEN finite_subset]) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 27 | apply (simp add: setrcos_part_G) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 28 | apply (simp add: card_cosets_equal subgroup.subset) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 29 | apply (simp add: rcos_disjoint) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 30 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 31 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 32 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 33 | text{*The combinatorial argument is in theory Exponent*}
 | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 34 | |
| 14747 | 35 | locale sylow = group + | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 36 | fixes p and a and m and calM and RelM | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 37 | assumes prime_p: "p \<in> prime" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 38 | and order_G: "order(G) = (p^a) * m" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 39 | and finite_G [iff]: "finite (carrier G)" | 
| 14747 | 40 |   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 | 41 |       and "RelM == {(N1,N2). N1 \<in> calM & N2 \<in> calM &
 | 
| 14666 | 42 | (\<exists>g \<in> carrier(G). N1 = (N2 #> g) )}" | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 43 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 44 | lemma (in sylow) RelM_refl: "refl calM RelM" | 
| 14666 | 45 | apply (auto simp add: refl_def RelM_def calM_def) | 
| 46 | apply (blast intro!: coset_mult_one [symmetric]) | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 47 | done | 
| 
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 | lemma (in sylow) RelM_sym: "sym RelM" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 50 | proof (unfold sym_def RelM_def, clarify) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 51 | fix y g | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 52 | assume "y \<in> calM" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 53 | and g: "g \<in> carrier G" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 54 | 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 | 55 | 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 | 56 | by (blast intro: g inv_closed) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 57 | qed | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 58 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 59 | lemma (in sylow) RelM_trans: "trans RelM" | 
| 14666 | 60 | 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 | 61 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 62 | lemma (in sylow) RelM_equiv: "equiv calM RelM" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 63 | apply (unfold equiv_def) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 64 | apply (blast intro: RelM_refl RelM_sym RelM_trans) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 65 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 66 | |
| 14747 | 67 | 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 | 68 | apply (unfold RelM_def) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 69 | apply (blast elim!: quotientE) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 70 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 71 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 72 | subsection{*Main Part of the Proof*}
 | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 73 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 74 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 75 | locale sylow_central = sylow + | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 76 | fixes H and M1 and M | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 77 | assumes M_in_quot: "M \<in> calM // RelM" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 78 | 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 | 79 | and M1_in_M: "M1 \<in> M" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 80 |   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 | 81 | |
| 14747 | 82 | 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 | 83 | 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 | 84 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 85 | 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 | 86 | apply (cut_tac M_subset_calM M1_in_M) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 87 | apply (simp add: calM_def, blast) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 88 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 89 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 90 | lemma card_nonempty: "0 < card(S) ==> S \<noteq> {}"
 | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 91 | by force | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 92 | |
| 14666 | 93 | lemma (in sylow_central) exists_x_in_M1: "\<exists>x. x\<in>M1" | 
| 94 | apply (subgoal_tac "0 < card M1") | |
| 95 | apply (blast dest: card_nonempty) | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 96 | 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 | 97 | apply (simp (no_asm_simp) add: card_M1) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 98 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 99 | |
| 14747 | 100 | 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 | 101 | apply (rule subsetD [THEN PowD]) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 102 | apply (rule_tac [2] M1_in_M) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 103 | apply (rule M_subset_calM [THEN subset_trans]) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 104 | apply (auto simp add: calM_def) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 105 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 106 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 107 | 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 | 108 | proof - | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 109 | 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 | 110 | 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 | 111 | show ?thesis | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 112 | proof | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 113 | show "inj_on (\<lambda>z\<in>H. m1 \<otimes> z) H" | 
| 14666 | 114 | 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 | 115 | 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 | 116 | proof (rule restrictI) | 
| 14666 | 117 | fix z assume zH: "z \<in> H" | 
| 118 | show "m1 \<otimes> z \<in> M1" | |
| 119 | proof - | |
| 120 | from zH | |
| 121 | have zG: "z \<in> carrier G" and M1zeq: "M1 #> z = M1" | |
| 122 | by (auto simp add: H_def) | |
| 123 | show ?thesis | |
| 124 | by (rule subst [OF M1zeq], simp add: m1M zG rcosI) | |
| 125 | qed | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 126 | qed | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 127 | qed | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 128 | qed | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 129 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 130 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 131 | subsection{*Discharging the Assumptions of @{text sylow_central}*}
 | 
| 
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) EmptyNotInEquivSet: "{} \<notin> calM // RelM"
 | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 134 | 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 | 135 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 136 | lemma (in sylow) existsM1inM: "M \<in> calM // RelM ==> \<exists>M1. M1 \<in> M" | 
| 14666 | 137 | apply (subgoal_tac "M \<noteq> {}")
 | 
| 138 | apply blast | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 139 | apply (cut_tac EmptyNotInEquivSet, blast) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 140 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 141 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 142 | 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 | 143 | apply (unfold order_def) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 144 | apply (blast intro: one_closed zero_less_card_empty) | 
| 
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) zero_less_m: "0 < m" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 148 | apply (cut_tac zero_less_o_G) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 149 | apply (simp add: order_G) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 150 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 151 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 152 | 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 | 153 | 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 | 154 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 155 | 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 | 156 | 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 | 157 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 158 | lemma (in sylow) max_p_div_calM: | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 159 | "~ (p ^ Suc(exponent p m) dvd card(calM))" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 160 | 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 | 161 | apply (cut_tac zero_less_card_calM prime_p) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 162 | apply (force dest: power_Suc_exponent_Not_dvd) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 163 | 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 | 164 | done | 
| 
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) finite_calM: "finite calM" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 167 | apply (unfold calM_def) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 168 | 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 | 169 | apply auto | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 170 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 171 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 172 | lemma (in sylow) lemma_A1: | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 173 | "\<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 | 174 | apply (rule max_p_div_calM [THEN contrapos_np]) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 175 | 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 | 176 | done | 
| 
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 | subsubsection{*Introduction and Destruct Rules for @{term H}*}
 | 
| 
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_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 | 182 | by (simp add: H_def) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 183 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 184 | 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 | 185 | by (simp add: H_def) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 186 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 187 | 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 | 188 | by (simp add: H_def) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 189 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 190 | 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 | 191 | apply (unfold H_def) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 192 | apply (simp add: coset_mult_assoc [symmetric] m_closed) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 193 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 194 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 195 | lemma (in sylow_central) H_not_empty: "H \<noteq> {}"
 | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 196 | apply (simp add: H_def) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 197 | apply (rule exI [of _ \<one>], simp) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 198 | done | 
| 
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) H_is_subgroup: "subgroup H G" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 201 | apply (rule subgroupI) | 
| 
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 H_into_carrier_G) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 204 | apply (rule H_not_empty) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 205 | apply (simp add: H_def, clarify) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 206 | 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 | 207 | apply (simp add: coset_mult_assoc ) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 208 | apply (blast intro: H_m_closed) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 209 | done | 
| 
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 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 212 | lemma (in sylow_central) rcosetGM1g_subset_G: | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 213 | "[| 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 | 214 | 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 | 215 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 216 | lemma (in sylow_central) finite_M1: "finite M1" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 217 | 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 | 218 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 219 | 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 | 220 | apply (rule finite_subset) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 221 | apply (rule subsetI) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 222 | apply (erule rcosetGM1g_subset_G, assumption) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 223 | apply (rule finite_G) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 224 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 225 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 226 | lemma (in sylow_central) M1_cardeq_rcosetGM1g: | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 227 | "g \<in> carrier G ==> card(M1 #> g) = card(M1)" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 228 | by (simp (no_asm_simp) add: M1_subset_G card_cosets_equal setrcosI) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 229 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 230 | lemma (in sylow_central) M1_RelM_rcosetGM1g: | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 231 | "g \<in> carrier G ==> (M1, M1 #> g) \<in> RelM" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 232 | 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 | 233 | apply (rule conjI) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 234 | apply (blast intro: rcosetGM1g_subset_G) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 235 | 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 | 236 | apply (rule bexI [of _ "inv g"]) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 237 | 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 | 238 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 239 | |
| 
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 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 242 | subsection{*Equal Cardinalities of @{term M} and @{term "rcosets G H"}*}
 | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 243 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 244 | text{*Injections between @{term M} and @{term "rcosets G H"} show that
 | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 245 | their cardinalities are equal.*} | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 246 | |
| 14666 | 247 | lemma ElemClassEquiv: | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 248 | "[| equiv A r; C\<in>A // r |] ==> \<forall>x \<in> C. \<forall>y \<in> C. (x,y)\<in>r" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 249 | apply (unfold equiv_def quotient_def sym_def trans_def, blast) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 250 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 251 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 252 | lemma (in sylow_central) M_elem_map: | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 253 | "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 | 254 | 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 | 255 | apply (simp add: RelM_def) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 256 | apply (blast dest!: bspec) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 257 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 258 | |
| 14666 | 259 | lemmas (in sylow_central) M_elem_map_carrier = | 
| 260 | M_elem_map [THEN someI_ex, THEN conjunct1] | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 261 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 262 | lemmas (in sylow_central) M_elem_map_eq = | 
| 14666 | 263 | M_elem_map [THEN someI_ex, THEN conjunct2] | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 264 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 265 | lemma (in sylow_central) M_funcset_setrcos_H: | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 266 | "(%x:M. H #> (SOME g. g \<in> carrier G & M1 #> g = x)) \<in> M \<rightarrow> rcosets G H" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 267 | apply (rule setrcosI [THEN restrictI]) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 268 | apply (rule H_is_subgroup [THEN subgroup.subset]) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 269 | apply (erule M_elem_map_carrier) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 270 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 271 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 272 | lemma (in sylow_central) inj_M_GmodH: "\<exists>f \<in> M\<rightarrow>rcosets G H. inj_on f M" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 273 | apply (rule bexI) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 274 | apply (rule_tac [2] M_funcset_setrcos_H) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 275 | apply (rule inj_onI, simp) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 276 | apply (rule trans [OF _ M_elem_map_eq]) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 277 | prefer 2 apply assumption | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 278 | 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 | 279 | apply (rule coset_mult_inv1) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 280 | apply (erule_tac [2] M_elem_map_carrier)+ | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 281 | apply (rule_tac [2] M1_subset_G) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 282 | apply (rule coset_join1 [THEN in_H_imp_eq]) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 283 | apply (rule_tac [3] H_is_subgroup) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 284 | 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 | 285 | 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 | 286 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 287 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 288 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 289 | (** the opposite injection **) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 290 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 291 | lemma (in sylow_central) H_elem_map: | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 292 | "H1\<in>rcosets G H ==> \<exists>g. g \<in> carrier G & H #> g = H1" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 293 | by (auto simp add: setrcos_eq) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 294 | |
| 14666 | 295 | lemmas (in sylow_central) H_elem_map_carrier = | 
| 296 | H_elem_map [THEN someI_ex, THEN conjunct1] | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 297 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 298 | lemmas (in sylow_central) H_elem_map_eq = | 
| 14666 | 299 | H_elem_map [THEN someI_ex, THEN conjunct2] | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 300 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 301 | |
| 14666 | 302 | lemma EquivElemClass: | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 303 | "[|equiv A r; M\<in>A // r; M1\<in>M; (M1, M2)\<in>r |] ==> M2\<in>M" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 304 | apply (unfold equiv_def quotient_def sym_def trans_def, blast) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 305 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 306 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 307 | lemma (in sylow_central) setrcos_H_funcset_M: | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 308 | "(\<lambda>C \<in> rcosets G H. M1 #> (@g. g \<in> carrier G \<and> H #> g = C)) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 309 | \<in> rcosets G H \<rightarrow> M" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 310 | apply (simp add: setrcos_eq) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 311 | apply (fast intro: someI2 | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 312 | intro!: restrictI M1_in_M | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 313 | 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 | 314 | done | 
| 
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 | 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 | 317 | lemma (in sylow_central) inj_GmodH_M: | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 318 | "\<exists>g \<in> rcosets G H\<rightarrow>M. inj_on g (rcosets G H)" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 319 | apply (rule bexI) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 320 | apply (rule_tac [2] setrcos_H_funcset_M) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 321 | apply (rule inj_onI) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 322 | apply (simp) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 323 | apply (rule trans [OF _ H_elem_map_eq]) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 324 | prefer 2 apply assumption | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 325 | 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 | 326 | apply (rule coset_mult_inv1) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 327 | apply (erule_tac [2] H_elem_map_carrier)+ | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 328 | 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 | 329 | apply (rule coset_join2) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 330 | apply (blast intro: m_closed inv_closed H_elem_map_carrier) | 
| 14666 | 331 | apply (rule H_is_subgroup) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 332 | 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 | 333 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 334 | |
| 14747 | 335 | 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 | 336 | by (auto simp add: calM_def) | 
| 
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 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 339 | lemma (in sylow_central) finite_M: "finite M" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 340 | apply (rule finite_subset) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 341 | apply (rule M_subset_calM [THEN subset_trans]) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 342 | apply (rule calM_subset_PowG, blast) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 343 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 344 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 345 | lemma (in sylow_central) cardMeqIndexH: "card(M) = card(rcosets G H)" | 
| 14666 | 346 | apply (insert inj_M_GmodH inj_GmodH_M) | 
| 347 | apply (blast intro: card_bij finite_M H_is_subgroup | |
| 348 | setrcos_subset_PowG [THEN finite_subset] | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 349 | finite_Pow_iff [THEN iffD2]) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 350 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 351 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 352 | 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 | 353 | by (simp add: cardMeqIndexH lagrange H_is_subgroup) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 354 | |
| 14747 | 355 | 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 | 356 | apply (rule dvd_imp_le) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 357 | 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 | 358 | 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 | 359 | 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 | 360 | zero_less_m) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 361 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 362 | |
| 14747 | 363 | 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 | 364 | apply (subst card_M1 [symmetric]) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 365 | apply (cut_tac M1_inj_H) | 
| 14666 | 366 | apply (blast intro!: M1_subset_G intro: | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 367 | 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 | 368 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 369 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 370 | 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 | 371 | 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 | 372 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 373 | lemma (in sylow) sylow_thm: "\<exists>H. subgroup H G & card(H) = p^a" | 
| 14666 | 374 | apply (cut_tac lemma_A1, clarify) | 
| 375 | apply (frule existsM1inM, clarify) | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 376 | 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 | 377 | apply (blast dest: sylow_central.H_is_subgroup sylow_central.card_H_eq) | 
| 14666 | 378 | 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 | 379 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 380 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 381 | text{*Needed because the locale's automatic definition refers to
 | 
| 14666 | 382 |    @{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 | 383 |   simply to @{term "group G"}.*}
 | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 384 | 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 | 385 | by (simp add: sylow_def group_def) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 386 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 387 | theorem sylow_thm: | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 388 | "[|p \<in> prime; group(G); order(G) = (p^a) * m; finite (carrier G)|] | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 389 | ==> \<exists>H. subgroup H G & card(H) = p^a" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 390 | apply (rule sylow.sylow_thm [of G p a m]) | 
| 14666 | 391 | apply (simp add: sylow_eq sylow_axioms_def) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 392 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 393 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 394 | end |