author | wenzelm |
Sat, 16 Nov 2024 15:04:41 +0100 | |
changeset 81460 | 6ea0055fa42d |
parent 78668 | d52934f126d4 |
permissions | -rw-r--r-- |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1 |
(* Title: HOL/Number_Theory/Pocklington.thy |
69785
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
2 |
Author: Amine Chaieb, Manuel Eberl |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
3 |
*) |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
4 |
|
60526 | 5 |
section \<open>Pocklington's Theorem for Primes\<close> |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
6 |
|
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
7 |
theory Pocklington |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
8 |
imports Residues |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
9 |
begin |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
10 |
|
66305 | 11 |
subsection \<open>Lemmas about previously defined terms\<close> |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
12 |
|
66305 | 13 |
lemma prime_nat_iff'': "prime (p::nat) \<longleftrightarrow> p \<noteq> 0 \<and> p \<noteq> 1 \<and> (\<forall>m. 0 < m \<and> m < p \<longrightarrow> coprime p m)" |
67051 | 14 |
apply (auto simp add: prime_nat_iff) |
15 |
apply (rule coprimeI) |
|
16 |
apply (auto dest: nat_dvd_not_less simp add: ac_simps) |
|
17 |
apply (metis One_nat_def dvd_1_iff_1 dvd_pos_nat gcd_nat.order_iff is_unit_gcd linorder_neqE_nat nat_dvd_not_less) |
|
18 |
done |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
19 |
|
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
20 |
lemma finite_number_segment: "card { m. 0 < m \<and> m < n } = n - 1" |
66305 | 21 |
proof - |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
22 |
have "{ m. 0 < m \<and> m < n } = {1..<n}" by auto |
66305 | 23 |
then show ?thesis by simp |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
24 |
qed |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
25 |
|
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
26 |
|
66305 | 27 |
subsection \<open>Some basic theorems about solving congruences\<close> |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
28 |
|
66305 | 29 |
lemma cong_solve: |
30 |
fixes n :: nat |
|
31 |
assumes an: "coprime a n" |
|
32 |
shows "\<exists>x. [a * x = b] (mod n)" |
|
33 |
proof (cases "a = 0") |
|
34 |
case True |
|
35 |
with an show ?thesis |
|
66888 | 36 |
by (simp add: cong_def) |
66305 | 37 |
next |
38 |
case False |
|
39 |
from bezout_add_strong_nat [OF this] |
|
40 |
obtain d x y where dxy: "d dvd a" "d dvd n" "a * x = n * y + d" by blast |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
41 |
from dxy(1,2) have d1: "d = 1" |
67051 | 42 |
using assms coprime_common_divisor [of a n d] by simp |
66305 | 43 |
with dxy(3) have "a * x * b = (n * y + 1) * b" |
44 |
by simp |
|
45 |
then have "a * (x * b) = n * (y * b) + b" |
|
46 |
by (auto simp: algebra_simps) |
|
47 |
then have "a * (x * b) mod n = (n * (y * b) + b) mod n" |
|
48 |
by simp |
|
49 |
then have "a * (x * b) mod n = b mod n" |
|
50 |
by (simp add: mod_add_left_eq) |
|
51 |
then have "[a * (x * b) = b] (mod n)" |
|
66888 | 52 |
by (simp only: cong_def) |
66305 | 53 |
then show ?thesis by blast |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
54 |
qed |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
55 |
|
66305 | 56 |
lemma cong_solve_unique: |
57 |
fixes n :: nat |
|
58 |
assumes an: "coprime a n" and nz: "n \<noteq> 0" |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
59 |
shows "\<exists>!x. x < n \<and> [a * x = b] (mod n)" |
66305 | 60 |
proof - |
61 |
from cong_solve[OF an] obtain x where x: "[a * x = b] (mod n)" |
|
62 |
by blast |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
63 |
let ?P = "\<lambda>x. x < n \<and> [a * x = b] (mod n)" |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
64 |
let ?x = "x mod n" |
66305 | 65 |
from x have *: "[a * ?x = b] (mod n)" |
66888 | 66 |
by (simp add: cong_def mod_mult_right_eq[of a x n]) |
66305 | 67 |
from mod_less_divisor[ of n x] nz * have Px: "?P ?x" by simp |
68 |
have "y = ?x" if Py: "y < n" "[a * y = b] (mod n)" for y |
|
69 |
proof - |
|
70 |
from Py(2) * have "[a * y = a * ?x] (mod n)" |
|
66888 | 71 |
by (simp add: cong_def) |
66305 | 72 |
then have "[y = ?x] (mod n)" |
73 |
by (metis an cong_mult_lcancel_nat) |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
74 |
with mod_less[OF Py(1)] mod_less_divisor[ of n x] nz |
66305 | 75 |
show ?thesis |
66888 | 76 |
by (simp add: cong_def) |
66305 | 77 |
qed |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
78 |
with Px show ?thesis by blast |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
79 |
qed |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
80 |
|
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
81 |
lemma cong_solve_unique_nontrivial: |
66305 | 82 |
fixes p :: nat |
83 |
assumes p: "prime p" |
|
84 |
and pa: "coprime p a" |
|
85 |
and x0: "0 < x" |
|
86 |
and xp: "x < p" |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
87 |
shows "\<exists>!y. 0 < y \<and> y < p \<and> [x * y = a] (mod p)" |
66305 | 88 |
proof - |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
89 |
from pa have ap: "coprime a p" |
67051 | 90 |
by (simp add: ac_simps) |
91 |
from x0 xp p have px: "coprime x p" |
|
92 |
by (auto simp add: prime_nat_iff'' ac_simps) |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
93 |
obtain y where y: "y < p" "[x * y = a] (mod p)" "\<forall>z. z < p \<and> [x * z = a] (mod p) \<longrightarrow> z = y" |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
94 |
by (metis cong_solve_unique neq0_conv p prime_gt_0_nat px) |
66305 | 95 |
have "y \<noteq> 0" |
96 |
proof |
|
97 |
assume "y = 0" |
|
98 |
with y(2) have "p dvd a" |
|
66888 | 99 |
using cong_dvd_iff by auto |
67051 | 100 |
with not_prime_1 p pa show False |
101 |
by (auto simp add: gcd_nat.order_iff) |
|
66305 | 102 |
qed |
103 |
with y show ?thesis |
|
104 |
by blast |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
105 |
qed |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
106 |
|
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
107 |
lemma cong_unique_inverse_prime: |
66305 | 108 |
fixes p :: nat |
109 |
assumes "prime p" and "0 < x" and "x < p" |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
110 |
shows "\<exists>!y. 0 < y \<and> y < p \<and> [x * y = 1] (mod p)" |
66305 | 111 |
by (rule cong_solve_unique_nontrivial) (use assms in simp_all) |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
112 |
|
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
113 |
lemma chinese_remainder_coprime_unique: |
66305 | 114 |
fixes a :: nat |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
115 |
assumes ab: "coprime a b" and az: "a \<noteq> 0" and bz: "b \<noteq> 0" |
66305 | 116 |
and ma: "coprime m a" and nb: "coprime n b" |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
117 |
shows "\<exists>!x. coprime x (a * b) \<and> x < a * b \<and> [x = m] (mod a) \<and> [x = n] (mod b)" |
66305 | 118 |
proof - |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
119 |
let ?P = "\<lambda>x. x < a * b \<and> [x = m] (mod a) \<and> [x = n] (mod b)" |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
120 |
from binary_chinese_remainder_unique_nat[OF ab az bz] |
66305 | 121 |
obtain x where x: "x < a * b" "[x = m] (mod a)" "[x = n] (mod b)" "\<forall>y. ?P y \<longrightarrow> y = x" |
122 |
by blast |
|
123 |
from ma nb x have "coprime x a" "coprime x b" |
|
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
68157
diff
changeset
|
124 |
using cong_imp_coprime cong_sym by blast+ |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
125 |
then have "coprime x (a*b)" |
67051 | 126 |
by simp |
66305 | 127 |
with x show ?thesis |
128 |
by blast |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
129 |
qed |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
130 |
|
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
131 |
|
66305 | 132 |
subsection \<open>Lucas's theorem\<close> |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
133 |
|
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
134 |
lemma lucas_coprime_lemma: |
66305 | 135 |
fixes n :: nat |
136 |
assumes m: "m \<noteq> 0" and am: "[a^m = 1] (mod n)" |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
137 |
shows "coprime a n" |
66305 | 138 |
proof - |
139 |
consider "n = 1" | "n = 0" | "n > 1" by arith |
|
140 |
then show ?thesis |
|
141 |
proof cases |
|
142 |
case 1 |
|
143 |
then show ?thesis by simp |
|
144 |
next |
|
145 |
case 2 |
|
146 |
with am m show ?thesis |
|
66888 | 147 |
by simp |
66305 | 148 |
next |
149 |
case 3 |
|
150 |
from m obtain m' where m': "m = Suc m'" by (cases m) blast+ |
|
151 |
have "d = 1" if d: "d dvd a" "d dvd n" for d |
|
152 |
proof - |
|
153 |
from am mod_less[OF \<open>n > 1\<close>] have am1: "a^m mod n = 1" |
|
66888 | 154 |
by (simp add: cong_def) |
66305 | 155 |
from dvd_mult2[OF d(1), of "a^m'"] have dam: "d dvd a^m" |
156 |
by (simp add: m') |
|
157 |
from dvd_mod_iff[OF d(2), of "a^m"] dam am1 show ?thesis |
|
158 |
by simp |
|
159 |
qed |
|
67051 | 160 |
then show ?thesis |
161 |
by (auto intro: coprimeI) |
|
66305 | 162 |
qed |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
163 |
qed |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
164 |
|
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
165 |
lemma lucas_weak: |
66305 | 166 |
fixes n :: nat |
167 |
assumes n: "n \<ge> 2" |
|
168 |
and an: "[a ^ (n - 1) = 1] (mod n)" |
|
169 |
and nm: "\<forall>m. 0 < m \<and> m < n - 1 \<longrightarrow> \<not> [a ^ m = 1] (mod n)" |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
170 |
shows "prime n" |
65726
f5d64d094efe
More material on totient function
eberlm <eberlm@in.tum.de>
parents:
65465
diff
changeset
|
171 |
proof (rule totient_imp_prime) |
65465
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
172 |
show "totient n = n - 1" |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
173 |
proof (rule ccontr) |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
174 |
have "[a ^ totient n = 1] (mod n)" |
66305 | 175 |
by (rule euler_theorem, rule lucas_coprime_lemma [of "n - 1"]) (use n an in auto) |
65465
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
176 |
moreover assume "totient n \<noteq> n - 1" |
66305 | 177 |
then have "totient n > 0" "totient n < n - 1" |
178 |
using \<open>n \<ge> 2\<close> and totient_less[of n] by simp_all |
|
65465
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
179 |
ultimately show False |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
180 |
using nm by auto |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
181 |
qed |
66305 | 182 |
qed (use n in auto) |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
183 |
|
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
184 |
lemma nat_exists_least_iff: "(\<exists>(n::nat). P n) \<longleftrightarrow> (\<exists>n. P n \<and> (\<forall>m < n. \<not> P m))" |
55337
5d45fb978d5a
Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents:
55321
diff
changeset
|
185 |
by (metis ex_least_nat_le not_less0) |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
186 |
|
66305 | 187 |
lemma nat_exists_least_iff': "(\<exists>(n::nat). P n) \<longleftrightarrow> P (Least P) \<and> (\<forall>m < (Least P). \<not> P m)" |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
188 |
(is "?lhs \<longleftrightarrow> ?rhs") |
66305 | 189 |
proof |
190 |
show ?lhs if ?rhs |
|
191 |
using that by blast |
|
192 |
show ?rhs if ?lhs |
|
193 |
proof - |
|
194 |
from \<open>?lhs\<close> obtain n where n: "P n" by blast |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
195 |
let ?x = "Least P" |
66305 | 196 |
have "\<not> P m" if "m < ?x" for m |
197 |
by (rule not_less_Least[OF that]) |
|
198 |
with LeastI_ex[OF \<open>?lhs\<close>] show ?thesis |
|
199 |
by blast |
|
200 |
qed |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
201 |
qed |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
202 |
|
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
203 |
theorem lucas: |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
204 |
assumes n2: "n \<ge> 2" and an1: "[a^(n - 1) = 1] (mod n)" |
66305 | 205 |
and pn: "\<forall>p. prime p \<and> p dvd n - 1 \<longrightarrow> [a^((n - 1) div p) \<noteq> 1] (mod n)" |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
206 |
shows "prime n" |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
207 |
proof- |
66305 | 208 |
from n2 have n01: "n \<noteq> 0" "n \<noteq> 1" "n - 1 \<noteq> 0" |
209 |
by arith+ |
|
210 |
from mod_less_divisor[of n 1] n01 have onen: "1 mod n = 1" |
|
211 |
by simp |
|
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
68157
diff
changeset
|
212 |
from lucas_coprime_lemma[OF n01(3) an1] cong_imp_coprime an1 |
67051 | 213 |
have an: "coprime a n" "coprime (a ^ (n - 1)) n" |
214 |
using \<open>n \<ge> 2\<close> by simp_all |
|
67091 | 215 |
have False if H0: "\<exists>m. 0 < m \<and> m < n - 1 \<and> [a ^ m = 1] (mod n)" (is "\<exists>m. ?P m") |
66305 | 216 |
proof - |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
217 |
from H0[unfolded nat_exists_least_iff[of ?P]] obtain m where |
66305 | 218 |
m: "0 < m" "m < n - 1" "[a ^ m = 1] (mod n)" "\<forall>k <m. \<not>?P k" |
219 |
by blast |
|
220 |
have False if nm1: "(n - 1) mod m > 0" |
|
221 |
proof - |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
222 |
from mod_less_divisor[OF m(1)] have th0:"(n - 1) mod m < m" by blast |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
223 |
let ?y = "a^ ((n - 1) div m * m)" |
64242
93c6f0da5c70
more standardized theorem names for facts involving the div and mod identity
haftmann
parents:
63905
diff
changeset
|
224 |
note mdeq = div_mult_mod_eq[of "(n - 1)" m] |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
225 |
have yn: "coprime ?y n" |
67051 | 226 |
using an(1) by (cases "(n - Suc 0) div m * m = 0") auto |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
227 |
have "?y mod n = (a^m)^((n - 1) div m) mod n" |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
228 |
by (simp add: algebra_simps power_mult) |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
229 |
also have "\<dots> = (a^m mod n)^((n - 1) div m) mod n" |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
230 |
using power_mod[of "a^m" n "(n - 1) div m"] by simp |
66888 | 231 |
also have "\<dots> = 1" using m(3)[unfolded cong_def onen] onen |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
232 |
by (metis power_one) |
66305 | 233 |
finally have *: "?y mod n = 1" . |
234 |
have **: "[?y * a ^ ((n - 1) mod m) = ?y* 1] (mod n)" |
|
66888 | 235 |
using an1[unfolded cong_def onen] onen |
64242
93c6f0da5c70
more standardized theorem names for facts involving the div and mod identity
haftmann
parents:
63905
diff
changeset
|
236 |
div_mult_mod_eq[of "(n - 1)" m, symmetric] |
66888 | 237 |
by (simp add:power_add[symmetric] cong_def * del: One_nat_def) |
66305 | 238 |
have "[a ^ ((n - 1) mod m) = 1] (mod n)" |
239 |
by (metis cong_mult_rcancel_nat mult.commute ** yn) |
|
240 |
with m(4)[rule_format, OF th0] nm1 |
|
241 |
less_trans[OF mod_less_divisor[OF m(1), of "n - 1"] m(2)] show ?thesis |
|
242 |
by blast |
|
243 |
qed |
|
244 |
then have "(n - 1) mod m = 0" by auto |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
245 |
then have mn: "m dvd n - 1" by presburger |
66305 | 246 |
then obtain r where r: "n - 1 = m * r" |
247 |
unfolding dvd_def by blast |
|
248 |
from n01 r m(2) have r01: "r \<noteq> 0" "r \<noteq> 1" by auto |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
249 |
obtain p where p: "prime p" "p dvd r" |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
250 |
by (metis prime_factor_nat r01(2)) |
66305 | 251 |
then have th: "prime p \<and> p dvd n - 1" |
252 |
unfolding r by (auto intro: dvd_mult) |
|
253 |
from r have "(a ^ ((n - 1) div p)) mod n = (a^(m*r div p)) mod n" |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
254 |
by (simp add: power_mult) |
66305 | 255 |
also have "\<dots> = (a^(m*(r div p))) mod n" |
256 |
using div_mult1_eq[of m r p] p(2)[unfolded dvd_eq_mod_eq_0] by simp |
|
257 |
also have "\<dots> = ((a^m)^(r div p)) mod n" |
|
258 |
by (simp add: power_mult) |
|
259 |
also have "\<dots> = ((a^m mod n)^(r div p)) mod n" |
|
260 |
using power_mod .. |
|
261 |
also from m(3) onen have "\<dots> = 1" |
|
66888 | 262 |
by (simp add: cong_def) |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
263 |
finally have "[(a ^ ((n - 1) div p))= 1] (mod n)" |
66888 | 264 |
using onen by (simp add: cong_def) |
66305 | 265 |
with pn th show ?thesis by blast |
266 |
qed |
|
267 |
then have "\<forall>m. 0 < m \<and> m < n - 1 \<longrightarrow> \<not> [a ^ m = 1] (mod n)" |
|
268 |
by blast |
|
269 |
then show ?thesis by (rule lucas_weak[OF n2 an1]) |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
270 |
qed |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
271 |
|
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
272 |
|
69785
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
273 |
subsection \<open>Definition of the order of a number mod \<open>n\<close>\<close> |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
274 |
|
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
275 |
definition "ord n a = (if coprime n a then Least (\<lambda>d. d > 0 \<and> [a ^d = 1] (mod n)) else 0)" |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
276 |
|
66305 | 277 |
text \<open>This has the expected properties.\<close> |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
278 |
|
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
279 |
lemma coprime_ord: |
66305 | 280 |
fixes n::nat |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
281 |
assumes "coprime n a" |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
282 |
shows "ord n a > 0 \<and> [a ^(ord n a) = 1] (mod n) \<and> (\<forall>m. 0 < m \<and> m < ord n a \<longrightarrow> [a^ m \<noteq> 1] (mod n))" |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
283 |
proof- |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
284 |
let ?P = "\<lambda>d. 0 < d \<and> [a ^ d = 1] (mod n)" |
66305 | 285 |
from bigger_prime[of a] obtain p where p: "prime p" "a < p" |
286 |
by blast |
|
287 |
from assms have o: "ord n a = Least ?P" |
|
288 |
by (simp add: ord_def) |
|
65465
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
289 |
have ex: "\<exists>m>0. ?P m" |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
290 |
proof (cases "n \<ge> 2") |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
291 |
case True |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
292 |
moreover from assms have "coprime a n" |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
293 |
by (simp add: ac_simps) |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
294 |
then have "[a ^ totient n = 1] (mod n)" |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
295 |
by (rule euler_theorem) |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
296 |
ultimately show ?thesis |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
297 |
by (auto intro: exI [where x = "totient n"]) |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
298 |
next |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
299 |
case False |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
300 |
then have "n = 0 \<or> n = 1" |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
301 |
by auto |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
302 |
with assms show ?thesis |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
303 |
by auto |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
304 |
qed |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
305 |
from nat_exists_least_iff'[of ?P] ex assms show ?thesis |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
306 |
unfolding o[symmetric] by auto |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
307 |
qed |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
308 |
|
66305 | 309 |
text \<open>With the special value \<open>0\<close> for non-coprime case, it's more convenient.\<close> |
310 |
lemma ord_works: "[a ^ (ord n a) = 1] (mod n) \<and> (\<forall>m. 0 < m \<and> m < ord n a \<longrightarrow> \<not> [a^ m = 1] (mod n))" |
|
311 |
for n :: nat |
|
66888 | 312 |
by (cases "coprime n a") (use coprime_ord[of n a] in \<open>auto simp add: ord_def cong_def\<close>) |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
313 |
|
66305 | 314 |
lemma ord: "[a^(ord n a) = 1] (mod n)" |
315 |
for n :: nat |
|
316 |
using ord_works by blast |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
317 |
|
66305 | 318 |
lemma ord_minimal: "0 < m \<Longrightarrow> m < ord n a \<Longrightarrow> \<not> [a^m = 1] (mod n)" |
319 |
for n :: nat |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
320 |
using ord_works by blast |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
321 |
|
66305 | 322 |
lemma ord_eq_0: "ord n a = 0 \<longleftrightarrow> \<not> coprime n a" |
323 |
for n :: nat |
|
324 |
by (cases "coprime n a") (simp add: coprime_ord, simp add: ord_def) |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
325 |
|
66305 | 326 |
lemma divides_rexp: "x dvd y \<Longrightarrow> x dvd (y ^ Suc n)" |
327 |
for x y :: nat |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
328 |
by (simp add: dvd_mult2[of x y]) |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
329 |
|
66305 | 330 |
lemma ord_divides:"[a ^ d = 1] (mod n) \<longleftrightarrow> ord n a dvd d" |
331 |
(is "?lhs \<longleftrightarrow> ?rhs") |
|
332 |
for n :: nat |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
333 |
proof |
66305 | 334 |
assume ?rhs |
335 |
then obtain k where "d = ord n a * k" |
|
336 |
unfolding dvd_def by blast |
|
337 |
then have "[a ^ d = (a ^ (ord n a) mod n)^k] (mod n)" |
|
66888 | 338 |
by (simp add : cong_def power_mult power_mod) |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
339 |
also have "[(a ^ (ord n a) mod n)^k = 1] (mod n)" |
66888 | 340 |
using ord[of a n, unfolded cong_def] |
341 |
by (simp add: cong_def power_mod) |
|
66305 | 342 |
finally show ?lhs . |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
343 |
next |
66305 | 344 |
assume ?lhs |
345 |
show ?rhs |
|
346 |
proof (cases "coprime n a") |
|
347 |
case prem: False |
|
348 |
then have o: "ord n a = 0" by (simp add: ord_def) |
|
349 |
show ?thesis |
|
350 |
proof (cases d) |
|
351 |
case 0 |
|
66888 | 352 |
with o prem show ?thesis by (simp add: cong_def) |
66305 | 353 |
next |
354 |
case (Suc d') |
|
355 |
then have d0: "d \<noteq> 0" by simp |
|
67051 | 356 |
from prem obtain p where p: "p dvd n" "p dvd a" "p \<noteq> 1" |
357 |
by (auto elim: not_coprimeE) |
|
66305 | 358 |
from \<open>?lhs\<close> obtain q1 q2 where q12: "a ^ d + n * q1 = 1 + n * q2" |
67051 | 359 |
using prem d0 lucas_coprime_lemma |
360 |
by (auto elim: not_coprimeE simp add: ac_simps) |
|
66305 | 361 |
then have "a ^ d + n * q1 - n * q2 = 1" by simp |
362 |
with dvd_diff_nat [OF dvd_add [OF divides_rexp]] dvd_mult2 Suc p have "p dvd 1" |
|
55337
5d45fb978d5a
Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents:
55321
diff
changeset
|
363 |
by metis |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
364 |
with p(3) have False by simp |
66305 | 365 |
then show ?thesis .. |
366 |
qed |
|
367 |
next |
|
368 |
case H: True |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
369 |
let ?o = "ord n a" |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
370 |
let ?q = "d div ord n a" |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
371 |
let ?r = "d mod ord n a" |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
372 |
have eqo: "[(a^?o)^?q = 1] (mod n)" |
66888 | 373 |
using cong_pow ord_works by fastforce |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
374 |
from H have onz: "?o \<noteq> 0" by (simp add: ord_eq_0) |
67345 | 375 |
then have opos: "?o > 0" by simp |
66305 | 376 |
from div_mult_mod_eq[of d "ord n a"] \<open>?lhs\<close> |
377 |
have "[a^(?o*?q + ?r) = 1] (mod n)" |
|
66888 | 378 |
by (simp add: cong_def mult.commute) |
66305 | 379 |
then have "[(a^?o)^?q * (a^?r) = 1] (mod n)" |
66888 | 380 |
by (simp add: cong_def power_mult[symmetric] power_add[symmetric]) |
66305 | 381 |
then have th: "[a^?r = 1] (mod n)" |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
382 |
using eqo mod_mult_left_eq[of "(a^?o)^?q" "a^?r" n] |
66888 | 383 |
by (simp add: cong_def del: One_nat_def) (metis mod_mult_left_eq nat_mult_1) |
66305 | 384 |
show ?thesis |
385 |
proof (cases "?r = 0") |
|
386 |
case True |
|
387 |
then show ?thesis by (simp add: dvd_eq_mod_eq_0) |
|
388 |
next |
|
389 |
case False |
|
67345 | 390 |
with mod_less_divisor[OF opos, of d] have r0o:"?r >0 \<and> ?r < ?o" by simp |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
391 |
from conjunct2[OF ord_works[of a n], rule_format, OF r0o] th |
66305 | 392 |
show ?thesis by blast |
393 |
qed |
|
394 |
qed |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
395 |
qed |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
396 |
|
67051 | 397 |
lemma order_divides_totient: |
398 |
"ord n a dvd totient n" if "coprime n a" |
|
399 |
using that euler_theorem [of a n] |
|
400 |
by (simp add: ord_divides [symmetric] ac_simps) |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
401 |
|
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
402 |
lemma order_divides_expdiff: |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
403 |
fixes n::nat and a::nat assumes na: "coprime n a" |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
404 |
shows "[a^d = a^e] (mod n) \<longleftrightarrow> [d = e] (mod (ord n a))" |
66305 | 405 |
proof - |
406 |
have th: "[a^d = a^e] (mod n) \<longleftrightarrow> [d = e] (mod (ord n a))" |
|
407 |
if na: "coprime n a" and ed: "(e::nat) \<le> d" |
|
408 |
for n a d e :: nat |
|
409 |
proof - |
|
410 |
from na ed have "\<exists>c. d = e + c" by presburger |
|
411 |
then obtain c where c: "d = e + c" .. |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
412 |
from na have an: "coprime a n" |
67051 | 413 |
by (simp add: ac_simps) |
414 |
then have aen: "coprime (a ^ e) n" |
|
415 |
by (cases "e > 0") simp_all |
|
416 |
from an have acn: "coprime (a ^ c) n" |
|
417 |
by (cases "c > 0") simp_all |
|
66305 | 418 |
from c have "[a^d = a^e] (mod n) \<longleftrightarrow> [a^(e + c) = a^(e + 0)] (mod n)" |
419 |
by simp |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
420 |
also have "\<dots> \<longleftrightarrow> [a^e* a^c = a^e *a^0] (mod n)" by (simp add: power_add) |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
421 |
also have "\<dots> \<longleftrightarrow> [a ^ c = 1] (mod n)" |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
422 |
using cong_mult_lcancel_nat [OF aen, of "a^c" "a^0"] by simp |
66305 | 423 |
also have "\<dots> \<longleftrightarrow> ord n a dvd c" |
424 |
by (simp only: ord_divides) |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
425 |
also have "\<dots> \<longleftrightarrow> [e + c = e + 0] (mod ord n a)" |
66888 | 426 |
by (auto simp add: cong_altdef_nat) |
66305 | 427 |
finally show ?thesis |
428 |
using c by simp |
|
429 |
qed |
|
430 |
consider "e \<le> d" | "d \<le> e" by arith |
|
431 |
then show ?thesis |
|
432 |
proof cases |
|
433 |
case 1 |
|
434 |
with na show ?thesis by (rule th) |
|
435 |
next |
|
436 |
case 2 |
|
437 |
from th[OF na this] show ?thesis |
|
66888 | 438 |
by (metis cong_sym) |
66305 | 439 |
qed |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
440 |
qed |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
441 |
|
69785
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
442 |
lemma ord_not_coprime [simp]: "\<not>coprime n a \<Longrightarrow> ord n a = 0" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
443 |
by (simp add: ord_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
444 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
445 |
lemma ord_1 [simp]: "ord 1 n = 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
446 |
proof - |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
447 |
have "(LEAST k. k > 0) = (1 :: nat)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
448 |
by (rule Least_equality) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
449 |
thus ?thesis by (simp add: ord_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
450 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
451 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
452 |
lemma ord_1_right [simp]: "ord (n::nat) 1 = 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
453 |
using ord_divides[of 1 1 n] by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
454 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
455 |
lemma ord_Suc_0_right [simp]: "ord (n::nat) (Suc 0) = 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
456 |
using ord_divides[of 1 1 n] by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
457 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
458 |
lemma ord_0_nat [simp]: "ord 0 (n :: nat) = (if n = 1 then 1 else 0)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
459 |
proof - |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
460 |
have "(LEAST k. k > 0) = (1 :: nat)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
461 |
by (rule Least_equality) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
462 |
thus ?thesis by (auto simp: ord_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
463 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
464 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
465 |
lemma ord_0_right_nat [simp]: "ord (n :: nat) 0 = (if n = 1 then 1 else 0)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
466 |
proof - |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
467 |
have "(LEAST k. k > 0) = (1 :: nat)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
468 |
by (rule Least_equality) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
469 |
thus ?thesis by (auto simp: ord_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
470 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
471 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
472 |
lemma ord_divides': "[a ^ d = Suc 0] (mod n) = (ord n a dvd d)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
473 |
using ord_divides[of a d n] by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
474 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
475 |
lemma ord_Suc_0 [simp]: "ord (Suc 0) n = 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
476 |
using ord_1[where 'a = nat] by (simp del: ord_1) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
477 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
478 |
lemma ord_mod [simp]: "ord n (k mod n) = ord n k" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
479 |
by (cases "n = 0") (auto simp add: ord_def cong_def power_mod) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
480 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
481 |
lemma ord_gt_0_iff [simp]: "ord (n::nat) x > 0 \<longleftrightarrow> coprime n x" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
482 |
using ord_eq_0[of n x] by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
483 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
484 |
lemma ord_eq_Suc_0_iff: "ord n (x::nat) = Suc 0 \<longleftrightarrow> [x = 1] (mod n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
485 |
using ord_divides[of x 1 n] by (auto simp: ord_divides') |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
486 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
487 |
lemma ord_cong: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
488 |
assumes "[k1 = k2] (mod n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
489 |
shows "ord n k1 = ord n k2" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
490 |
proof - |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
491 |
have "ord n (k1 mod n) = ord n (k2 mod n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
492 |
by (simp only: assms[unfolded cong_def]) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
493 |
thus ?thesis by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
494 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
495 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
496 |
lemma ord_nat_code [code_unfold]: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
497 |
"ord n a = |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
498 |
(if n = 0 then if a = 1 then 1 else 0 else |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
499 |
if coprime n a then Min (Set.filter (\<lambda>k. [a ^ k = 1] (mod n)) {0<..n}) else 0)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
500 |
proof (cases "coprime n a \<and> n > 0") |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
501 |
case True |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
502 |
define A where "A = {k\<in>{0<..n}. [a ^ k = 1] (mod n)}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
503 |
define k where "k = (LEAST k. k > 0 \<and> [a ^ k = 1] (mod n))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
504 |
have totient: "totient n \<in> A" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
505 |
using euler_theorem[of a n] True |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
506 |
by (auto simp: A_def coprime_commute intro!: Nat.gr0I totient_le) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
507 |
moreover have "finite A" by (auto simp: A_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
508 |
ultimately have *: "Min A \<in> A" and "\<forall>y. y \<in> A \<longrightarrow> Min A \<le> y" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
509 |
by (auto intro: Min_in) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
510 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
511 |
have "k > 0 \<and> [a ^ k = 1] (mod n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
512 |
unfolding k_def by (rule LeastI[of _ "totient n"]) (use totient in \<open>auto simp: A_def\<close>) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
513 |
moreover have "k \<le> totient n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
514 |
unfolding k_def by (intro Least_le) (use totient in \<open>auto simp: A_def\<close>) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
515 |
ultimately have "k \<in> A" using totient_le[of n] by (auto simp: A_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
516 |
hence "Min A \<le> k" by (intro Min_le) (auto simp: \<open>finite A\<close>) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
517 |
moreover from * have "k \<le> Min A" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
518 |
unfolding k_def by (intro Least_le) (auto simp: A_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
519 |
ultimately show ?thesis using True by (simp add: ord_def k_def A_def Set.filter_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
520 |
qed auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
521 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
522 |
theorem ord_modulus_mult_coprime: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
523 |
fixes x :: nat |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
524 |
assumes "coprime m n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
525 |
shows "ord (m * n) x = lcm (ord m x) (ord n x)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
526 |
proof (intro dvd_antisym) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
527 |
have "[x ^ lcm (ord m x) (ord n x) = 1] (mod (m * n))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
528 |
using assms by (intro coprime_cong_mult_nat assms) (auto simp: ord_divides') |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
529 |
thus "ord (m * n) x dvd lcm (ord m x) (ord n x)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
530 |
by (simp add: ord_divides') |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
531 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
532 |
show "lcm (ord m x) (ord n x) dvd ord (m * n) x" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
533 |
proof (intro lcm_least) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
534 |
show "ord m x dvd ord (m * n) x" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
535 |
using cong_modulus_mult_nat[of "x ^ ord (m * n) x" 1 m n] assms |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
536 |
by (simp add: ord_divides') |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
537 |
show "ord n x dvd ord (m * n) x" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
538 |
using cong_modulus_mult_nat[of "x ^ ord (m * n) x" 1 n m] assms |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
539 |
by (simp add: ord_divides' mult.commute) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
540 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
541 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
542 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
543 |
corollary ord_modulus_prod_coprime: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
544 |
assumes "finite A" "\<And>i j. i \<in> A \<Longrightarrow> j \<in> A \<Longrightarrow> i \<noteq> j \<Longrightarrow> coprime (f i) (f j)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
545 |
shows "ord (\<Prod>i\<in>A. f i :: nat) x = (LCM i\<in>A. ord (f i) x)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
546 |
using assms by (induction A rule: finite_induct) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
547 |
(simp, simp, subst ord_modulus_mult_coprime, auto intro!: prod_coprime_right) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
548 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
549 |
lemma ord_power_aux: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
550 |
fixes m x k a :: nat |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
551 |
defines "l \<equiv> ord m a" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
552 |
shows "ord m (a ^ k) * gcd k l = l" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
553 |
proof (rule dvd_antisym) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
554 |
have "[a ^ lcm k l = 1] (mod m)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
555 |
unfolding ord_divides by (simp add: l_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
556 |
also have "lcm k l = k * (l div gcd k l)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
557 |
by (simp add: lcm_nat_def div_mult_swap) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
558 |
finally have "ord m (a ^ k) dvd l div gcd k l" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
559 |
unfolding ord_divides [symmetric] by (simp add: power_mult [symmetric]) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
560 |
thus "ord m (a ^ k) * gcd k l dvd l" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
561 |
by (cases "l = 0") (auto simp: dvd_div_iff_mult) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
562 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
563 |
have "[(a ^ k) ^ ord m (a ^ k) = 1] (mod m)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
564 |
by (rule ord) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
565 |
also have "(a ^ k) ^ ord m (a ^ k) = a ^ (k * ord m (a ^ k))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
566 |
by (simp add: power_mult) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
567 |
finally have "ord m a dvd k * ord m (a ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
568 |
by (simp add: ord_divides') |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
569 |
hence "l dvd gcd (k * ord m (a ^ k)) (l * ord m (a ^ k))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
570 |
by (intro gcd_greatest dvd_triv_left) (auto simp: l_def ord_divides') |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
571 |
also have "gcd (k * ord m (a ^ k)) (l * ord m (a ^ k)) = ord m (a ^ k) * gcd k l" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
572 |
by (subst gcd_mult_distrib_nat) (auto simp: mult_ac) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
573 |
finally show "l dvd ord m (a ^ k) * gcd k l" . |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
574 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
575 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
576 |
theorem ord_power: "coprime m a \<Longrightarrow> ord m (a ^ k :: nat) = ord m a div gcd k (ord m a)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
577 |
using ord_power_aux[of m a k] by (metis div_mult_self_is_m gcd_pos_nat ord_eq_0) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
578 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
579 |
lemma inj_power_mod: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
580 |
assumes "coprime n (a :: nat)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
581 |
shows "inj_on (\<lambda>k. a ^ k mod n) {..<ord n a}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
582 |
proof |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
583 |
fix k l assume *: "k \<in> {..<ord n a}" "l \<in> {..<ord n a}" "a ^ k mod n = a ^ l mod n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
584 |
have "k = l" if "k < l" "l < ord n a" "[a ^ k = a ^ l] (mod n)" for k l |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
585 |
proof - |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
586 |
have "l = k + (l - k)" using that by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
587 |
also have "a ^ \<dots> = a ^ k * a ^ (l - k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
588 |
by (simp add: power_add) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
589 |
also have "[\<dots> = a ^ l * a ^ (l - k)] (mod n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
590 |
using that by (intro cong_mult) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
591 |
finally have "[a ^ l * a ^ (l - k) = a ^ l * 1] (mod n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
592 |
by (simp add: cong_sym_eq) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
593 |
with assms have "[a ^ (l - k) = 1] (mod n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
594 |
by (subst (asm) cong_mult_lcancel_nat) (auto simp: coprime_commute) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
595 |
hence "ord n a dvd l - k" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
596 |
by (simp add: ord_divides') |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
597 |
from dvd_imp_le[OF this] and \<open>l < ord n a\<close> have "l - k = 0" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
598 |
by (cases "l - k = 0") auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
599 |
with \<open>k < l\<close> show "k = l" by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
600 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
601 |
from this[of k l] and this[of l k] and * show "k = l" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
602 |
by (cases k l rule: linorder_cases) (auto simp: cong_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
603 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
604 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
605 |
lemma ord_eq_2_iff: "ord n (x :: nat) = 2 \<longleftrightarrow> [x \<noteq> 1] (mod n) \<and> [x\<^sup>2 = 1] (mod n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
606 |
proof |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
607 |
assume x: "[x \<noteq> 1] (mod n) \<and> [x\<^sup>2 = 1] (mod n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
608 |
hence "coprime n x" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
609 |
by (metis coprime_commute lucas_coprime_lemma zero_neq_numeral) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
610 |
with x have "ord n x dvd 2" "ord n x \<noteq> 1" "ord n x > 0" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
611 |
by (auto simp: ord_divides' ord_eq_Suc_0_iff) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
612 |
thus "ord n x = 2" by (auto dest!: dvd_imp_le simp del: ord_gt_0_iff) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
613 |
qed (use ord_divides[of _ 2] ord_divides[of _ 1] in auto) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
614 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
615 |
lemma square_mod_8_eq_1_iff: "[x\<^sup>2 = 1] (mod 8) \<longleftrightarrow> odd (x :: nat)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
616 |
proof - |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
617 |
have "[x\<^sup>2 = 1] (mod 8) \<longleftrightarrow> ((x mod 8)\<^sup>2 mod 8 = 1)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
618 |
by (simp add: power_mod cong_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
619 |
also have "\<dots> \<longleftrightarrow> x mod 8 \<in> {1, 3, 5, 7}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
620 |
proof |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
621 |
assume x: "(x mod 8)\<^sup>2 mod 8 = 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
622 |
have "x mod 8 \<in> {..<8}" by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
623 |
also have "{..<8} = {0, 1, 2, 3, 4, 5, 6, 7::nat}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
624 |
by (simp add: lessThan_nat_numeral lessThan_Suc insert_commute) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
625 |
finally have x_cases: "x mod 8 \<in> {0, 1, 2, 3, 4, 5, 6, 7}" . |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
626 |
from x have "x mod 8 \<notin> {0, 2, 4, 6}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
627 |
using x by (auto intro: Nat.gr0I) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
628 |
with x_cases show "x mod 8 \<in> {1, 3, 5, 7}" by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
629 |
qed auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
630 |
also have "\<dots> \<longleftrightarrow> odd (x mod 8)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
631 |
by (auto elim!: oddE) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
632 |
also have "\<dots> \<longleftrightarrow> odd x" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
633 |
by presburger |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
634 |
finally show ?thesis . |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
635 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
636 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
637 |
lemma ord_twopow_aux: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
638 |
assumes "k \<ge> 3" and "odd (x :: nat)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
639 |
shows "[x ^ (2 ^ (k - 2)) = 1] (mod (2 ^ k))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
640 |
using assms(1) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
641 |
proof (induction k rule: dec_induct) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
642 |
case base |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
643 |
from assms have "[x\<^sup>2 = 1] (mod 8)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
644 |
by (subst square_mod_8_eq_1_iff) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
645 |
thus ?case by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
646 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
647 |
case (step k) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
648 |
define k' where "k' = k - 2" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
649 |
have k: "k = Suc (Suc k')" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
650 |
using \<open>k \<ge> 3\<close> by (simp add: k'_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
651 |
from \<open>k \<ge> 3\<close> have "2 * k \<ge> Suc k" by presburger |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
652 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
653 |
from \<open>odd x\<close> have "x > 0" by (intro Nat.gr0I) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
654 |
from step.IH have "2 ^ k dvd (x ^ (2 ^ (k - 2)) - 1)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
655 |
by (rule cong_to_1_nat) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
656 |
then obtain t where "x ^ (2 ^ (k - 2)) - 1 = t * 2 ^ k" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
657 |
by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
658 |
hence "x ^ (2 ^ (k - 2)) = t * 2 ^ k + 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
659 |
by (metis \<open>0 < x\<close> add.commute add_diff_inverse_nat less_one neq0_conv power_eq_0_iff) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
660 |
hence "(x ^ (2 ^ (k - 2))) ^ 2 = (t * 2 ^ k + 1) ^ 2" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
661 |
by (rule arg_cong) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
662 |
hence "[(x ^ (2 ^ (k - 2))) ^ 2 = (t * 2 ^ k + 1) ^ 2] (mod (2 ^ Suc k))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
663 |
by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
664 |
also have "(x ^ (2 ^ (k - 2))) ^ 2 = x ^ (2 ^ (k - 1))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
665 |
by (simp_all add: power_even_eq[symmetric] power_mult k ) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
666 |
also have "(t * 2 ^ k + 1) ^ 2 = t\<^sup>2 * 2 ^ (2 * k) + t * 2 ^ Suc k + 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
667 |
by (subst power2_eq_square) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
668 |
(auto simp: algebra_simps k power2_eq_square[of t] |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
669 |
power_even_eq[symmetric] power_add [symmetric]) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
670 |
also have "[\<dots> = 0 + 0 + 1] (mod 2 ^ Suc k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
671 |
using \<open>2 * k \<ge> Suc k\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
672 |
by (intro cong_add) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
673 |
(auto simp: cong_0_iff intro: dvd_mult[OF le_imp_power_dvd] simp del: power_Suc) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
674 |
finally show ?case by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
675 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
676 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
677 |
lemma ord_twopow_3_5: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
678 |
assumes "k \<ge> 3" "x mod 8 \<in> {3, 5 :: nat}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
679 |
shows "ord (2 ^ k) x = 2 ^ (k - 2)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
680 |
using assms(1) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
681 |
proof (induction k rule: less_induct) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
682 |
have "x mod 8 = 3 \<or> x mod 8 = 5" using assms by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
683 |
hence "odd x" by presburger |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
684 |
case (less k) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
685 |
from \<open>k \<ge> 3\<close> consider "k = 3" | "k = 4" | "k \<ge> 5" by force |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
686 |
thus ?case |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
687 |
proof cases |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
688 |
case 1 |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
689 |
thus ?thesis using assms |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
690 |
by (auto simp: ord_eq_2_iff cong_def simp flip: power_mod[of x]) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
691 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
692 |
case 2 |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
693 |
from assms have "x mod 8 = 3 \<or> x mod 8 = 5" by auto |
76231 | 694 |
then have x': "x mod 16 = 3 \<or> x mod 16 = 5 \<or> x mod 16 = 11 \<or> x mod 16 = 13" |
78668 | 695 |
using mod_double_nat [of x 8] by auto |
69785
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
696 |
hence "[x ^ 4 = 1] (mod 16)" using assms |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
697 |
by (auto simp: cong_def simp flip: power_mod[of x]) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
698 |
hence "ord 16 x dvd 2\<^sup>2" by (simp add: ord_divides') |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
699 |
then obtain l where l: "ord 16 x = 2 ^ l" "l \<le> 2" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
700 |
by (subst (asm) divides_primepow_nat) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
701 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
702 |
have "[x ^ 2 \<noteq> 1] (mod 16)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
703 |
using x' by (auto simp: cong_def simp flip: power_mod[of x]) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
704 |
hence "\<not>ord 16 x dvd 2" by (simp add: ord_divides') |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
705 |
with l have "l = 2" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
706 |
using le_imp_power_dvd[of l 1 2] by (cases "l \<le> 1") auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
707 |
with l show ?thesis by (simp add: \<open>k = 4\<close>) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
708 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
709 |
case 3 |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
710 |
define k' where "k' = k - 2" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
711 |
have k': "k' \<ge> 2" and [simp]: "k = Suc (Suc k')" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
712 |
using 3 by (simp_all add: k'_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
713 |
have IH: "ord (2 ^ k') x = 2 ^ (k' - 2)" "ord (2 ^ Suc k') x = 2 ^ (k' - 1)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
714 |
using less.IH[of k'] less.IH[of "Suc k'"] 3 by simp_all |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
715 |
from IH have cong: "[x ^ (2 ^ (k' - 2)) = 1] (mod (2 ^ k'))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
716 |
by (simp_all add: ord_divides') |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
717 |
have notcong: "[x ^ (2 ^ (k' - 2)) \<noteq> 1] (mod (2 ^ Suc k'))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
718 |
proof |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
719 |
assume "[x ^ (2 ^ (k' - 2)) = 1] (mod (2 ^ Suc k'))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
720 |
hence "ord (2 ^ Suc k') x dvd 2 ^ (k' - 2)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
721 |
by (simp add: ord_divides') |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
722 |
also have "ord (2 ^ Suc k') x = 2 ^ (k' - 1)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
723 |
using IH by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
724 |
finally have "k' - 1 \<le> k' - 2" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
725 |
by (rule power_dvd_imp_le) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
726 |
with \<open>k' \<ge> 2\<close> show False by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
727 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
728 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
729 |
have "2 ^ k' + 1 < 2 ^ k' + (2 ^ k' :: nat)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
730 |
using one_less_power[of "2::nat" k'] k' by (intro add_strict_left_mono) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
731 |
with cong notcong have cong': "x ^ (2 ^ (k' - 2)) mod 2 ^ Suc k' = 1 + 2 ^ k'" |
78668 | 732 |
using mod_double_nat [of \<open>x ^ 2 ^ (k' - 2)\<close> \<open>2 ^ k'\<close>] k' by (auto simp: cong_def) |
69785
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
733 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
734 |
hence "x ^ (2 ^ (k' - 2)) mod 2 ^ k = 1 + 2 ^ k' \<or> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
735 |
x ^ (2 ^ (k' - 2)) mod 2 ^ k = 1 + 2 ^ k' + 2 ^ Suc k'" |
78668 | 736 |
using mod_double_nat [of \<open>x ^ 2 ^ (k' - 2)\<close> \<open>2 ^ Suc k'\<close>] by auto |
69785
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
737 |
hence eq: "[x ^ 2 ^ (k' - 1) = 1 + 2 ^ (k - 1)] (mod 2 ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
738 |
proof |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
739 |
assume *: "x ^ (2 ^ (k' - 2)) mod (2 ^ k) = 1 + 2 ^ k'" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
740 |
have "[x ^ (2 ^ (k' - 2)) = x ^ (2 ^ (k' - 2)) mod 2 ^ k] (mod 2 ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
741 |
by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
742 |
also have "[x ^ (2 ^ (k' - 2)) mod (2 ^ k) = 1 + 2 ^ k'] (mod 2 ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
743 |
by (subst *) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
744 |
finally have "[(x ^ 2 ^ (k' - 2)) ^ 2 = (1 + 2 ^ k') ^ 2] (mod 2 ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
745 |
by (rule cong_pow) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
746 |
hence "[x ^ 2 ^ Suc (k' - 2) = (1 + 2 ^ k') ^ 2] (mod 2 ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
747 |
by (simp add: power_mult [symmetric] power_Suc2 [symmetric] del: power_Suc) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
748 |
also have "Suc (k' - 2) = k' - 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
749 |
using k' by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
750 |
also have "(1 + 2 ^ k' :: nat)\<^sup>2 = 1 + 2 ^ (k - 1) + 2 ^ (2 * k')" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
751 |
by (subst power2_eq_square) (simp add: algebra_simps flip: power_add) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
752 |
also have "(2 ^ k :: nat) dvd 2 ^ (2 * k')" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
753 |
using k' by (intro le_imp_power_dvd) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
754 |
hence "[1 + 2 ^ (k - 1) + 2 ^ (2 * k') = 1 + 2 ^ (k - 1) + (0 :: nat)] (mod 2 ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
755 |
by (intro cong_add) (auto simp: cong_0_iff) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
756 |
finally show "[x ^ 2 ^ (k' - 1) = 1 + 2 ^ (k - 1)] (mod 2 ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
757 |
by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
758 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
759 |
assume *: "x ^ (2 ^ (k' - 2)) mod 2 ^ k = 1 + 2 ^ k' + 2 ^ Suc k'" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
760 |
have "[x ^ (2 ^ (k' - 2)) = x ^ (2 ^ (k' - 2)) mod 2 ^ k] (mod 2 ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
761 |
by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
762 |
also have "[x ^ (2 ^ (k' - 2)) mod (2 ^ k) = 1 + 3 * 2 ^ k'] (mod 2 ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
763 |
by (subst *) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
764 |
finally have "[(x ^ 2 ^ (k' - 2)) ^ 2 = (1 + 3 * 2 ^ k') ^ 2] (mod 2 ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
765 |
by (rule cong_pow) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
766 |
hence "[x ^ 2 ^ Suc (k' - 2) = (1 + 3 * 2 ^ k') ^ 2] (mod 2 ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
767 |
by (simp add: power_mult [symmetric] power_Suc2 [symmetric] del: power_Suc) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
768 |
also have "Suc (k' - 2) = k' - 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
769 |
using k' by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
770 |
also have "(1 + 3 * 2 ^ k' :: nat)\<^sup>2 = 1 + 2 ^ (k - 1) + 2 ^ k + 9 * 2 ^ (2 * k')" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
771 |
by (subst power2_eq_square) (simp add: algebra_simps flip: power_add) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
772 |
also have "(2 ^ k :: nat) dvd 9 * 2 ^ (2 * k')" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
773 |
using k' by (intro dvd_mult le_imp_power_dvd) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
774 |
hence "[1 + 2 ^ (k - 1) + 2 ^ k + 9 * 2 ^ (2 * k') = 1 + 2 ^ (k - 1) + 0 + (0 :: nat)] |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
775 |
(mod 2 ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
776 |
by (intro cong_add) (auto simp: cong_0_iff) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
777 |
finally show "[x ^ 2 ^ (k' - 1) = 1 + 2 ^ (k - 1)] (mod 2 ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
778 |
by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
779 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
780 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
781 |
have notcong': "[x ^ 2 ^ (k - 3) \<noteq> 1] (mod 2 ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
782 |
proof |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
783 |
assume "[x ^ 2 ^ (k - 3) = 1] (mod 2 ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
784 |
hence "[x ^ 2 ^ (k' - 1) - x ^ 2 ^ (k' - 1) = 1 + 2 ^ (k - 1) - 1] (mod 2 ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
785 |
by (intro cong_diff_nat eq) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
786 |
hence "[2 ^ (k - 1) = (0 :: nat)] (mod 2 ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
787 |
by (simp add: cong_sym_eq) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
788 |
hence "2 ^ k dvd 2 ^ (k - 1)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
789 |
by (simp add: cong_0_iff) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
790 |
hence "k \<le> k - 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
791 |
by (rule power_dvd_imp_le) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
792 |
thus False by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
793 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
794 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
795 |
have "[x ^ 2 ^ (k - 2) = 1] (mod 2 ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
796 |
using ord_twopow_aux[of k x] \<open>odd x\<close> \<open>k \<ge> 3\<close> by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
797 |
hence "ord (2 ^ k) x dvd 2 ^ (k - 2)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
798 |
by (simp add: ord_divides') |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
799 |
then obtain l where l: "l \<le> k - 2" "ord (2 ^ k) x = 2 ^ l" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
800 |
using divides_primepow_nat[of 2 "ord (2 ^ k) x" "k - 2"] by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
801 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
802 |
from notcong' have "\<not>ord (2 ^ k) x dvd 2 ^ (k - 3)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
803 |
by (simp add: ord_divides') |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
804 |
with l have "l = k - 2" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
805 |
using le_imp_power_dvd[of l "k - 3" 2] by (cases "l \<le> k - 3") auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
806 |
with l show ?thesis by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
807 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
808 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
809 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
810 |
lemma ord_4_3 [simp]: "ord 4 (3::nat) = 2" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
811 |
proof - |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
812 |
have "[3 ^ 2 = (1 :: nat)] (mod 4)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
813 |
by (simp add: cong_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
814 |
hence "ord 4 (3::nat) dvd 2" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
815 |
by (subst (asm) ord_divides) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
816 |
hence "ord 4 (3::nat) \<le> 2" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
817 |
by (intro dvd_imp_le) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
818 |
moreover have "ord 4 (3::nat) \<noteq> 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
819 |
by (auto simp: ord_eq_Suc_0_iff cong_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
820 |
moreover have "ord 4 (3::nat) \<noteq> 0" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
821 |
by (auto simp: gcd_non_0_nat coprime_iff_gcd_eq_1) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
822 |
ultimately show "ord 4 (3 :: nat) = 2" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
823 |
by linarith |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
824 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
825 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
826 |
lemma elements_with_ord_1: "n > 0 \<Longrightarrow> {x\<in>totatives n. ord n x = Suc 0} = {1}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
827 |
by (auto simp: ord_eq_Suc_0_iff cong_def totatives_less) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
828 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
829 |
lemma residue_prime_has_primroot: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
830 |
fixes p :: nat |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
831 |
assumes "prime p" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
832 |
shows "\<exists>a\<in>totatives p. ord p a = p - 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
833 |
proof - |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
834 |
from residue_prime_mult_group_has_gen[OF assms] |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
835 |
obtain a where a: "a \<in> {1..p-1}" "{1..p-1} = {a ^ i mod p |i. i \<in> UNIV}" by blast |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
836 |
from a have "coprime p a" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
837 |
using a assms by (intro prime_imp_coprime) (auto dest: dvd_imp_le) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
838 |
with a(1) have "a \<in> totatives p" by (auto simp: totatives_def coprime_commute) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
839 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
840 |
have "p - 1 = card {1..p-1}" by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
841 |
also have "{1..p-1} = {a ^ i mod p |i. i \<in> UNIV}" by fact |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
842 |
also have "{a ^ i mod p |i. i \<in> UNIV} = (\<lambda>i. a ^ i mod p) ` {..<ord p a}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
843 |
proof (intro equalityI subsetI) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
844 |
fix x assume "x \<in> {a ^ i mod p |i. i \<in> UNIV}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
845 |
then obtain i where [simp]: "x = a ^ i mod p" by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
846 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
847 |
have "[a ^ i = a ^ (i mod ord p a)] (mod p)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
848 |
using \<open>coprime p a\<close> by (subst order_divides_expdiff) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
849 |
hence "\<exists>j. a ^ i mod p = a ^ j mod p \<and> j < ord p a" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
850 |
using \<open>coprime p a\<close> by (intro exI[of _ "i mod ord p a"]) (auto simp: cong_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
851 |
thus "x \<in> (\<lambda>i. a ^ i mod p) ` {..<ord p a}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
852 |
by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
853 |
qed auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
854 |
also have "card \<dots> = ord p a" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
855 |
using inj_power_mod[OF \<open>coprime p a\<close>] by (subst card_image) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
856 |
finally show ?thesis using \<open>a \<in> totatives p\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
857 |
by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
858 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
859 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69064
diff
changeset
|
860 |
|
66305 | 861 |
|
862 |
subsection \<open>Another trivial primality characterization\<close> |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
863 |
|
66305 | 864 |
lemma prime_prime_factor: "prime n \<longleftrightarrow> n \<noteq> 1 \<and> (\<forall>p. prime p \<and> p dvd n \<longrightarrow> p = n)" |
55337
5d45fb978d5a
Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents:
55321
diff
changeset
|
865 |
(is "?lhs \<longleftrightarrow> ?rhs") |
66305 | 866 |
for n :: nat |
867 |
proof (cases "n = 0 \<or> n = 1") |
|
55337
5d45fb978d5a
Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents:
55321
diff
changeset
|
868 |
case True |
5d45fb978d5a
Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents:
55321
diff
changeset
|
869 |
then show ?thesis |
63633 | 870 |
by (metis bigger_prime dvd_0_right not_prime_1 not_prime_0) |
55337
5d45fb978d5a
Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents:
55321
diff
changeset
|
871 |
next |
5d45fb978d5a
Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents:
55321
diff
changeset
|
872 |
case False |
5d45fb978d5a
Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents:
55321
diff
changeset
|
873 |
show ?thesis |
5d45fb978d5a
Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents:
55321
diff
changeset
|
874 |
proof |
5d45fb978d5a
Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents:
55321
diff
changeset
|
875 |
assume "prime n" |
5d45fb978d5a
Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents:
55321
diff
changeset
|
876 |
then show ?rhs |
66305 | 877 |
by (metis not_prime_1 prime_nat_iff) |
55337
5d45fb978d5a
Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents:
55321
diff
changeset
|
878 |
next |
5d45fb978d5a
Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents:
55321
diff
changeset
|
879 |
assume ?rhs |
5d45fb978d5a
Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents:
55321
diff
changeset
|
880 |
with False show "prime n" |
63633 | 881 |
by (auto simp: prime_nat_iff) (metis One_nat_def prime_factor_nat prime_nat_iff) |
55337
5d45fb978d5a
Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents:
55321
diff
changeset
|
882 |
qed |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
883 |
qed |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
884 |
|
66305 | 885 |
lemma prime_divisor_sqrt: "prime n \<longleftrightarrow> n \<noteq> 1 \<and> (\<forall>d. d dvd n \<and> d\<^sup>2 \<le> n \<longrightarrow> d = 1)" |
886 |
for n :: nat |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
887 |
proof - |
66305 | 888 |
consider "n = 0" | "n = 1" | "n \<noteq> 0" "n \<noteq> 1" by blast |
889 |
then show ?thesis |
|
890 |
proof cases |
|
891 |
case 1 |
|
892 |
then show ?thesis by simp |
|
893 |
next |
|
894 |
case 2 |
|
895 |
then show ?thesis by simp |
|
896 |
next |
|
897 |
case n: 3 |
|
898 |
then have np: "n > 1" by arith |
|
899 |
{ |
|
900 |
fix d |
|
901 |
assume d: "d dvd n" "d\<^sup>2 \<le> n" |
|
902 |
and H: "\<forall>m. m dvd n \<longrightarrow> m = 1 \<or> m = n" |
|
903 |
from H d have d1n: "d = 1 \<or> d = n" by blast |
|
904 |
then have "d = 1" |
|
905 |
proof |
|
906 |
assume dn: "d = n" |
|
907 |
from n have "n\<^sup>2 > n * 1" |
|
908 |
by (simp add: power2_eq_square) |
|
909 |
with dn d(2) show ?thesis by simp |
|
910 |
qed |
|
911 |
} |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
912 |
moreover |
66305 | 913 |
{ |
914 |
fix d assume d: "d dvd n" and H: "\<forall>d'. d' dvd n \<and> d'\<^sup>2 \<le> n \<longrightarrow> d' = 1" |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
915 |
from d n have "d \<noteq> 0" |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
916 |
by (metis dvd_0_left_iff) |
66305 | 917 |
then have dp: "d > 0" by simp |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
918 |
from d[unfolded dvd_def] obtain e where e: "n= d*e" by blast |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
919 |
from n dp e have ep:"e > 0" by simp |
66305 | 920 |
from dp ep have "d\<^sup>2 \<le> n \<or> e\<^sup>2 \<le> n" |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
921 |
by (auto simp add: e power2_eq_square mult_le_cancel_left) |
66305 | 922 |
then have "d = 1 \<or> d = n" |
923 |
proof |
|
924 |
assume "d\<^sup>2 \<le> n" |
|
925 |
with H[rule_format, of d] d have "d = 1" by blast |
|
926 |
then show ?thesis .. |
|
927 |
next |
|
928 |
assume h: "e\<^sup>2 \<le> n" |
|
929 |
from e have "e dvd n" by (simp add: dvd_def mult.commute) |
|
930 |
with H[rule_format, of e] h have "e = 1" by simp |
|
931 |
with e have "d = n" by simp |
|
932 |
then show ?thesis .. |
|
933 |
qed |
|
934 |
} |
|
935 |
ultimately show ?thesis |
|
936 |
unfolding prime_nat_iff using np n(2) by blast |
|
937 |
qed |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
938 |
qed |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
939 |
|
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
940 |
lemma prime_prime_factor_sqrt: |
66305 | 941 |
"prime (n::nat) \<longleftrightarrow> n \<noteq> 0 \<and> n \<noteq> 1 \<and> (\<nexists>p. prime p \<and> p dvd n \<and> p\<^sup>2 \<le> n)" |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
942 |
(is "?lhs \<longleftrightarrow>?rhs") |
66305 | 943 |
proof - |
944 |
consider "n = 0" | "n = 1" | "n \<noteq> 0" "n \<noteq> 1" |
|
945 |
by blast |
|
946 |
then show ?thesis |
|
947 |
proof cases |
|
948 |
case 1 |
|
949 |
then show ?thesis by (metis not_prime_0) |
|
950 |
next |
|
951 |
case 2 |
|
952 |
then show ?thesis by (metis not_prime_1) |
|
953 |
next |
|
954 |
case n: 3 |
|
955 |
show ?thesis |
|
956 |
proof |
|
957 |
assume ?lhs |
|
958 |
from this[unfolded prime_divisor_sqrt] n show ?rhs |
|
959 |
by (metis prime_prime_factor) |
|
960 |
next |
|
961 |
assume ?rhs |
|
962 |
{ |
|
963 |
fix d |
|
964 |
assume d: "d dvd n" "d\<^sup>2 \<le> n" "d \<noteq> 1" |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
965 |
then obtain p where p: "prime p" "p dvd d" |
66305 | 966 |
by (metis prime_factor_nat) |
55337
5d45fb978d5a
Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents:
55321
diff
changeset
|
967 |
from d(1) n have dp: "d > 0" |
66305 | 968 |
by (metis dvd_0_left neq0_conv) |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
969 |
from mult_mono[OF dvd_imp_le[OF p(2) dp] dvd_imp_le[OF p(2) dp]] d(2) |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
970 |
have "p\<^sup>2 \<le> n" unfolding power2_eq_square by arith |
66305 | 971 |
with \<open>?rhs\<close> n p(1) dvd_trans[OF p(2) d(1)] have False |
972 |
by blast |
|
973 |
} |
|
974 |
with n prime_divisor_sqrt show ?lhs by auto |
|
975 |
qed |
|
976 |
qed |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
977 |
qed |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
978 |
|
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
979 |
|
66305 | 980 |
subsection \<open>Pocklington theorem\<close> |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
981 |
|
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
982 |
lemma pocklington_lemma: |
66305 | 983 |
fixes p :: nat |
984 |
assumes n: "n \<ge> 2" and nqr: "n - 1 = q * r" |
|
985 |
and an: "[a^ (n - 1) = 1] (mod n)" |
|
986 |
and aq: "\<forall>p. prime p \<and> p dvd q \<longrightarrow> coprime (a ^ ((n - 1) div p) - 1) n" |
|
987 |
and pp: "prime p" and pn: "p dvd n" |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
988 |
shows "[p = 1] (mod q)" |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
989 |
proof - |
66305 | 990 |
have p01: "p \<noteq> 0" "p \<noteq> 1" |
991 |
using pp by (auto intro: prime_gt_0_nat) |
|
992 |
obtain k where k: "a ^ (q * r) - 1 = n * k" |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
993 |
by (metis an cong_to_1_nat dvd_def nqr) |
66305 | 994 |
from pn[unfolded dvd_def] obtain l where l: "n = p * l" |
995 |
by blast |
|
996 |
have a0: "a \<noteq> 0" |
|
997 |
proof |
|
998 |
assume "a = 0" |
|
999 |
with n have "a^ (n - 1) = 0" |
|
1000 |
by (simp add: power_0_left) |
|
1001 |
with n an mod_less[of 1 n] show False |
|
66888 | 1002 |
by (simp add: power_0_left cong_def) |
66305 | 1003 |
qed |
1004 |
with n nqr have aqr0: "a ^ (q * r) \<noteq> 0" |
|
1005 |
by simp |
|
1006 |
then have "(a ^ (q * r) - 1) + 1 = a ^ (q * r)" |
|
1007 |
by simp |
|
1008 |
with k l have "a ^ (q * r) = p * l * k + 1" |
|
1009 |
by simp |
|
1010 |
then have "a ^ (r * q) + p * 0 = 1 + p * (l * k)" |
|
1011 |
by (simp add: ac_simps) |
|
1012 |
then have odq: "ord p (a^r) dvd q" |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1013 |
unfolding ord_divides[symmetric] power_mult[symmetric] |
66305 | 1014 |
by (metis an cong_dvd_modulus_nat mult.commute nqr pn) |
1015 |
from odq[unfolded dvd_def] obtain d where d: "q = ord p (a^r) * d" |
|
1016 |
by blast |
|
1017 |
have d1: "d = 1" |
|
1018 |
proof (rule ccontr) |
|
1019 |
assume d1: "d \<noteq> 1" |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1020 |
obtain P where P: "prime P" "P dvd d" |
66305 | 1021 |
by (metis d1 prime_factor_nat) |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1022 |
from d dvd_mult[OF P(2), of "ord p (a^r)"] have Pq: "P dvd q" by simp |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1023 |
from aq P(1) Pq have caP:"coprime (a^ ((n - 1) div P) - 1) n" by blast |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1024 |
from Pq obtain s where s: "q = P*s" unfolding dvd_def by blast |
66305 | 1025 |
from P(1) have P0: "P \<noteq> 0" |
1026 |
by (metis not_prime_0) |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1027 |
from P(2) obtain t where t: "d = P*t" unfolding dvd_def by blast |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1028 |
from d s t P0 have s': "ord p (a^r) * t = s" |
66305 | 1029 |
by (metis mult.commute mult_cancel1 mult.assoc) |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1030 |
have "ord p (a^r) * t*r = r * ord p (a^r) * t" |
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
55370
diff
changeset
|
1031 |
by (metis mult.assoc mult.commute) |
66305 | 1032 |
then have exps: "a^(ord p (a^r) * t*r) = ((a ^ r) ^ ord p (a^r)) ^ t" |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1033 |
by (simp only: power_mult) |
66305 | 1034 |
then have "[((a ^ r) ^ ord p (a^r)) ^ t= 1] (mod p)" |
66888 | 1035 |
by (metis cong_pow ord power_one) |
66305 | 1036 |
then have pd0: "p dvd a^(ord p (a^r) * t*r) - 1" |
1037 |
by (metis cong_to_1_nat exps) |
|
1038 |
from nqr s s' have "(n - 1) div P = ord p (a^r) * t*r" |
|
1039 |
using P0 by simp |
|
67051 | 1040 |
with caP have "coprime (a ^ (ord p (a ^ r) * t * r) - 1) n" |
1041 |
by simp |
|
1042 |
with p01 pn pd0 coprime_common_divisor [of _ n p] show False |
|
66305 | 1043 |
by auto |
1044 |
qed |
|
1045 |
with d have o: "ord p (a^r) = q" by simp |
|
1046 |
from pp totient_prime [of p] have totient_eq: "totient p = p - 1" |
|
1047 |
by simp |
|
1048 |
{ |
|
1049 |
fix d |
|
1050 |
assume d: "d dvd p" "d dvd a" "d \<noteq> 1" |
|
63633 | 1051 |
from pp[unfolded prime_nat_iff] d have dp: "d = p" by blast |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1052 |
from n have "n \<noteq> 0" by simp |
67051 | 1053 |
then have False using d dp pn an |
1054 |
by auto (metis One_nat_def Suc_lessI |
|
1055 |
\<open>1 < p \<and> (\<forall>m. m dvd p \<longrightarrow> m = 1 \<or> m = p)\<close> \<open>a ^ (q * r) = p * l * k + 1\<close> add_diff_cancel_left' dvd_diff_nat dvd_power dvd_triv_left gcd_nat.trans nat_dvd_not_less nqr zero_less_diff zero_less_one) |
|
66305 | 1056 |
} |
67051 | 1057 |
then have cpa: "coprime p a" |
1058 |
by (auto intro: coprimeI) |
|
1059 |
then have arp: "coprime (a ^ r) p" |
|
1060 |
by (cases "r > 0") (simp_all add: ac_simps) |
|
66305 | 1061 |
from euler_theorem [OF arp, simplified ord_divides] o totient_eq have "q dvd (p - 1)" |
1062 |
by simp |
|
1063 |
then obtain d where d:"p - 1 = q * d" |
|
55337
5d45fb978d5a
Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents:
55321
diff
changeset
|
1064 |
unfolding dvd_def by blast |
66305 | 1065 |
have "p \<noteq> 0" |
1066 |
by (metis p01(1)) |
|
1067 |
with d have "p + q * 0 = 1 + q * d" by simp |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1068 |
then show ?thesis |
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
55370
diff
changeset
|
1069 |
by (metis cong_iff_lin_nat mult.commute) |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1070 |
qed |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1071 |
|
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1072 |
theorem pocklington: |
66305 | 1073 |
assumes n: "n \<ge> 2" and nqr: "n - 1 = q * r" and sqr: "n \<le> q\<^sup>2" |
1074 |
and an: "[a^ (n - 1) = 1] (mod n)" |
|
1075 |
and aq: "\<forall>p. prime p \<and> p dvd q \<longrightarrow> coprime (a^ ((n - 1) div p) - 1) n" |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1076 |
shows "prime n" |
66305 | 1077 |
unfolding prime_prime_factor_sqrt[of n] |
1078 |
proof - |
|
1079 |
let ?ths = "n \<noteq> 0 \<and> n \<noteq> 1 \<and> (\<nexists>p. prime p \<and> p dvd n \<and> p\<^sup>2 \<le> n)" |
|
1080 |
from n have n01: "n \<noteq> 0" "n \<noteq> 1" by arith+ |
|
1081 |
{ |
|
1082 |
fix p |
|
1083 |
assume p: "prime p" "p dvd n" "p\<^sup>2 \<le> n" |
|
1084 |
from p(3) sqr have "p^(Suc 1) \<le> q^(Suc 1)" |
|
1085 |
by (simp add: power2_eq_square) |
|
1086 |
then have pq: "p \<le> q" |
|
1087 |
by (metis le0 power_le_imp_le_base) |
|
1088 |
from pocklington_lemma[OF n nqr an aq p(1,2)] have *: "q dvd p - 1" |
|
1089 |
by (metis cong_to_1_nat) |
|
1090 |
have "p - 1 \<noteq> 0" |
|
1091 |
using prime_ge_2_nat [OF p(1)] by arith |
|
1092 |
with pq * have False |
|
1093 |
by (simp add: nat_dvd_not_less) |
|
1094 |
} |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1095 |
with n01 show ?ths by blast |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1096 |
qed |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1097 |
|
66305 | 1098 |
text \<open>Variant for application, to separate the exponentiation.\<close> |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1099 |
lemma pocklington_alt: |
66305 | 1100 |
assumes n: "n \<ge> 2" and nqr: "n - 1 = q * r" and sqr: "n \<le> q\<^sup>2" |
1101 |
and an: "[a^ (n - 1) = 1] (mod n)" |
|
1102 |
and aq: "\<forall>p. prime p \<and> p dvd q \<longrightarrow> (\<exists>b. [a^((n - 1) div p) = b] (mod n) \<and> coprime (b - 1) n)" |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1103 |
shows "prime n" |
66305 | 1104 |
proof - |
1105 |
{ |
|
1106 |
fix p |
|
1107 |
assume p: "prime p" "p dvd q" |
|
1108 |
from aq[rule_format] p obtain b where b: "[a^((n - 1) div p) = b] (mod n)" "coprime (b - 1) n" |
|
1109 |
by blast |
|
1110 |
have a0: "a \<noteq> 0" |
|
1111 |
proof |
|
1112 |
assume a0: "a = 0" |
|
1113 |
from n an have "[0 = 1] (mod n)" |
|
1114 |
unfolding a0 power_0_left by auto |
|
1115 |
then show False |
|
66888 | 1116 |
using n by (simp add: cong_def dvd_eq_mod_eq_0[symmetric]) |
66305 | 1117 |
qed |
1118 |
then have a1: "a \<ge> 1" by arith |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1119 |
from one_le_power[OF a1] have ath: "1 \<le> a ^ ((n - 1) div p)" . |
66305 | 1120 |
have b0: "b \<noteq> 0" |
1121 |
proof |
|
1122 |
assume b0: "b = 0" |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1123 |
from p(2) nqr have "(n - 1) mod p = 0" |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1124 |
by (metis mod_0 mod_mod_cancel mod_mult_self1_is_0) |
64242
93c6f0da5c70
more standardized theorem names for facts involving the div and mod identity
haftmann
parents:
63905
diff
changeset
|
1125 |
with div_mult_mod_eq[of "n - 1" p] |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1126 |
have "(n - 1) div p * p= n - 1" by auto |
66305 | 1127 |
then have eq: "(a^((n - 1) div p))^p = a^(n - 1)" |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1128 |
by (simp only: power_mult[symmetric]) |
66305 | 1129 |
have "p - 1 \<noteq> 0" |
1130 |
using prime_ge_2_nat [OF p(1)] by arith |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1131 |
then have pS: "Suc (p - 1) = p" by arith |
66305 | 1132 |
from b have d: "n dvd a^((n - 1) div p)" |
1133 |
unfolding b0 by auto |
|
66888 | 1134 |
from divides_rexp[OF d, of "p - 1"] pS eq cong_dvd_iff [OF an] n show False |
66305 | 1135 |
by simp |
1136 |
qed |
|
1137 |
then have b1: "b \<ge> 1" by arith |
|
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
68157
diff
changeset
|
1138 |
from cong_imp_coprime[OF Cong.cong_diff_nat[OF cong_sym [OF b(1)] cong_refl [of 1] b1]] |
66305 | 1139 |
ath b1 b nqr |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1140 |
have "coprime (a ^ ((n - 1) div p) - 1) n" |
66305 | 1141 |
by simp |
1142 |
} |
|
1143 |
then have "\<forall>p. prime p \<and> p dvd q \<longrightarrow> coprime (a ^ ((n - 1) div p) - 1) n " |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1144 |
by blast |
66305 | 1145 |
then show ?thesis by (rule pocklington[OF n nqr sqr an]) |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1146 |
qed |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1147 |
|
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1148 |
|
66305 | 1149 |
subsection \<open>Prime factorizations\<close> |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1150 |
|
55370 | 1151 |
(* FIXME some overlap with material in UniqueFactorization, class unique_factorization *) |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1152 |
|
69064
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68707
diff
changeset
|
1153 |
definition "primefact ps n \<longleftrightarrow> foldr (*) ps 1 = n \<and> (\<forall>p\<in> set ps. prime p)" |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1154 |
|
66305 | 1155 |
lemma primefact: |
1156 |
fixes n :: nat |
|
1157 |
assumes n: "n \<noteq> 0" |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1158 |
shows "\<exists>ps. primefact ps n" |
63534
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
62429
diff
changeset
|
1159 |
proof - |
66305 | 1160 |
obtain xs where xs: "mset xs = prime_factorization n" |
1161 |
using ex_mset [of "prime_factorization n"] by blast |
|
1162 |
from assms have "n = prod_mset (prime_factorization n)" |
|
63830 | 1163 |
by (simp add: prod_mset_prime_factorization) |
1164 |
also have "\<dots> = prod_mset (mset xs)" by (simp add: xs) |
|
69064
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68707
diff
changeset
|
1165 |
also have "\<dots> = foldr (*) xs 1" by (induct xs) simp_all |
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68707
diff
changeset
|
1166 |
finally have "foldr (*) xs 1 = n" .. |
66305 | 1167 |
moreover from xs have "\<forall>p\<in>#mset xs. prime p" by auto |
63534
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
62429
diff
changeset
|
1168 |
ultimately have "primefact xs n" by (auto simp: primefact_def) |
66305 | 1169 |
then show ?thesis .. |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1170 |
qed |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1171 |
|
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1172 |
lemma primefact_contains: |
66305 | 1173 |
fixes p :: nat |
1174 |
assumes pf: "primefact ps n" |
|
1175 |
and p: "prime p" |
|
1176 |
and pn: "p dvd n" |
|
1177 |
shows "p \<in> set ps" |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1178 |
using pf p pn |
66305 | 1179 |
proof (induct ps arbitrary: p n) |
1180 |
case Nil |
|
1181 |
then show ?case by (auto simp: primefact_def) |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1182 |
next |
66305 | 1183 |
case (Cons q qs) |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1184 |
from Cons.prems[unfolded primefact_def] |
69064
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68707
diff
changeset
|
1185 |
have q: "prime q" "q * foldr (*) qs 1 = n" "\<forall>p \<in>set qs. prime p" |
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68707
diff
changeset
|
1186 |
and p: "prime p" "p dvd q * foldr (*) qs 1" |
66305 | 1187 |
by simp_all |
69064
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68707
diff
changeset
|
1188 |
consider "p dvd q" | "p dvd foldr (*) qs 1" |
66305 | 1189 |
by (metis p prime_dvd_mult_eq_nat) |
1190 |
then show ?case |
|
1191 |
proof cases |
|
1192 |
case 1 |
|
1193 |
with p(1) q(1) have "p = q" |
|
1194 |
unfolding prime_nat_iff by auto |
|
1195 |
then show ?thesis by simp |
|
1196 |
next |
|
1197 |
case prem: 2 |
|
69064
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68707
diff
changeset
|
1198 |
from q(3) have pqs: "primefact qs (foldr (*) qs 1)" |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1199 |
by (simp add: primefact_def) |
66305 | 1200 |
from Cons.hyps[OF pqs p(1) prem] show ?thesis by simp |
1201 |
qed |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1202 |
qed |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1203 |
|
69064
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68707
diff
changeset
|
1204 |
lemma primefact_variant: "primefact ps n \<longleftrightarrow> foldr (*) ps 1 = n \<and> list_all prime ps" |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1205 |
by (auto simp add: primefact_def list_all_iff) |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1206 |
|
66305 | 1207 |
text \<open>Variant of Lucas theorem.\<close> |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1208 |
lemma lucas_primefact: |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1209 |
assumes n: "n \<ge> 2" and an: "[a^(n - 1) = 1] (mod n)" |
69064
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68707
diff
changeset
|
1210 |
and psn: "foldr (*) ps 1 = n - 1" |
66305 | 1211 |
and psp: "list_all (\<lambda>p. prime p \<and> \<not> [a^((n - 1) div p) = 1] (mod n)) ps" |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1212 |
shows "prime n" |
66305 | 1213 |
proof - |
1214 |
{ |
|
1215 |
fix p |
|
1216 |
assume p: "prime p" "p dvd n - 1" "[a ^ ((n - 1) div p) = 1] (mod n)" |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1217 |
from psn psp have psn1: "primefact ps (n - 1)" |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1218 |
by (auto simp add: list_all_iff primefact_variant) |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1219 |
from p(3) primefact_contains[OF psn1 p(1,2)] psp |
66305 | 1220 |
have False by (induct ps) auto |
1221 |
} |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1222 |
with lucas[OF n an] show ?thesis by blast |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1223 |
qed |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1224 |
|
66305 | 1225 |
text \<open>Variant of Pocklington theorem.\<close> |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1226 |
lemma pocklington_primefact: |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1227 |
assumes n: "n \<ge> 2" and qrn: "q*r = n - 1" and nq2: "n \<le> q\<^sup>2" |
69064
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68707
diff
changeset
|
1228 |
and arnb: "(a^r) mod n = b" and psq: "foldr (*) ps 1 = q" |
66305 | 1229 |
and bqn: "(b^q) mod n = 1" |
1230 |
and psp: "list_all (\<lambda>p. prime p \<and> coprime ((b^(q div p)) mod n - 1) n) ps" |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1231 |
shows "prime n" |
66305 | 1232 |
proof - |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1233 |
from bqn psp qrn |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1234 |
have bqn: "a ^ (n - 1) mod n = 1" |
66305 | 1235 |
and psp: "list_all (\<lambda>p. prime p \<and> coprime (a^(r *(q div p)) mod n - 1) n) ps" |
1236 |
unfolding arnb[symmetric] power_mod |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1237 |
by (simp_all add: power_mult[symmetric] algebra_simps) |
66305 | 1238 |
from n have n0: "n > 0" by arith |
64242
93c6f0da5c70
more standardized theorem names for facts involving the div and mod identity
haftmann
parents:
63905
diff
changeset
|
1239 |
from div_mult_mod_eq[of "a^(n - 1)" n] |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1240 |
mod_less_divisor[OF n0, of "a^(n - 1)"] |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1241 |
have an1: "[a ^ (n - 1) = 1] (mod n)" |
66888 | 1242 |
by (metis bqn cong_def mod_mod_trivial) |
66305 | 1243 |
have "coprime (a ^ ((n - 1) div p) - 1) n" if p: "prime p" "p dvd q" for p |
1244 |
proof - |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1245 |
from psp psq have pfpsq: "primefact ps q" |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1246 |
by (auto simp add: primefact_variant list_all_iff) |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1247 |
from psp primefact_contains[OF pfpsq p] |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1248 |
have p': "coprime (a ^ (r * (q div p)) mod n - 1) n" |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1249 |
by (simp add: list_all_iff) |
66305 | 1250 |
from p prime_nat_iff have p01: "p \<noteq> 0" "p \<noteq> 1" "p = Suc (p - 1)" |
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1251 |
by auto |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1252 |
from div_mult1_eq[of r q p] p(2) |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1253 |
have eq1: "r* (q div p) = (n - 1) div p" |
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
55370
diff
changeset
|
1254 |
unfolding qrn[symmetric] dvd_eq_mod_eq_0 by (simp add: mult.commute) |
66305 | 1255 |
have ath: "a \<le> b \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> 1 \<le> a \<and> 1 \<le> b" for a b :: nat |
1256 |
by arith |
|
1257 |
{ |
|
1258 |
assume "a ^ ((n - 1) div p) mod n = 0" |
|
1259 |
then obtain s where s: "a ^ ((n - 1) div p) = n * s" |
|
68157 | 1260 |
by blast |
66305 | 1261 |
then have eq0: "(a^((n - 1) div p))^p = (n*s)^p" by simp |
1262 |
from qrn[symmetric] have qn1: "q dvd n - 1" |
|
1263 |
by (auto simp: dvd_def) |
|
1264 |
from dvd_trans[OF p(2) qn1] have npp: "(n - 1) div p * p = n - 1" |
|
1265 |
by simp |
|
1266 |
with eq0 have "a ^ (n - 1) = (n * s) ^ p" |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1267 |
by (simp add: power_mult[symmetric]) |
66305 | 1268 |
with bqn p01 have "1 = (n * s)^(Suc (p - 1)) mod n" |
1269 |
by simp |
|
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
55370
diff
changeset
|
1270 |
also have "\<dots> = 0" by (simp add: mult.assoc) |
66305 | 1271 |
finally have False by simp |
1272 |
} |
|
1273 |
then have *: "a ^ ((n - 1) div p) mod n \<noteq> 0" by auto |
|
1274 |
have "[a ^ ((n - 1) div p) mod n = a ^ ((n - 1) div p)] (mod n)" |
|
66888 | 1275 |
by (simp add: cong_def) |
66305 | 1276 |
with ath[OF mod_less_eq_dividend *] |
1277 |
have "[a ^ ((n - 1) div p) mod n - 1 = a ^ ((n - 1) div p) - 1] (mod n)" |
|
66888 | 1278 |
by (simp add: cong_diff_nat) |
66305 | 1279 |
then show ?thesis |
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
68157
diff
changeset
|
1280 |
by (metis cong_imp_coprime eq1 p') |
66305 | 1281 |
qed |
1282 |
with pocklington[OF n qrn[symmetric] nq2 an1] show ?thesis |
|
1283 |
by blast |
|
55321
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1284 |
qed |
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1285 |
|
eadea363deb6
Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1286 |
end |