src/HOL/Algebra/Sylow.thy
author wenzelm
Sat, 14 Sep 2019 22:13:36 +0200
changeset 70696 47ca5c7550e4
parent 69597 ff784d5a5bfb
child 76987 4c275405faae
permissions -rw-r--r--
potentially more robust: read under lock if not yet set;
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
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
     6
  imports Coset Exponent
35849
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
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
     9
text \<open>See also @{cite "Kammueller-Paulson:1999"}.\<close>
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    10
69272
15e9ed5b28fb isabelle update_cartouches -t;
wenzelm
parents: 68561
diff changeset
    11
text \<open>The combinatorial argument is in theory \<open>Exponent\<close>.\<close>
14706
71590b7733b7 tuned document;
wenzelm
parents: 14666
diff changeset
    12
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    13
lemma le_extend_mult: "\<lbrakk>0 < c; a \<le> b\<rbrakk> \<Longrightarrow> a \<le> b * c"
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    14
  for c :: nat
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    15
  by (metis divisors_zero dvd_triv_left leI less_le_trans nat_dvd_not_less zero_less_iff_neq_zero)
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    16
14747
2eaff69d3d10 removal of locale coset
paulson
parents: 14706
diff changeset
    17
locale sylow = group +
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    18
  fixes p and a and m and calM and RelM
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    19
  assumes prime_p: "prime p"
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    20
    and order_G: "order G = (p^a) * m"
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    21
    and finite_G[iff]: "finite (carrier G)"
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    22
  defines "calM \<equiv> {s. s \<subseteq> carrier G \<and> card s = p^a}"
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    23
    and "RelM \<equiv> {(N1, N2). N1 \<in> calM \<and> N2 \<in> calM \<and> (\<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
    24
begin
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    25
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    26
lemma RelM_refl_on: "refl_on calM RelM"
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    27
  by (auto simp: refl_on_def RelM_def calM_def) (blast intro!: coset_mult_one [symmetric])
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_sym: "sym RelM"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    30
proof (unfold sym_def RelM_def, clarify)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    31
  fix y g
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    32
  assume "y \<in> calM"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    33
    and g: "g \<in> carrier G"
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    34
  then have "y = y #> g #> (inv g)"
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    35
    by (simp add: coset_mult_assoc calM_def)
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    36
  then show "\<exists>g'\<in>carrier G. y = y #> g #> g'"
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    37
    by (blast intro: g)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    38
qed
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    39
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    40
lemma RelM_trans: "trans RelM"
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    41
  by (auto simp add: trans_def RelM_def calM_def coset_mult_assoc)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    42
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    43
lemma RelM_equiv: "equiv calM RelM"
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    44
  unfolding equiv_def by (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
    45
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    46
lemma M_subset_calM_prep: "M' \<in> calM // RelM  \<Longrightarrow> M' \<subseteq> calM"
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    47
  unfolding RelM_def by (blast elim!: quotientE)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    48
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    49
end
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 16663
diff changeset
    50
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    51
subsection \<open>Main Part of the Proof\<close>
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    52
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    53
locale sylow_central = sylow +
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    54
  fixes H and M1 and M
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    55
  assumes M_in_quot: "M \<in> calM // RelM"
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    56
    and not_dvd_M: "\<not> (p ^ Suc (multiplicity p m) dvd card M)"
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    57
    and M1_in_M: "M1 \<in> M"
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    58
  defines "H \<equiv> {g. g \<in> carrier G \<and> M1 #> g = M1}"
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    59
begin
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    60
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    61
lemma M_subset_calM: "M \<subseteq> calM"
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    62
  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
    63
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    64
lemma card_M1: "card M1 = p^a"
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    65
  using M1_in_M M_subset_calM calM_def by blast
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    66
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    67
lemma exists_x_in_M1: "\<exists>x. x \<in> M1"
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    68
  using prime_p [THEN prime_gt_Suc_0_nat] card_M1
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    69
  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
    70
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    71
lemma M1_subset_G [simp]: "M1 \<subseteq> carrier G"
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    72
  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
    73
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    74
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
    75
proof -
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    76
  from exists_x_in_M1 obtain m1 where m1M: "m1 \<in> M1"..
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    77
  have m1: "m1 \<in> carrier G"
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    78
    by (simp add: m1M M1_subset_G [THEN subsetD])
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    79
  show ?thesis
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    80
  proof
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    81
    show "inj_on (\<lambda>z\<in>H. m1 \<otimes> z) H"
68399
0b71d08528f0 resolution of name clashes in Algebra
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
    82
      by (simp add: H_def inj_on_def m1)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 64914
diff changeset
    83
    show "restrict ((\<otimes>) m1) H \<in> H \<rightarrow> M1"
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    84
    proof (rule restrictI)
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    85
      fix z
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    86
      assume zH: "z \<in> H"
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    87
      show "m1 \<otimes> z \<in> M1"
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    88
      proof -
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    89
        from zH
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    90
        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
    91
          by (auto simp add: H_def)
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    92
        show ?thesis
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
    93
          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
    94
      qed
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    95
    qed
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    96
  qed
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    97
qed
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    98
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
    99
end
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   100
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   101
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   102
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
   103
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   104
context sylow
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   105
begin
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   106
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   107
lemma EmptyNotInEquivSet: "{} \<notin> calM // RelM"
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   108
  by (blast elim!: quotientE dest: RelM_equiv [THEN equiv_class_self])
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   109
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   110
lemma existsM1inM: "M \<in> calM // RelM \<Longrightarrow> \<exists>M1. M1 \<in> M"
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   111
  using RelM_equiv equiv_Eps_in by blast
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   112
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   113
lemma zero_less_o_G: "0 < order G"
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   114
  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
   115
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   116
lemma zero_less_m: "m > 0"
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   117
  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
   118
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   119
lemma card_calM: "card calM = (p^a) * m choose p^a"
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   120
  by (simp add: calM_def n_subsets order_G [symmetric] order_def)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   121
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   122
lemma zero_less_card_calM: "card calM > 0"
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   123
  by (simp add: card_calM zero_less_binomial le_extend_mult zero_less_m)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   124
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   125
lemma max_p_div_calM: "\<not> (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
   126
proof
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   127
  assume "p ^ Suc (multiplicity p m) dvd card calM"
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   128
  with zero_less_card_calM prime_p
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   129
  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
   130
    by (intro multiplicity_geI) auto
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   131
  then have "multiplicity p m < multiplicity p (card calM)" by simp
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   132
  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
   133
    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
   134
  finally show False by simp
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   135
qed
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   136
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   137
lemma finite_calM: "finite calM"
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   138
  unfolding calM_def by (rule finite_subset [where B = "Pow (carrier G)"]) auto
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   139
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   140
lemma lemma_A1: "\<exists>M \<in> calM // RelM. \<not> (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
   141
  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
   142
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   143
end
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   144
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   145
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   146
subsubsection \<open>Introduction and Destruct Rules for \<open>H\<close>\<close>
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   147
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   148
context sylow_central
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   149
begin
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   150
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   151
lemma H_I: "\<lbrakk>g \<in> carrier G; M1 #> g = M1\<rbrakk> \<Longrightarrow> g \<in> H"
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   152
  by (simp add: H_def)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   153
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   154
lemma H_into_carrier_G: "x \<in> H \<Longrightarrow> x \<in> carrier G"
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   155
  by (simp add: H_def)
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   156
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   157
lemma in_H_imp_eq: "g \<in> H \<Longrightarrow> M1 #> g = M1"
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   158
  by (simp add: H_def)
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   159
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   160
lemma H_m_closed: "\<lbrakk>x \<in> H; y \<in> H\<rbrakk> \<Longrightarrow> x \<otimes> y \<in> H"
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   161
  by (simp add: H_def coset_mult_assoc [symmetric])
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   162
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   163
lemma H_not_empty: "H \<noteq> {}"
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   164
  by (force simp add: H_def intro: exI [of _ \<one>])
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   165
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   166
lemma H_is_subgroup: "subgroup H G"
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   167
proof (rule subgroupI)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   168
  show "H \<subseteq> carrier G"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   169
    using H_into_carrier_G by blast
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   170
  show "\<And>a. a \<in> H \<Longrightarrow> inv a \<in> H"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   171
    by (metis H_I H_into_carrier_G H_m_closed M1_subset_G Units_eq Units_inv_closed Units_inv_inv coset_mult_inv1 in_H_imp_eq)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   172
  show "\<And>a b. \<lbrakk>a \<in> H; b \<in> H\<rbrakk> \<Longrightarrow> a \<otimes> b \<in> H"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   173
    by (blast intro: H_m_closed)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   174
qed (use H_not_empty in auto)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   175
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   176
lemma rcosetGM1g_subset_G: "\<lbrakk>g \<in> carrier G; x \<in> M1 #> g\<rbrakk> \<Longrightarrow> x \<in> carrier G"
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   177
  by (blast intro: M1_subset_G [THEN r_coset_subset_G, THEN subsetD])
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   178
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   179
lemma finite_M1: "finite M1"
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   180
  by (rule finite_subset [OF M1_subset_G finite_G])
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   181
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   182
lemma finite_rcosetGM1g: "g \<in> carrier G \<Longrightarrow> finite (M1 #> g)"
62410
2fc7a8d9c529 partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents: 61382
diff changeset
   183
  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
   184
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   185
lemma M1_cardeq_rcosetGM1g: "g \<in> carrier G \<Longrightarrow> card (M1 #> g) = card M1"
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68399
diff changeset
   186
  by (metis M1_subset_G card_rcosets_equal rcosetsI)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   187
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   188
lemma M1_RelM_rcosetGM1g: 
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   189
  assumes "g \<in> carrier G"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   190
  shows "(M1, M1 #> g) \<in> RelM"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   191
proof -
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   192
  have "M1 #> g \<subseteq> carrier G"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   193
    by (simp add: assms r_coset_subset_G)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   194
  moreover have "card (M1 #> g) = p ^ a"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   195
    using assms by (simp add: card_M1 M1_cardeq_rcosetGM1g)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   196
  moreover have "\<exists>h\<in>carrier G. M1 = M1 #> g #> h"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   197
    by (metis assms M1_subset_G coset_mult_assoc coset_mult_one r_inv_ex)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   198
  ultimately show ?thesis
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   199
    by (simp add: RelM_def calM_def card_M1)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   200
qed
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   201
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   202
end
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   203
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   204
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   205
subsection \<open>Equal Cardinalities of \<open>M\<close> and the Set of Cosets\<close>
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   206
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69272
diff changeset
   207
text \<open>Injections between \<^term>\<open>M\<close> and \<^term>\<open>rcosets\<^bsub>G\<^esub> H\<close> show that
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 59807
diff changeset
   208
 their cardinalities are equal.\<close>
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   209
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   210
lemma ElemClassEquiv: "\<lbrakk>equiv A r; C \<in> A // r\<rbrakk> \<Longrightarrow> \<forall>x \<in> C. \<forall>y \<in> C. (x, y) \<in> r"
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   211
  unfolding equiv_def quotient_def sym_def trans_def by blast
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   212
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   213
context sylow_central
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   214
begin
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   215
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   216
lemma M_elem_map: "M2 \<in> M \<Longrightarrow> \<exists>g. g \<in> carrier G \<and> M1 #> g = M2"
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   217
  using M1_in_M M_in_quot [THEN RelM_equiv [THEN ElemClassEquiv]]
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   218
  by (simp add: RelM_def) (blast dest!: bspec)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   219
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   220
lemmas M_elem_map_carrier = M_elem_map [THEN someI_ex, THEN conjunct1]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   221
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   222
lemmas M_elem_map_eq = M_elem_map [THEN someI_ex, THEN conjunct2]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   223
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   224
lemma M_funcset_rcosets_H:
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   225
  "(\<lambda>x\<in>M. H #> (SOME g. g \<in> carrier G \<and> M1 #> g = x)) \<in> M \<rightarrow> rcosets H"
68445
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   226
  by (metis (lifting) H_is_subgroup M_elem_map_carrier rcosetsI restrictI subgroup.subset)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   227
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   228
lemma inj_M_GmodH: "\<exists>f \<in> M \<rightarrow> rcosets H. inj_on f M"
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   229
proof
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   230
  let ?inv = "\<lambda>x. SOME g. g \<in> carrier G \<and> M1 #> g = x"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   231
  show "inj_on (\<lambda>x\<in>M. H #> ?inv x) M"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   232
  proof (rule inj_onI, simp)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   233
    fix x y
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   234
    assume eq: "H #> ?inv x = H #> ?inv y" and xy: "x \<in> M" "y \<in> M"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   235
    have "x = M1 #> ?inv x"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   236
      by (simp add: M_elem_map_eq \<open>x \<in> M\<close>)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   237
    also have "... = M1 #> ?inv y"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   238
    proof (rule coset_mult_inv1 [OF in_H_imp_eq [OF coset_join1]])
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   239
      show "H #> ?inv x \<otimes> inv (?inv y) = H"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   240
        by (simp add: H_into_carrier_G M_elem_map_carrier xy coset_mult_inv2 eq subsetI)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   241
    qed (simp_all add: H_is_subgroup M_elem_map_carrier xy)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   242
    also have "... = y"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   243
      using M_elem_map_eq \<open>y \<in> M\<close> by simp
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   244
    finally show "x=y" .
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   245
  qed
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   246
  show "(\<lambda>x\<in>M. H #> ?inv x) \<in> M \<rightarrow> rcosets H"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   247
    by (rule M_funcset_rcosets_H)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   248
qed
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   249
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   250
end
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   251
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   252
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
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
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   255
context sylow_central
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   256
begin
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   257
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   258
lemma H_elem_map: "H1 \<in> rcosets H \<Longrightarrow> \<exists>g. g \<in> carrier G \<and> H #> g = H1"
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   259
  by (auto simp: RCOSETS_def)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   260
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   261
lemmas H_elem_map_carrier = H_elem_map [THEN someI_ex, THEN conjunct1]
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   262
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   263
lemmas H_elem_map_eq = 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
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   265
lemma rcosets_H_funcset_M:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   266
  "(\<lambda>C \<in> rcosets H. M1 #> (SOME g. g \<in> carrier G \<and> H #> g = C)) \<in> rcosets H \<rightarrow> M"
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   267
  using in_quotient_imp_closed [OF RelM_equiv M_in_quot _  M1_RelM_rcosetGM1g]
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   268
  by (simp add: M1_in_M H_elem_map_carrier RCOSETS_def)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   269
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   270
lemma inj_GmodH_M: "\<exists>g \<in> rcosets H\<rightarrow>M. inj_on g (rcosets H)"
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   271
proof
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   272
  let ?inv = "\<lambda>x. SOME g. g \<in> carrier G \<and> H #> g = x"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   273
  show "inj_on (\<lambda>C\<in>rcosets H. M1 #> ?inv C) (rcosets H)"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   274
  proof (rule inj_onI, simp)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   275
    fix x y
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   276
    assume eq: "M1 #> ?inv x = M1 #> ?inv y" and xy: "x \<in> rcosets H" "y \<in> rcosets H"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   277
    have "x = H #> ?inv x"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   278
      by (simp add: H_elem_map_eq \<open>x \<in> rcosets H\<close>)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   279
    also have "... = H #> ?inv y"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   280
    proof (rule coset_mult_inv1 [OF coset_join2])
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   281
      show "?inv x \<otimes> inv (?inv y) \<in> carrier G"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   282
        by (simp add: H_elem_map_carrier \<open>x \<in> rcosets H\<close> \<open>y \<in> rcosets H\<close>)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   283
      then show "(?inv x) \<otimes> inv (?inv y) \<in> H"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   284
        by (simp add: H_I H_elem_map_carrier xy coset_mult_inv2 eq)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   285
      show "H \<subseteq> carrier G"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   286
        by (simp add: H_is_subgroup subgroup.subset)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   287
    qed (simp_all add: H_is_subgroup H_elem_map_carrier xy)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   288
    also have "... = y"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   289
      by (simp add: H_elem_map_eq \<open>y \<in> rcosets H\<close>)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   290
    finally show "x=y" .
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   291
  qed
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   292
  show "(\<lambda>C\<in>rcosets H. M1 #> ?inv C) \<in> rcosets H \<rightarrow> M"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   293
    using rcosets_H_funcset_M by blast
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   294
qed
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   295
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   296
lemma calM_subset_PowG: "calM \<subseteq> Pow (carrier G)"
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   297
  by (auto simp: calM_def)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   298
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   299
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   300
lemma finite_M: "finite M"
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   301
  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
   302
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   303
lemma cardMeqIndexH: "card M = card (rcosets H)"
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   304
  using inj_M_GmodH inj_GmodH_M
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   305
  by (blast intro: card_bij finite_M H_is_subgroup rcosets_subset_PowG [THEN finite_subset])
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   306
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   307
lemma index_lem: "card M * card H = order G"
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   308
  by (simp add: cardMeqIndexH lagrange H_is_subgroup)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   309
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   310
lemma card_H_eq: "card H = p^a"
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   311
proof (rule antisym)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   312
  show "p^a \<le> card H"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   313
  proof (rule dvd_imp_le)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   314
    show "p ^ a dvd card H"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   315
      apply (rule div_combine [OF prime_imp_prime_elem[OF prime_p] not_dvd_M])
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   316
      by (simp add: index_lem multiplicity_dvd order_G power_add)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   317
    show "0 < card H"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   318
      by (blast intro: subgroup.finite_imp_card_positive H_is_subgroup)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   319
  qed
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   320
next
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   321
  show "card H \<le> p^a"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   322
    using M1_inj_H card_M1 card_inj finite_M1 by fastforce
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   323
qed
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   324
64914
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   325
end
51f015bd4565 prefer context groups;
wenzelm
parents: 64912
diff changeset
   326
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   327
lemma (in sylow) sylow_thm: "\<exists>H. subgroup H G \<and> card H = p^a"
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   328
proof -
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   329
  obtain M where M: "M \<in> calM // RelM" "\<not> (p ^ Suc (multiplicity p m) dvd card M)"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   330
    using lemma_A1 by blast
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   331
  then obtain M1 where "M1 \<in> M"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   332
    by (metis existsM1inM) 
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   333
  define H where "H \<equiv> {g. g \<in> carrier G \<and> M1 #> g = M1}"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   334
  with M \<open>M1 \<in> M\<close>
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   335
  interpret sylow_central G p a m calM RelM H M1 M
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   336
    by unfold_locales (auto simp add: H_def calM_def RelM_def)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   337
  show ?thesis
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   338
    using H_is_subgroup card_H_eq by blast
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 68484
diff changeset
   339
qed
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   340
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   341
text \<open>Needed because the locale's automatic definition refers to
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69272
diff changeset
   342
  \<^term>\<open>semigroup G\<close> and \<^term>\<open>group_axioms G\<close> rather than
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69272
diff changeset
   343
  simply to \<^term>\<open>group G\<close>.\<close>
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   344
lemma sylow_eq: "sylow G p a m \<longleftrightarrow> group G \<and> sylow_axioms G p a m"
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   345
  by (simp add: sylow_def group_def)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   346
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 16663
diff changeset
   347
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 59807
diff changeset
   348
subsection \<open>Sylow's Theorem\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 16663
diff changeset
   349
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   350
theorem sylow_thm:
64912
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   351
  "\<lbrakk>prime p; group G; order G = (p^a) * m; finite (carrier G)\<rbrakk>
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   352
    \<Longrightarrow> \<exists>H. subgroup H G \<and> card H = p^a"
68f0465d956b misc tuning and modernization;
wenzelm
parents: 63537
diff changeset
   353
  by (rule sylow.sylow_thm [of G p a m]) (simp add: sylow_eq sylow_axioms_def)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   354
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   355
end