author | haftmann |
Thu, 02 Jul 2020 12:10:58 +0000 | |
changeset 71989 | bad75618fb82 |
parent 66453 | cc19f7ca2ed6 |
permissions | -rw-r--r-- |
14706 | 1 |
(* Title: HOL/Algebra/Exponent.thy |
35849 | 2 |
Author: Florian Kammueller |
3 |
Author: L C Paulson |
|
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
4 |
|
35849 | 5 |
exponent p s yields the greatest power of p that divides s. |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
6 |
*) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
7 |
|
27105
5f139027c365
slightly tuning of some proofs involving case distinction and induction on natural numbers and similar
haftmann
parents:
25162
diff
changeset
|
8 |
theory Exponent |
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
65417
diff
changeset
|
9 |
imports Main "HOL-Computational_Algebra.Primes" |
27105
5f139027c365
slightly tuning of some proofs involving case distinction and induction on natural numbers and similar
haftmann
parents:
25162
diff
changeset
|
10 |
begin |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
20282
diff
changeset
|
11 |
|
61382 | 12 |
section \<open>Sylow's Theorem\<close> |
27717
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
ballarin
parents:
27651
diff
changeset
|
13 |
|
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
14 |
text \<open>The Combinatorial Argument Underlying the First Sylow Theorem\<close> |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
15 |
|
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
16 |
text\<open>needed in this form to prove Sylow's theorem\<close> |
63537
831816778409
Removed redundant material related to primes
eberlm <eberlm@in.tum.de>
parents:
63534
diff
changeset
|
17 |
corollary (in algebraic_semidom) div_combine: |
63633 | 18 |
"\<lbrakk>prime_elem p; \<not> p ^ Suc r dvd n; p ^ (a + r) dvd n * k\<rbrakk> \<Longrightarrow> p ^ a dvd k" |
19 |
by (metis add_Suc_right mult.commute prime_elem_power_dvd_cases) |
|
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
20 |
|
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
21 |
lemma exponent_p_a_m_k_equation: |
63534
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
62410
diff
changeset
|
22 |
fixes p :: nat |
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
23 |
assumes "0 < m" "0 < k" "p \<noteq> 0" "k < p^a" |
63537
831816778409
Removed redundant material related to primes
eberlm <eberlm@in.tum.de>
parents:
63534
diff
changeset
|
24 |
shows "multiplicity p (p^a * m - k) = multiplicity p (p^a - k)" |
831816778409
Removed redundant material related to primes
eberlm <eberlm@in.tum.de>
parents:
63534
diff
changeset
|
25 |
proof (rule multiplicity_cong [OF iffI]) |
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
26 |
fix r |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
27 |
assume *: "p ^ r dvd p ^ a * m - k" |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
28 |
show "p ^ r dvd p ^ a - k" |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
29 |
proof - |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
30 |
have "k \<le> p ^ a * m" using assms |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
31 |
by (meson nat_dvd_not_less dvd_triv_left leI mult_pos_pos order.strict_trans) |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
32 |
then have "r \<le> a" |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
33 |
by (meson "*" \<open>0 < k\<close> \<open>k < p^a\<close> dvd_diffD1 dvd_triv_left leI less_imp_le_nat nat_dvd_not_less power_le_dvd) |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
34 |
then have "p^r dvd p^a * m" by (simp add: le_imp_power_dvd) |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
35 |
thus ?thesis |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
36 |
by (meson \<open>k \<le> p ^ a * m\<close> \<open>r \<le> a\<close> * dvd_diffD1 dvd_diff_nat le_imp_power_dvd) |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
37 |
qed |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
38 |
next |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
39 |
fix r |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
40 |
assume *: "p ^ r dvd p ^ a - k" |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
41 |
with assms have "r \<le> a" |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
42 |
by (metis diff_diff_cancel less_imp_le_nat nat_dvd_not_less nat_le_linear power_le_dvd zero_less_diff) |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
43 |
show "p ^ r dvd p ^ a * m - k" |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
44 |
proof - |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
45 |
have "p^r dvd p^a*m" |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
46 |
by (simp add: \<open>r \<le> a\<close> le_imp_power_dvd) |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
47 |
then show ?thesis |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
48 |
by (meson assms * dvd_diffD1 dvd_diff_nat le_imp_power_dvd less_imp_le_nat \<open>r \<le> a\<close>) |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
49 |
qed |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
50 |
qed |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
51 |
|
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
52 |
lemma p_not_div_choose_lemma: |
63534
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
62410
diff
changeset
|
53 |
fixes p :: nat |
63537
831816778409
Removed redundant material related to primes
eberlm <eberlm@in.tum.de>
parents:
63534
diff
changeset
|
54 |
assumes eeq: "\<And>i. Suc i < K \<Longrightarrow> multiplicity p (Suc i) = multiplicity p (Suc (j + i))" |
63534
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
62410
diff
changeset
|
55 |
and "k < K" and p: "prime p" |
63537
831816778409
Removed redundant material related to primes
eberlm <eberlm@in.tum.de>
parents:
63534
diff
changeset
|
56 |
shows "multiplicity p (j + k choose k) = 0" |
63534
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
62410
diff
changeset
|
57 |
using \<open>k < K\<close> |
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
62410
diff
changeset
|
58 |
proof (induction k) |
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
62410
diff
changeset
|
59 |
case 0 then show ?case by simp |
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
60 |
next |
63534
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
62410
diff
changeset
|
61 |
case (Suc k) |
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
62410
diff
changeset
|
62 |
then have *: "(Suc (j+k) choose Suc k) > 0" by simp |
63537
831816778409
Removed redundant material related to primes
eberlm <eberlm@in.tum.de>
parents:
63534
diff
changeset
|
63 |
then have "multiplicity p ((Suc (j+k) choose Suc k) * Suc k) = multiplicity p (Suc k)" |
63633 | 64 |
by (subst Suc_times_binomial_eq [symmetric], subst prime_elem_multiplicity_mult_distrib) |
63534
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
62410
diff
changeset
|
65 |
(insert p Suc.prems, simp_all add: eeq [symmetric] Suc.IH) |
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
62410
diff
changeset
|
66 |
with p * show ?case |
63633 | 67 |
by (subst (asm) prime_elem_multiplicity_mult_distrib) simp_all |
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
68 |
qed |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
69 |
|
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
70 |
text\<open>The lemma above, with two changes of variables\<close> |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
71 |
lemma p_not_div_choose: |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
72 |
assumes "k < K" and "k \<le> n" |
63633 | 73 |
and eeq: "\<And>j. \<lbrakk>0<j; j<K\<rbrakk> \<Longrightarrow> multiplicity p (n - k + (K - j)) = multiplicity p (K - j)" "prime p" |
63537
831816778409
Removed redundant material related to primes
eberlm <eberlm@in.tum.de>
parents:
63534
diff
changeset
|
74 |
shows "multiplicity p (n choose k) = 0" |
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
75 |
apply (rule p_not_div_choose_lemma [of K p "n-k" k, simplified assms nat_minus_add_max max_absorb1]) |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
76 |
apply (metis add_Suc_right eeq diff_diff_cancel order_less_imp_le zero_less_Suc zero_less_diff) |
63534
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
62410
diff
changeset
|
77 |
apply (rule TrueI)+ |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
78 |
done |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
79 |
|
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
80 |
proposition const_p_fac: |
63633 | 81 |
assumes "m>0" and prime: "prime p" |
63537
831816778409
Removed redundant material related to primes
eberlm <eberlm@in.tum.de>
parents:
63534
diff
changeset
|
82 |
shows "multiplicity p (p^a * m choose p^a) = multiplicity p m" |
63534
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
62410
diff
changeset
|
83 |
proof- |
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
62410
diff
changeset
|
84 |
from assms have p: "0 < p ^ a" "0 < p^a * m" "p^a \<le> p^a * m" |
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
85 |
by (auto simp: prime_gt_0_nat) |
63537
831816778409
Removed redundant material related to primes
eberlm <eberlm@in.tum.de>
parents:
63534
diff
changeset
|
86 |
have *: "multiplicity p ((p^a * m - 1) choose (p^a - 1)) = 0" |
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
87 |
apply (rule p_not_div_choose [where K = "p^a"]) |
63534
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
62410
diff
changeset
|
88 |
using p exponent_p_a_m_k_equation by (auto simp: diff_le_mono prime) |
63537
831816778409
Removed redundant material related to primes
eberlm <eberlm@in.tum.de>
parents:
63534
diff
changeset
|
89 |
have "multiplicity p ((p ^ a * m choose p ^ a) * p ^ a) = a + multiplicity p m" |
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
90 |
proof - |
63534
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
62410
diff
changeset
|
91 |
have "(p ^ a * m choose p ^ a) * p ^ a = p ^ a * m * (p ^ a * m - 1 choose (p ^ a - 1))" |
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
62410
diff
changeset
|
92 |
(is "_ = ?rhs") using prime |
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
62410
diff
changeset
|
93 |
by (subst times_binomial_minus1_eq [symmetric]) (auto simp: prime_gt_0_nat) |
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
62410
diff
changeset
|
94 |
also from p have "p ^ a - Suc 0 \<le> p ^ a * m - Suc 0" by linarith |
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
62410
diff
changeset
|
95 |
with prime * p have "multiplicity p ?rhs = multiplicity p (p ^ a * m)" |
63633 | 96 |
by (subst prime_elem_multiplicity_mult_distrib) auto |
63534
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
62410
diff
changeset
|
97 |
also have "\<dots> = a + multiplicity p m" |
63633 | 98 |
using prime p by (subst prime_elem_multiplicity_mult_distrib) simp_all |
63534
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
62410
diff
changeset
|
99 |
finally show ?thesis . |
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
100 |
qed |
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
101 |
then show ?thesis |
63633 | 102 |
using prime p by (subst (asm) prime_elem_multiplicity_mult_distrib) simp_all |
62410
2fc7a8d9c529
partial tidy-up of Sylow's theorem
paulson <lp15@cam.ac.uk>
parents:
62349
diff
changeset
|
103 |
qed |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
104 |
|
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff
changeset
|
105 |
end |