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