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