src/HOL/Algebra/Sylow.thy
author wenzelm
Sun, 11 Sep 2016 00:13:25 +0200
changeset 63832 a400b127853c
parent 63537 831816778409
child 64912 68f0465d956b
permissions -rw-r--r--
misc tuning and modernization;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14706
71590b7733b7 tuned document;
wenzelm
parents: 14666
diff changeset
     1
(*  Title:      HOL/Algebra/Sylow.thy
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
     2
    Author:     Florian Kammueller, with new proofs by L C Paulson
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
     3
*)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
     4
35849
b5522b51cb1e standard headers;
wenzelm
parents: 33657
diff changeset
     5
theory Sylow
b5522b51cb1e standard headers;
wenzelm
parents: 33657
diff changeset
     6
imports Coset Exponent
b5522b51cb1e standard headers;
wenzelm
parents: 33657
diff changeset
     7
begin
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
     8
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 59807
diff changeset
     9
text \<open>
58622
aa99568f56de more antiquotations;
wenzelm
parents: 55157
diff changeset
    10
  See also @{cite "Kammueller-Paulson:1999"}.
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 59807
diff changeset
    11
\<close>
14706
71590b7733b7 tuned document;
wenzelm
parents: 14666
diff changeset
    12
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 59807
diff changeset
    13
text\<open>The combinatorial argument is in theory Exponent\<close>
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    14
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    15
lemma le_extend_mult: 
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    16
  fixes c::nat shows "\<lbrakk>0 < c; a \<le> b\<rbrakk> \<Longrightarrow> a \<le> b * c"
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    17
by (metis divisors_zero dvd_triv_left leI less_le_trans nat_dvd_not_less zero_less_iff_neq_zero)
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    18
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
    19
locale sylow = group +
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    20
  fixes p and a and m and calM and RelM
16663
13e9c402308b prime is a predicate now.
nipkow
parents: 16417
diff changeset
    21
  assumes prime_p:   "prime p"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    22
      and order_G:   "order(G) = (p^a) * m"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    23
      and finite_G [iff]:  "finite (carrier G)"
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
    24
  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
    25
      and "RelM == {(N1,N2). N1 \<in> calM & N2 \<in> calM &
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    26
                             (\<exists>g \<in> carrier(G). N1 = (N2 #> g) )}"
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    27
begin
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    28
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    29
lemma RelM_refl_on: "refl_on calM RelM"
30198
922f944f03b2 name changes
nipkow
parents: 27717
diff changeset
    30
apply (auto simp add: refl_on_def RelM_def calM_def)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    31
apply (blast intro!: coset_mult_one [symmetric])
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    32
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    33
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    34
lemma RelM_sym: "sym RelM"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    35
proof (unfold sym_def RelM_def, clarify)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    36
  fix y g
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    37
  assume   "y \<in> calM"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    38
    and g: "g \<in> carrier G"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    39
  hence "y = y #> g #> (inv g)" by (simp add: coset_mult_assoc calM_def)
41541
1fa4725c4656 eliminated global prems;
wenzelm
parents: 35849
diff changeset
    40
  thus "\<exists>g'\<in>carrier G. y = y #> g #> g'" by (blast intro: g)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    41
qed
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    42
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    43
lemma RelM_trans: "trans RelM"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    44
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
    45
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    46
lemma RelM_equiv: "equiv calM RelM"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    47
apply (unfold equiv_def)
30198
922f944f03b2 name changes
nipkow
parents: 27717
diff changeset
    48
apply (blast intro: RelM_refl_on RelM_sym RelM_trans)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    49
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    50
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    51
lemma 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
    52
apply (unfold RelM_def)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    53
apply (blast elim!: quotientE)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    54
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    55
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    56
end
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 16663
diff changeset
    57
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 59807
diff changeset
    58
subsection\<open>Main Part of the Proof\<close>
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    59
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    60
locale sylow_central = sylow +
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    61
  fixes H and M1 and M
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    62
  assumes M_in_quot:  "M \<in> calM // RelM"
63537
831816778409 Removed redundant material related to primes
eberlm <eberlm@in.tum.de>
parents: 63534
diff changeset
    63
      and not_dvd_M:  "~(p ^ Suc(multiplicity p m) dvd card(M))"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    64
      and M1_in_M:    "M1 \<in> M"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    65
  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
    66
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    67
begin
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    68
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    69
lemma M_subset_calM: "M \<subseteq> calM"
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    70
  by (rule M_in_quot [THEN M_subset_calM_prep])
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    71
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    72
lemma card_M1: "card(M1) = p^a"
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    73
  using M1_in_M M_subset_calM calM_def by blast
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    74
 
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    75
lemma exists_x_in_M1: "\<exists>x. x \<in> M1"
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    76
using prime_p [THEN prime_gt_Suc_0_nat] card_M1
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    77
by (metis Suc_lessD card_eq_0_iff empty_subsetI equalityI gr_implies_not0 nat_zero_less_power_iff subsetI)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    78
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    79
lemma M1_subset_G [simp]: "M1 \<subseteq> carrier G"
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    80
  using M1_in_M  M_subset_calM calM_def mem_Collect_eq subsetCE by blast
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    81
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    82
lemma M1_inj_H: "\<exists>f \<in> H\<rightarrow>M1. inj_on f H"
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    83
proof -
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    84
  from exists_x_in_M1 obtain m1 where m1M: "m1 \<in> M1"..
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    85
  have m1G: "m1 \<in> carrier G" by (simp add: m1M M1_subset_G [THEN subsetD])
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    86
  show ?thesis
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    87
  proof
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    88
    show "inj_on (\<lambda>z\<in>H. m1 \<otimes> z) H"
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    89
      by (simp add: inj_on_def l_cancel [of m1 x y, THEN iffD1] H_def m1G)
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    90
    show "restrict (op \<otimes> m1) H \<in> H \<rightarrow> M1"
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    91
    proof (rule restrictI)
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    92
      fix z assume zH: "z \<in> H"
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    93
      show "m1 \<otimes> z \<in> M1"
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    94
      proof -
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    95
        from zH
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    96
        have zG: "z \<in> carrier G" and M1zeq: "M1 #> z = M1"
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    97
          by (auto simp add: H_def)
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    98
        show ?thesis
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    99
          by (rule subst [OF M1zeq], simp add: m1M zG rcosI)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   100
      qed
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   101
    qed
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   102
  qed
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   103
qed
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   104
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   105
end
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   106
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62410
diff changeset
   107
subsection\<open>Discharging the Assumptions of \<open>sylow_central\<close>\<close>
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   108
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   109
context sylow
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   110
begin
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   111
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   112
lemma EmptyNotInEquivSet: "{} \<notin> calM // RelM"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   113
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
   114
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   115
lemma existsM1inM: "M \<in> calM // RelM ==> \<exists>M1. M1 \<in> M"
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   116
  using RelM_equiv equiv_Eps_in by blast
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   117
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   118
lemma zero_less_o_G: "0 < order(G)"
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   119
  by (simp add: order_def card_gt_0_iff carrier_not_empty)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   120
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   121
lemma zero_less_m: "m > 0"
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   122
  using zero_less_o_G by (simp add: order_G)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   123
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   124
lemma card_calM: "card(calM) = (p^a) * m choose p^a"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   125
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
   126
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   127
lemma zero_less_card_calM: "card calM > 0"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   128
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
   129
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   130
lemma max_p_div_calM:
63537
831816778409 Removed redundant material related to primes
eberlm <eberlm@in.tum.de>
parents: 63534
diff changeset
   131
     "~ (p ^ Suc(multiplicity p m) dvd card(calM))"
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   132
proof
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   133
  assume "p ^ Suc (multiplicity p m) dvd card calM"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   134
  with zero_less_card_calM prime_p 
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   135
  have "Suc (multiplicity p m) \<le> multiplicity p (card calM)"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   136
    by (intro multiplicity_geI) auto
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   137
  hence "multiplicity p m < multiplicity p (card calM)" by simp
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   138
  also have "multiplicity p m = multiplicity p (card calM)"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   139
    by (simp add: const_p_fac prime_p zero_less_m card_calM)
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   140
  finally show False by simp
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   141
qed
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   142
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   143
lemma finite_calM: "finite calM"
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   144
  unfolding calM_def
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   145
  by (rule_tac B = "Pow (carrier G) " in finite_subset) auto
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   146
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   147
lemma lemma_A1:
63537
831816778409 Removed redundant material related to primes
eberlm <eberlm@in.tum.de>
parents: 63534
diff changeset
   148
     "\<exists>M \<in> calM // RelM. ~ (p ^ Suc(multiplicity p m) dvd card(M))"
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   149
  using RelM_equiv equiv_imp_dvd_card finite_calM max_p_div_calM by blast
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   150
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   151
end
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   152
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 59807
diff changeset
   153
subsubsection\<open>Introduction and Destruct Rules for @{term H}\<close>
13870
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_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
   156
by (simp add: H_def)
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_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
   159
by (simp add: H_def)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   160
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   161
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
   162
by (simp add: H_def)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   163
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   164
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
   165
apply (unfold H_def)
41541
1fa4725c4656 eliminated global prems;
wenzelm
parents: 35849
diff changeset
   166
apply (simp add: coset_mult_assoc [symmetric])
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   167
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   168
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   169
lemma (in sylow_central) H_not_empty: "H \<noteq> {}"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   170
apply (simp add: H_def)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   171
apply (rule exI [of _ \<one>], simp)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   172
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   173
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   174
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
   175
apply (rule subgroupI)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   176
apply (rule subsetI)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   177
apply (erule H_into_carrier_G)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   178
apply (rule H_not_empty)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   179
apply (simp add: H_def, clarify)
59807
22bc39064290 prefer local fixes;
wenzelm
parents: 58622
diff changeset
   180
apply (erule_tac P = "%z. lhs(z) = M1" for lhs in subst)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   181
apply (simp add: coset_mult_assoc )
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   182
apply (blast intro: H_m_closed)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   183
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   184
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   185
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   186
lemma (in sylow_central) rcosetGM1g_subset_G:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   187
     "[| 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
   188
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
   189
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   190
lemma (in sylow_central) finite_M1: "finite M1"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   191
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
   192
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   193
lemma (in sylow_central) finite_rcosetGM1g: "g\<in>carrier G ==> finite (M1 #> g)"
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   194
  using rcosetGM1g_subset_G finite_G M1_subset_G cosets_finite rcosetsI by blast
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   195
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   196
lemma (in sylow_central) M1_cardeq_rcosetGM1g:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   197
     "g \<in> carrier G ==> card(M1 #> g) = card(M1)"
41541
1fa4725c4656 eliminated global prems;
wenzelm
parents: 35849
diff changeset
   198
by (simp (no_asm_simp) add: card_cosets_equal rcosetsI)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   199
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   200
lemma (in sylow_central) M1_RelM_rcosetGM1g:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   201
     "g \<in> carrier G ==> (M1, M1 #> g) \<in> RelM"
55157
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41541
diff changeset
   202
apply (simp add: RelM_def calM_def card_M1)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   203
apply (rule conjI)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   204
 apply (blast intro: rcosetGM1g_subset_G)
55157
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41541
diff changeset
   205
apply (simp add: card_M1 M1_cardeq_rcosetGM1g)
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41541
diff changeset
   206
apply (metis M1_subset_G coset_mult_assoc coset_mult_one r_inv_ex)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   207
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   208
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   209
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 59807
diff changeset
   210
subsection\<open>Equal Cardinalities of @{term M} and the Set of Cosets\<close>
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   211
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 59807
diff changeset
   212
text\<open>Injections between @{term M} and @{term "rcosets\<^bsub>G\<^esub> H"} show that
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 59807
diff changeset
   213
 their cardinalities are equal.\<close>
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   214
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   215
lemma ElemClassEquiv:
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   216
     "[| equiv A r; C \<in> A // r |] ==> \<forall>x \<in> C. \<forall>y \<in> C. (x,y)\<in>r"
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   217
by (unfold equiv_def quotient_def sym_def trans_def, blast)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   218
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   219
lemma (in sylow_central) M_elem_map:
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   220
     "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
   221
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
   222
apply (simp add: RelM_def)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   223
apply (blast dest!: bspec)
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
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   226
lemmas (in sylow_central) M_elem_map_carrier =
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   227
        M_elem_map [THEN someI_ex, THEN conjunct1]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   228
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   229
lemmas (in sylow_central) M_elem_map_eq =
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   230
        M_elem_map [THEN someI_ex, THEN conjunct2]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   231
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   232
lemma (in sylow_central) M_funcset_rcosets_H:
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   233
     "(%x:M. H #> (SOME g. g \<in> carrier G & M1 #> g = x)) \<in> M \<rightarrow> rcosets H"
55157
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41541
diff changeset
   234
  by (metis (lifting) H_is_subgroup M_elem_map_carrier rcosetsI restrictI subgroup_imp_subset)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   235
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   236
lemma (in sylow_central) inj_M_GmodH: "\<exists>f \<in> M \<rightarrow> rcosets H. inj_on f M"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   237
apply (rule bexI)
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   238
apply (rule_tac [2] M_funcset_rcosets_H)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   239
apply (rule inj_onI, simp)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   240
apply (rule trans [OF _ M_elem_map_eq])
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   241
prefer 2 apply assumption
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   242
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
   243
apply (rule coset_mult_inv1)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   244
apply (erule_tac [2] M_elem_map_carrier)+
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   245
apply (rule_tac [2] M1_subset_G)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   246
apply (rule coset_join1 [THEN in_H_imp_eq])
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   247
apply (rule_tac [3] H_is_subgroup)
41541
1fa4725c4656 eliminated global prems;
wenzelm
parents: 35849
diff changeset
   248
prefer 2 apply (blast intro: M_elem_map_carrier)
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 25162
diff changeset
   249
apply (simp add: coset_mult_inv2 H_def M_elem_map_carrier subset_eq)
13870
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
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 59807
diff changeset
   253
subsubsection\<open>The Opposite Injection\<close>
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   254
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   255
lemma (in sylow_central) H_elem_map:
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   256
     "H1 \<in> rcosets H ==> \<exists>g. g \<in> carrier G & H #> g = H1"
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   257
by (auto simp add: RCOSETS_def)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   258
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   259
lemmas (in sylow_central) H_elem_map_carrier =
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   260
        H_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) H_elem_map_eq =
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   263
        H_elem_map [THEN someI_ex, THEN conjunct2]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   264
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   265
lemma (in sylow_central) rcosets_H_funcset_M:
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   266
  "(\<lambda>C \<in> rcosets H. M1 #> (@g. g \<in> carrier G \<and> H #> g = C)) \<in> rcosets H \<rightarrow> M"
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   267
apply (simp add: RCOSETS_def)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   268
apply (fast intro: someI2
55157
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41541
diff changeset
   269
            intro!: M1_in_M in_quotient_imp_closed [OF RelM_equiv M_in_quot _  M1_RelM_rcosetGM1g])
13870
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
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62410
diff changeset
   272
text\<open>close to a duplicate of \<open>inj_M_GmodH\<close>\<close>
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   273
lemma (in sylow_central) inj_GmodH_M:
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   274
     "\<exists>g \<in> rcosets H\<rightarrow>M. inj_on g (rcosets H)"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   275
apply (rule bexI)
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   276
apply (rule_tac [2] rcosets_H_funcset_M)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   277
apply (rule inj_onI)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   278
apply (simp)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   279
apply (rule trans [OF _ H_elem_map_eq])
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   280
prefer 2 apply assumption
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   281
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
   282
apply (rule coset_mult_inv1)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   283
apply (erule_tac [2] H_elem_map_carrier)+
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   284
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
   285
apply (rule coset_join2)
41541
1fa4725c4656 eliminated global prems;
wenzelm
parents: 35849
diff changeset
   286
apply (blast intro: H_elem_map_carrier)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   287
apply (rule H_is_subgroup)
41541
1fa4725c4656 eliminated global prems;
wenzelm
parents: 35849
diff changeset
   288
apply (simp add: H_I coset_mult_inv2 H_elem_map_carrier)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   289
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   290
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   291
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
   292
by (auto simp add: calM_def)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   293
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   294
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   295
lemma (in sylow_central) finite_M: "finite M"
55157
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41541
diff changeset
   296
by (metis M_subset_calM finite_calM rev_finite_subset)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   297
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   298
lemma (in sylow_central) cardMeqIndexH: "card(M) = card(rcosets H)"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   299
apply (insert inj_M_GmodH inj_GmodH_M)
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   300
apply (blast intro: card_bij finite_M H_is_subgroup
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14803
diff changeset
   301
             rcosets_subset_PowG [THEN finite_subset]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   302
             finite_Pow_iff [THEN iffD2])
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   303
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   304
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   305
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
   306
by (simp add: cardMeqIndexH lagrange H_is_subgroup)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   307
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   308
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
   309
apply (rule dvd_imp_le)
63537
831816778409 Removed redundant material related to primes
eberlm <eberlm@in.tum.de>
parents: 63534
diff changeset
   310
 apply (rule div_combine [OF prime_imp_prime_elem[OF prime_p] not_dvd_M])
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   311
 prefer 2 apply (blast intro: subgroup.finite_imp_card_positive H_is_subgroup)
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   312
apply (simp add: index_lem order_G power_add mult_dvd_mono multiplicity_dvd
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   313
                 zero_less_m)
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
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
   316
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
   317
apply (subst card_M1 [symmetric])
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   318
apply (cut_tac M1_inj_H)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   319
apply (blast intro!: M1_subset_G intro:
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   320
             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
   321
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   322
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   323
lemma (in sylow_central) card_H_eq: "card(H) = p^a"
33657
a4179bf442d1 renamed lemmas "anti_sym" -> "antisym"
nipkow
parents: 31754
diff changeset
   324
by (blast intro: le_antisym lemma_leq1 lemma_leq2)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   325
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   326
lemma (in sylow) sylow_thm: "\<exists>H. subgroup H G & card(H) = p^a"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   327
apply (cut_tac lemma_A1, clarify)
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   328
apply (frule existsM1inM, clarify)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   329
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
   330
 apply (blast dest:  sylow_central.H_is_subgroup sylow_central.card_H_eq)
41541
1fa4725c4656 eliminated global prems;
wenzelm
parents: 35849
diff changeset
   331
apply (simp add: sylow_central_def sylow_central_axioms_def sylow_axioms calM_def RelM_def)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   332
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   333
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 59807
diff changeset
   334
text\<open>Needed because the locale's automatic definition refers to
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   335
   @{term "semigroup G"} and @{term "group_axioms G"} rather than
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 59807
diff changeset
   336
  simply to @{term "group G"}.\<close>
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   337
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
   338
by (simp add: sylow_def group_def)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   339
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 16663
diff changeset
   340
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 59807
diff changeset
   341
subsection \<open>Sylow's Theorem\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 16663
diff changeset
   342
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   343
theorem sylow_thm:
16663
13e9c402308b prime is a predicate now.
nipkow
parents: 16417
diff changeset
   344
     "[| prime p;  group(G);  order(G) = (p^a) * m; finite (carrier G)|]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   345
      ==> \<exists>H. subgroup H G & card(H) = p^a"
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   346
apply (rule sylow.sylow_thm [of G p a m])
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   347
apply (simp add: sylow_eq sylow_axioms_def)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   348
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   349
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   350
end