| author | wenzelm | 
| Fri, 06 Jan 2012 17:44:42 +0100 | |
| changeset 46138 | 85f8d8a8c711 | 
| parent 41541 | 1fa4725c4656 | 
| child 49962 | a8cc904a6820 | 
| permissions | -rw-r--r-- | 
| 38159 | 1  | 
(* Title: HOL/Old_Number_Theory/Pocklington.thy  | 
| 30488 | 2  | 
Author: Amine Chaieb  | 
| 26126 | 3  | 
*)  | 
4  | 
||
5  | 
header {* Pocklington's Theorem for Primes *}
 | 
|
6  | 
||
7  | 
theory Pocklington  | 
|
| 38159 | 8  | 
imports Primes  | 
| 26126 | 9  | 
begin  | 
10  | 
||
11  | 
definition modeq:: "nat => nat => nat => bool"    ("(1[_ = _] '(mod _'))")
 | 
|
12  | 
where "[a = b] (mod p) == ((a mod p) = (b mod p))"  | 
|
13  | 
||
14  | 
definition modneq:: "nat => nat => nat => bool"    ("(1[_ \<noteq> _] '(mod _'))")
 | 
|
15  | 
where "[a \<noteq> b] (mod p) == ((a mod p) \<noteq> (b mod p))"  | 
|
16  | 
||
17  | 
lemma modeq_trans:  | 
|
18  | 
"\<lbrakk> [a = b] (mod p); [b = c] (mod p) \<rbrakk> \<Longrightarrow> [a = c] (mod p)"  | 
|
19  | 
by (simp add:modeq_def)  | 
|
20  | 
||
21  | 
||
22  | 
lemma nat_mod_lemma: assumes xyn: "[x = y] (mod n)" and xy:"y \<le> x"  | 
|
23  | 
shows "\<exists>q. x = y + n * q"  | 
|
| 27668 | 24  | 
using xyn xy unfolding modeq_def using nat_mod_eq_lemma by blast  | 
| 26126 | 25  | 
|
| 30488 | 26  | 
lemma nat_mod[algebra]: "[x = y] (mod n) \<longleftrightarrow> (\<exists>q1 q2. x + n * q1 = y + n * q2)"  | 
| 27668 | 27  | 
unfolding modeq_def nat_mod_eq_iff ..  | 
| 26126 | 28  | 
|
29  | 
(* Lemmas about previously defined terms. *)  | 
|
30  | 
||
| 30488 | 31  | 
lemma prime: "prime p \<longleftrightarrow> p \<noteq> 0 \<and> p\<noteq>1 \<and> (\<forall>m. 0 < m \<and> m < p \<longrightarrow> coprime p m)"  | 
32  | 
(is "?lhs \<longleftrightarrow> ?rhs")  | 
|
| 26126 | 33  | 
proof-  | 
34  | 
  {assume "p=0 \<or> p=1" hence ?thesis using prime_0 prime_1 by (cases "p=0", simp_all)}
 | 
|
35  | 
moreover  | 
|
36  | 
  {assume p0: "p\<noteq>0" "p\<noteq>1"
 | 
|
37  | 
    {assume H: "?lhs"
 | 
|
38  | 
      {fix m assume m: "m > 0" "m < p"
 | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
39  | 
        {assume "m=1" hence "coprime p m" by simp}
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
40  | 
moreover  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
41  | 
        {assume "p dvd m" hence "p \<le> m" using dvd_imp_le m by blast with m(2)
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
42  | 
have "coprime p m" by simp}  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
43  | 
ultimately have "coprime p m" using prime_coprime[OF H, of m] by blast}  | 
| 26126 | 44  | 
hence ?rhs using p0 by auto}  | 
45  | 
moreover  | 
|
46  | 
    { assume H: "\<forall>m. 0 < m \<and> m < p \<longrightarrow> coprime p m"
 | 
|
47  | 
from prime_factor[OF p0(2)] obtain q where q: "prime q" "q dvd p" by blast  | 
|
48  | 
from prime_ge_2[OF q(1)] have q0: "q > 0" by arith  | 
|
49  | 
from dvd_imp_le[OF q(2)] p0 have qp: "q \<le> p" by arith  | 
|
50  | 
      {assume "q = p" hence ?lhs using q(1) by blast}
 | 
|
51  | 
moreover  | 
|
52  | 
      {assume "q\<noteq>p" with qp have qplt: "q < p" by arith
 | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
53  | 
from H[rule_format, of q] qplt q0 have "coprime p q" by arith  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
54  | 
with coprime_prime[of p q q] q have False by simp hence ?lhs by blast}  | 
| 26126 | 55  | 
ultimately have ?lhs by blast}  | 
56  | 
ultimately have ?thesis by blast}  | 
|
57  | 
ultimately show ?thesis by (cases"p=0 \<or> p=1", auto)  | 
|
58  | 
qed  | 
|
59  | 
||
60  | 
lemma finite_number_segment: "card { m. 0 < m \<and> m < n } = n - 1"
 | 
|
61  | 
proof-  | 
|
62  | 
  have "{ m. 0 < m \<and> m < n } = {1..<n}" by auto
 | 
|
63  | 
thus ?thesis by simp  | 
|
64  | 
qed  | 
|
65  | 
||
66  | 
lemma coprime_mod: assumes n: "n \<noteq> 0" shows "coprime (a mod n) n \<longleftrightarrow> coprime a n"  | 
|
67  | 
using n dvd_mod_iff[of _ n a] by (auto simp add: coprime)  | 
|
68  | 
||
69  | 
(* Congruences. *)  | 
|
70  | 
||
| 30488 | 71  | 
lemma cong_mod_01[simp,presburger]:  | 
| 26126 | 72  | 
"[x = y] (mod 0) \<longleftrightarrow> x = y" "[x = y] (mod 1)" "[x = 0] (mod n) \<longleftrightarrow> n dvd x"  | 
73  | 
by (simp_all add: modeq_def, presburger)  | 
|
74  | 
||
| 30488 | 75  | 
lemma cong_sub_cases:  | 
| 26126 | 76  | 
"[x = y] (mod n) \<longleftrightarrow> (if x <= y then [y - x = 0] (mod n) else [x - y = 0] (mod n))"  | 
77  | 
apply (auto simp add: nat_mod)  | 
|
78  | 
apply (rule_tac x="q2" in exI)  | 
|
79  | 
apply (rule_tac x="q1" in exI, simp)  | 
|
80  | 
apply (rule_tac x="q2" in exI)  | 
|
81  | 
apply (rule_tac x="q1" in exI, simp)  | 
|
82  | 
apply (rule_tac x="q1" in exI)  | 
|
83  | 
apply (rule_tac x="q2" in exI, simp)  | 
|
84  | 
apply (rule_tac x="q1" in exI)  | 
|
85  | 
apply (rule_tac x="q2" in exI, simp)  | 
|
86  | 
done  | 
|
87  | 
||
88  | 
lemma cong_mult_lcancel: assumes an: "coprime a n" and axy:"[a * x = a * y] (mod n)"  | 
|
89  | 
shows "[x = y] (mod n)"  | 
|
90  | 
proof-  | 
|
91  | 
  {assume "a = 0" with an axy coprime_0'[of n] have ?thesis by (simp add: modeq_def) }
 | 
|
92  | 
moreover  | 
|
93  | 
  {assume az: "a\<noteq>0"
 | 
|
94  | 
    {assume xy: "x \<le> y" hence axy': "a*x \<le> a*y" by simp
 | 
|
95  | 
with axy cong_sub_cases[of "a*x" "a*y" n] have "[a*(y - x) = 0] (mod n)"  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
96  | 
by (simp only: if_True diff_mult_distrib2)  | 
| 30488 | 97  | 
hence th: "n dvd a*(y -x)" by simp  | 
| 26126 | 98  | 
from coprime_divprod[OF th] an have "n dvd y - x"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
99  | 
by (simp add: coprime_commute)  | 
| 26126 | 100  | 
hence ?thesis using xy cong_sub_cases[of x y n] by simp}  | 
101  | 
moreover  | 
|
| 30488 | 102  | 
    {assume H: "\<not>x \<le> y" hence xy: "y \<le> x"  by arith
 | 
| 26126 | 103  | 
from H az have axy': "\<not> a*x \<le> a*y" by auto  | 
104  | 
with axy H cong_sub_cases[of "a*x" "a*y" n] have "[a*(x - y) = 0] (mod n)"  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
105  | 
by (simp only: if_False diff_mult_distrib2)  | 
| 30488 | 106  | 
hence th: "n dvd a*(x - y)" by simp  | 
| 26126 | 107  | 
from coprime_divprod[OF th] an have "n dvd x - y"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
108  | 
by (simp add: coprime_commute)  | 
| 26126 | 109  | 
hence ?thesis using xy cong_sub_cases[of x y n] by simp}  | 
110  | 
ultimately have ?thesis by blast}  | 
|
111  | 
ultimately show ?thesis by blast  | 
|
112  | 
qed  | 
|
113  | 
||
114  | 
lemma cong_mult_rcancel: assumes an: "coprime a n" and axy:"[x*a = y*a] (mod n)"  | 
|
115  | 
shows "[x = y] (mod n)"  | 
|
116  | 
using cong_mult_lcancel[OF an axy[unfolded mult_commute[of _a]]] .  | 
|
117  | 
||
118  | 
lemma cong_refl: "[x = x] (mod n)" by (simp add: modeq_def)  | 
|
119  | 
||
120  | 
lemma eq_imp_cong: "a = b \<Longrightarrow> [a = b] (mod n)" by (simp add: cong_refl)  | 
|
121  | 
||
| 30488 | 122  | 
lemma cong_commute: "[x = y] (mod n) \<longleftrightarrow> [y = x] (mod n)"  | 
| 26126 | 123  | 
by (auto simp add: modeq_def)  | 
124  | 
||
125  | 
lemma cong_trans[trans]: "[x = y] (mod n) \<Longrightarrow> [y = z] (mod n) \<Longrightarrow> [x = z] (mod n)"  | 
|
126  | 
by (simp add: modeq_def)  | 
|
127  | 
||
128  | 
lemma cong_add: assumes xx': "[x = x'] (mod n)" and yy':"[y = y'] (mod n)"  | 
|
129  | 
shows "[x + y = x' + y'] (mod n)"  | 
|
130  | 
proof-  | 
|
131  | 
have "(x + y) mod n = (x mod n + y mod n) mod n"  | 
|
132  | 
by (simp add: mod_add_left_eq[of x y n] mod_add_right_eq[of "x mod n" y n])  | 
|
| 30488 | 133  | 
also have "\<dots> = (x' mod n + y' mod n) mod n" using xx' yy' modeq_def by simp  | 
| 26126 | 134  | 
also have "\<dots> = (x' + y') mod n"  | 
135  | 
by (simp add: mod_add_left_eq[of x' y' n] mod_add_right_eq[of "x' mod n" y' n])  | 
|
| 30488 | 136  | 
finally show ?thesis unfolding modeq_def .  | 
| 26126 | 137  | 
qed  | 
138  | 
||
139  | 
lemma cong_mult: assumes xx': "[x = x'] (mod n)" and yy':"[y = y'] (mod n)"  | 
|
140  | 
shows "[x * y = x' * y'] (mod n)"  | 
|
141  | 
proof-  | 
|
| 30488 | 142  | 
have "(x * y) mod n = (x mod n) * (y mod n) mod n"  | 
| 30224 | 143  | 
by (simp add: mod_mult_left_eq[of x y n] mod_mult_right_eq[of "x mod n" y n])  | 
| 30488 | 144  | 
also have "\<dots> = (x' mod n) * (y' mod n) mod n" using xx'[unfolded modeq_def] yy'[unfolded modeq_def] by simp  | 
| 26126 | 145  | 
also have "\<dots> = (x' * y') mod n"  | 
| 30224 | 146  | 
by (simp add: mod_mult_left_eq[of x' y' n] mod_mult_right_eq[of "x' mod n" y' n])  | 
| 30488 | 147  | 
finally show ?thesis unfolding modeq_def .  | 
| 26126 | 148  | 
qed  | 
149  | 
||
150  | 
lemma cong_exp: "[x = y] (mod n) \<Longrightarrow> [x^k = y^k] (mod n)"  | 
|
151  | 
by (induct k, auto simp add: cong_refl cong_mult)  | 
|
152  | 
lemma cong_sub: assumes xx': "[x = x'] (mod n)" and yy': "[y = y'] (mod n)"  | 
|
153  | 
and yx: "y <= x" and yx': "y' <= x'"  | 
|
154  | 
shows "[x - y = x' - y'] (mod n)"  | 
|
155  | 
proof-  | 
|
| 30488 | 156  | 
  { fix x a x' a' y b y' b'
 | 
| 26126 | 157  | 
have "(x::nat) + a = x' + a' \<Longrightarrow> y + b = y' + b' \<Longrightarrow> y <= x \<Longrightarrow> y' <= x'  | 
158  | 
\<Longrightarrow> (x - y) + (a + b') = (x' - y') + (a' + b)" by arith}  | 
|
159  | 
note th = this  | 
|
| 30488 | 160  | 
from xx' yy' obtain q1 q2 q1' q2' where q12: "x + n*q1 = x'+n*q2"  | 
| 26126 | 161  | 
and q12': "y + n*q1' = y'+n*q2'" unfolding nat_mod by blast+  | 
162  | 
from th[OF q12 q12' yx yx']  | 
|
| 30488 | 163  | 
have "(x - y) + n*(q1 + q2') = (x' - y') + n*(q2 + q1')"  | 
| 26126 | 164  | 
by (simp add: right_distrib)  | 
165  | 
thus ?thesis unfolding nat_mod by blast  | 
|
166  | 
qed  | 
|
167  | 
||
| 30488 | 168  | 
lemma cong_mult_lcancel_eq: assumes an: "coprime a n"  | 
| 26126 | 169  | 
shows "[a * x = a * y] (mod n) \<longleftrightarrow> [x = y] (mod n)" (is "?lhs \<longleftrightarrow> ?rhs")  | 
170  | 
proof  | 
|
171  | 
assume H: "?rhs" from cong_mult[OF cong_refl[of a n] H] show ?lhs .  | 
|
172  | 
next  | 
|
173  | 
assume H: "?lhs" hence H': "[x*a = y*a] (mod n)" by (simp add: mult_commute)  | 
|
174  | 
from cong_mult_rcancel[OF an H'] show ?rhs .  | 
|
175  | 
qed  | 
|
176  | 
||
| 30488 | 177  | 
lemma cong_mult_rcancel_eq: assumes an: "coprime a n"  | 
| 26126 | 178  | 
shows "[x * a = y * a] (mod n) \<longleftrightarrow> [x = y] (mod n)"  | 
179  | 
using cong_mult_lcancel_eq[OF an, of x y] by (simp add: mult_commute)  | 
|
180  | 
||
| 30488 | 181  | 
lemma cong_add_lcancel_eq: "[a + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)"  | 
| 26126 | 182  | 
by (simp add: nat_mod)  | 
183  | 
||
184  | 
lemma cong_add_rcancel_eq: "[x + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)"  | 
|
185  | 
by (simp add: nat_mod)  | 
|
186  | 
||
| 30488 | 187  | 
lemma cong_add_rcancel: "[x + a = y + a] (mod n) \<Longrightarrow> [x = y] (mod n)"  | 
| 26126 | 188  | 
by (simp add: nat_mod)  | 
189  | 
||
190  | 
lemma cong_add_lcancel: "[a + x = a + y] (mod n) \<Longrightarrow> [x = y] (mod n)"  | 
|
191  | 
by (simp add: nat_mod)  | 
|
192  | 
||
| 30488 | 193  | 
lemma cong_add_lcancel_eq_0: "[a + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"  | 
| 26126 | 194  | 
by (simp add: nat_mod)  | 
195  | 
||
196  | 
lemma cong_add_rcancel_eq_0: "[x + a = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"  | 
|
197  | 
by (simp add: nat_mod)  | 
|
198  | 
||
199  | 
lemma cong_imp_eq: assumes xn: "x < n" and yn: "y < n" and xy: "[x = y] (mod n)"  | 
|
200  | 
shows "x = y"  | 
|
| 30488 | 201  | 
using xy[unfolded modeq_def mod_less[OF xn] mod_less[OF yn]] .  | 
| 26126 | 202  | 
|
203  | 
lemma cong_divides_modulus: "[x = y] (mod m) \<Longrightarrow> n dvd m ==> [x = y] (mod n)"  | 
|
204  | 
apply (auto simp add: nat_mod dvd_def)  | 
|
205  | 
apply (rule_tac x="k*q1" in exI)  | 
|
206  | 
apply (rule_tac x="k*q2" in exI)  | 
|
207  | 
by simp  | 
|
| 30488 | 208  | 
|
| 26126 | 209  | 
lemma cong_0_divides: "[x = 0] (mod n) \<longleftrightarrow> n dvd x" by simp  | 
210  | 
||
211  | 
lemma cong_1_divides:"[x = 1] (mod n) ==> n dvd x - 1"  | 
|
212  | 
apply (cases "x\<le>1", simp_all)  | 
|
213  | 
using cong_sub_cases[of x 1 n] by auto  | 
|
214  | 
||
215  | 
lemma cong_divides: "[x = y] (mod n) \<Longrightarrow> n dvd x \<longleftrightarrow> n dvd y"  | 
|
216  | 
apply (auto simp add: nat_mod dvd_def)  | 
|
217  | 
apply (rule_tac x="k + q1 - q2" in exI, simp add: add_mult_distrib2 diff_mult_distrib2)  | 
|
218  | 
apply (rule_tac x="k + q2 - q1" in exI, simp add: add_mult_distrib2 diff_mult_distrib2)  | 
|
219  | 
done  | 
|
220  | 
||
| 30488 | 221  | 
lemma cong_coprime: assumes xy: "[x = y] (mod n)"  | 
| 26126 | 222  | 
shows "coprime n x \<longleftrightarrow> coprime n y"  | 
223  | 
proof-  | 
|
224  | 
  {assume "n=0" hence ?thesis using xy by simp}
 | 
|
225  | 
moreover  | 
|
226  | 
  {assume nz: "n \<noteq> 0"
 | 
|
| 30488 | 227  | 
have "coprime n x \<longleftrightarrow> coprime (x mod n) n"  | 
| 26126 | 228  | 
by (simp add: coprime_mod[OF nz, of x] coprime_commute[of n x])  | 
229  | 
also have "\<dots> \<longleftrightarrow> coprime (y mod n) n" using xy[unfolded modeq_def] by simp  | 
|
230  | 
also have "\<dots> \<longleftrightarrow> coprime y n" by (simp add: coprime_mod[OF nz, of y])  | 
|
231  | 
finally have ?thesis by (simp add: coprime_commute) }  | 
|
232  | 
ultimately show ?thesis by blast  | 
|
233  | 
qed  | 
|
234  | 
||
235  | 
lemma cong_mod: "~(n = 0) \<Longrightarrow> [a mod n = a] (mod n)" by (simp add: modeq_def)  | 
|
236  | 
||
| 30488 | 237  | 
lemma mod_mult_cong: "~(a = 0) \<Longrightarrow> ~(b = 0)  | 
| 26126 | 238  | 
\<Longrightarrow> [x mod (a * b) = y] (mod a) \<longleftrightarrow> [x = y] (mod a)"  | 
239  | 
by (simp add: modeq_def mod_mult2_eq mod_add_left_eq)  | 
|
240  | 
||
241  | 
lemma cong_mod_mult: "[x = y] (mod n) \<Longrightarrow> m dvd n \<Longrightarrow> [x = y] (mod m)"  | 
|
242  | 
apply (auto simp add: nat_mod dvd_def)  | 
|
243  | 
apply (rule_tac x="k*q1" in exI)  | 
|
244  | 
apply (rule_tac x="k*q2" in exI, simp)  | 
|
245  | 
done  | 
|
246  | 
||
247  | 
(* Some things when we know more about the order. *)  | 
|
248  | 
||
249  | 
lemma cong_le: "y <= x \<Longrightarrow> [x = y] (mod n) \<longleftrightarrow> (\<exists>q. x = q * n + y)"  | 
|
250  | 
using nat_mod_lemma[of x y n]  | 
|
251  | 
apply auto  | 
|
252  | 
apply (simp add: nat_mod)  | 
|
253  | 
apply (rule_tac x="q" in exI)  | 
|
254  | 
apply (rule_tac x="q + q" in exI)  | 
|
| 29667 | 255  | 
by (auto simp: algebra_simps)  | 
| 26126 | 256  | 
|
257  | 
lemma cong_to_1: "[a = 1] (mod n) \<longleftrightarrow> a = 0 \<and> n = 1 \<or> (\<exists>m. a = 1 + m * n)"  | 
|
258  | 
proof-  | 
|
| 30488 | 259  | 
  {assume "n = 0 \<or> n = 1\<or> a = 0 \<or> a = 1" hence ?thesis
 | 
| 26126 | 260  | 
apply (cases "n=0", simp_all add: cong_commute)  | 
261  | 
apply (cases "n=1", simp_all add: cong_commute modeq_def)  | 
|
| 30488 | 262  | 
apply arith  | 
| 41541 | 263  | 
apply (cases "a=1")  | 
264  | 
apply (simp_all add: modeq_def cong_commute)  | 
|
265  | 
done }  | 
|
| 26126 | 266  | 
moreover  | 
267  | 
  {assume n: "n\<noteq>0" "n\<noteq>1" and a:"a\<noteq>0" "a \<noteq> 1" hence a': "a \<ge> 1" by simp
 | 
|
268  | 
hence ?thesis using cong_le[OF a', of n] by auto }  | 
|
269  | 
ultimately show ?thesis by auto  | 
|
270  | 
qed  | 
|
271  | 
||
272  | 
(* Some basic theorems about solving congruences. *)  | 
|
273  | 
||
274  | 
||
275  | 
lemma cong_solve: assumes an: "coprime a n" shows "\<exists>x. [a * x = b] (mod n)"  | 
|
276  | 
proof-  | 
|
277  | 
  {assume "a=0" hence ?thesis using an by (simp add: modeq_def)}
 | 
|
278  | 
moreover  | 
|
279  | 
  {assume az: "a\<noteq>0"
 | 
|
| 30488 | 280  | 
from bezout_add_strong[OF az, of n]  | 
| 26126 | 281  | 
obtain d x y where dxy: "d dvd a" "d dvd n" "a*x = n*y + d" by blast  | 
282  | 
from an[unfolded coprime, rule_format, of d] dxy(1,2) have d1: "d = 1" by blast  | 
|
283  | 
hence "a*x*b = (n*y + 1)*b" using dxy(3) by simp  | 
|
284  | 
hence "a*(x*b) = n*(y*b) + b" by algebra  | 
|
285  | 
hence "a*(x*b) mod n = (n*(y*b) + b) mod n" by simp  | 
|
286  | 
hence "a*(x*b) mod n = b mod n" by (simp add: mod_add_left_eq)  | 
|
287  | 
hence "[a*(x*b) = b] (mod n)" unfolding modeq_def .  | 
|
288  | 
hence ?thesis by blast}  | 
|
289  | 
ultimately show ?thesis by blast  | 
|
290  | 
qed  | 
|
291  | 
||
292  | 
lemma cong_solve_unique: assumes an: "coprime a n" and nz: "n \<noteq> 0"  | 
|
293  | 
shows "\<exists>!x. x < n \<and> [a * x = b] (mod n)"  | 
|
294  | 
proof-  | 
|
295  | 
let ?P = "\<lambda>x. x < n \<and> [a * x = b] (mod n)"  | 
|
296  | 
from cong_solve[OF an] obtain x where x: "[a*x = b] (mod n)" by blast  | 
|
297  | 
let ?x = "x mod n"  | 
|
298  | 
from x have th: "[a * ?x = b] (mod n)"  | 
|
| 30224 | 299  | 
by (simp add: modeq_def mod_mult_right_eq[of a x n])  | 
| 26126 | 300  | 
from mod_less_divisor[ of n x] nz th have Px: "?P ?x" by simp  | 
301  | 
  {fix y assume Py: "y < n" "[a * y = b] (mod n)"
 | 
|
302  | 
from Py(2) th have "[a * y = a*?x] (mod n)" by (simp add: modeq_def)  | 
|
303  | 
hence "[y = ?x] (mod n)" by (simp add: cong_mult_lcancel_eq[OF an])  | 
|
304  | 
with mod_less[OF Py(1)] mod_less_divisor[ of n x] nz  | 
|
305  | 
have "y = ?x" by (simp add: modeq_def)}  | 
|
306  | 
with Px show ?thesis by blast  | 
|
307  | 
qed  | 
|
308  | 
||
309  | 
lemma cong_solve_unique_nontrivial:  | 
|
310  | 
assumes p: "prime p" and pa: "coprime p a" and x0: "0 < x" and xp: "x < p"  | 
|
311  | 
shows "\<exists>!y. 0 < y \<and> y < p \<and> [x * y = a] (mod p)"  | 
|
312  | 
proof-  | 
|
313  | 
from p have p1: "p > 1" using prime_ge_2[OF p] by arith  | 
|
314  | 
hence p01: "p \<noteq> 0" "p \<noteq> 1" by arith+  | 
|
315  | 
from pa have ap: "coprime a p" by (simp add: coprime_commute)  | 
|
316  | 
from prime_coprime[OF p, of x] dvd_imp_le[of p x] x0 xp have px:"coprime x p"  | 
|
317  | 
by (auto simp add: coprime_commute)  | 
|
| 30488 | 318  | 
from cong_solve_unique[OF px p01(1)]  | 
| 26126 | 319  | 
obtain y where y: "y < p" "[x * y = a] (mod p)" "\<forall>z. z < p \<and> [x * z = a] (mod p) \<longrightarrow> z = y" by blast  | 
320  | 
  {assume y0: "y = 0"
 | 
|
321  | 
with y(2) have th: "p dvd a" by (simp add: cong_commute[of 0 a p])  | 
|
322  | 
with p coprime_prime[OF pa, of p] have False by simp}  | 
|
| 30488 | 323  | 
with y show ?thesis unfolding Ex1_def using neq0_conv by blast  | 
| 26126 | 324  | 
qed  | 
325  | 
lemma cong_unique_inverse_prime:  | 
|
326  | 
assumes p: "prime p" and x0: "0 < x" and xp: "x < p"  | 
|
327  | 
shows "\<exists>!y. 0 < y \<and> y < p \<and> [x * y = 1] (mod p)"  | 
|
328  | 
using cong_solve_unique_nontrivial[OF p coprime_1[of p] x0 xp] .  | 
|
329  | 
||
330  | 
(* Forms of the Chinese remainder theorem. *)  | 
|
331  | 
||
| 30488 | 332  | 
lemma cong_chinese:  | 
333  | 
assumes ab: "coprime a b" and xya: "[x = y] (mod a)"  | 
|
| 26126 | 334  | 
and xyb: "[x = y] (mod b)"  | 
335  | 
shows "[x = y] (mod a*b)"  | 
|
336  | 
using ab xya xyb  | 
|
| 30488 | 337  | 
by (simp add: cong_sub_cases[of x y a] cong_sub_cases[of x y b]  | 
338  | 
cong_sub_cases[of x y "a*b"])  | 
|
| 26126 | 339  | 
(cases "x \<le> y", simp_all add: divides_mul[of a _ b])  | 
340  | 
||
341  | 
lemma chinese_remainder_unique:  | 
|
342  | 
assumes ab: "coprime a b" and az: "a \<noteq> 0" and bz: "b\<noteq>0"  | 
|
343  | 
shows "\<exists>!x. x < a * b \<and> [x = m] (mod a) \<and> [x = n] (mod b)"  | 
|
344  | 
proof-  | 
|
345  | 
from az bz have abpos: "a*b > 0" by simp  | 
|
| 30488 | 346  | 
from chinese_remainder[OF ab az bz] obtain x q1 q2 where  | 
| 26126 | 347  | 
xq12: "x = m + q1 * a" "x = n + q2 * b" by blast  | 
| 30488 | 348  | 
let ?w = "x mod (a*b)"  | 
| 26126 | 349  | 
have wab: "?w < a*b" by (simp add: mod_less_divisor[OF abpos])  | 
350  | 
from xq12(1) have "?w mod a = ((m + q1 * a) mod (a*b)) mod a" by simp  | 
|
351  | 
also have "\<dots> = m mod a" apply (simp add: mod_mult2_eq)  | 
|
352  | 
apply (subst mod_add_left_eq)  | 
|
353  | 
by simp  | 
|
354  | 
finally have th1: "[?w = m] (mod a)" by (simp add: modeq_def)  | 
|
355  | 
from xq12(2) have "?w mod b = ((n + q2 * b) mod (a*b)) mod b" by simp  | 
|
356  | 
also have "\<dots> = ((n + q2 * b) mod (b*a)) mod b" by (simp add: mult_commute)  | 
|
357  | 
also have "\<dots> = n mod b" apply (simp add: mod_mult2_eq)  | 
|
358  | 
apply (subst mod_add_left_eq)  | 
|
359  | 
by simp  | 
|
360  | 
finally have th2: "[?w = n] (mod b)" by (simp add: modeq_def)  | 
|
361  | 
  {fix y assume H: "y < a*b" "[y = m] (mod a)" "[y = n] (mod b)"
 | 
|
362  | 
with th1 th2 have H': "[y = ?w] (mod a)" "[y = ?w] (mod b)"  | 
|
363  | 
by (simp_all add: modeq_def)  | 
|
| 30488 | 364  | 
from cong_chinese[OF ab H'] mod_less[OF H(1)] mod_less[OF wab]  | 
| 26126 | 365  | 
have "y = ?w" by (simp add: modeq_def)}  | 
366  | 
with th1 th2 wab show ?thesis by blast  | 
|
367  | 
qed  | 
|
368  | 
||
369  | 
lemma chinese_remainder_coprime_unique:  | 
|
| 30488 | 370  | 
assumes ab: "coprime a b" and az: "a \<noteq> 0" and bz: "b \<noteq> 0"  | 
| 26126 | 371  | 
and ma: "coprime m a" and nb: "coprime n b"  | 
372  | 
shows "\<exists>!x. coprime x (a * b) \<and> x < a * b \<and> [x = m] (mod a) \<and> [x = n] (mod b)"  | 
|
373  | 
proof-  | 
|
374  | 
let ?P = "\<lambda>x. x < a * b \<and> [x = m] (mod a) \<and> [x = n] (mod b)"  | 
|
375  | 
from chinese_remainder_unique[OF ab az bz]  | 
|
| 30488 | 376  | 
obtain x where x: "x < a * b" "[x = m] (mod a)" "[x = n] (mod b)"  | 
| 26126 | 377  | 
"\<forall>y. ?P y \<longrightarrow> y = x" by blast  | 
378  | 
from ma nb cong_coprime[OF x(2)] cong_coprime[OF x(3)]  | 
|
379  | 
have "coprime x a" "coprime x b" by (simp_all add: coprime_commute)  | 
|
380  | 
with coprime_mul[of x a b] have "coprime x (a*b)" by simp  | 
|
381  | 
with x show ?thesis by blast  | 
|
382  | 
qed  | 
|
383  | 
||
384  | 
(* Euler totient function. *)  | 
|
385  | 
||
386  | 
definition phi_def: "\<phi> n = card { m. 0 < m \<and> m <= n \<and> coprime m n }"
 | 
|
| 
31197
 
c1c163ec6c44
fine-tuned elimination of comprehensions involving x=t.
 
nipkow 
parents: 
31021 
diff
changeset
 | 
387  | 
|
| 26126 | 388  | 
lemma phi_0[simp]: "\<phi> 0 = 0"  | 
| 
31197
 
c1c163ec6c44
fine-tuned elimination of comprehensions involving x=t.
 
nipkow 
parents: 
31021 
diff
changeset
 | 
389  | 
unfolding phi_def by auto  | 
| 26126 | 390  | 
|
391  | 
lemma phi_finite[simp]: "finite ({ m. 0 < m \<and> m <= n \<and> coprime m n })"
 | 
|
392  | 
proof-  | 
|
393  | 
  have "{ m. 0 < m \<and> m <= n \<and> coprime m n } \<subseteq> {0..n}" by auto
 | 
|
394  | 
thus ?thesis by (auto intro: finite_subset)  | 
|
395  | 
qed  | 
|
396  | 
||
397  | 
declare coprime_1[presburger]  | 
|
398  | 
lemma phi_1[simp]: "\<phi> 1 = 1"  | 
|
399  | 
proof-  | 
|
| 30488 | 400  | 
  {fix m
 | 
| 26126 | 401  | 
have "0 < m \<and> m <= 1 \<and> coprime m 1 \<longleftrightarrow> m = 1" by presburger }  | 
402  | 
thus ?thesis by (simp add: phi_def)  | 
|
403  | 
qed  | 
|
404  | 
||
405  | 
lemma [simp]: "\<phi> (Suc 0) = Suc 0" using phi_1 by simp  | 
|
406  | 
||
407  | 
lemma phi_alt: "\<phi>(n) = card { m. coprime m n \<and> m < n}"
 | 
|
408  | 
proof-  | 
|
409  | 
  {assume "n=0 \<or> n=1" hence ?thesis by (cases "n=0", simp_all)}
 | 
|
410  | 
moreover  | 
|
411  | 
  {assume n: "n\<noteq>0" "n\<noteq>1"
 | 
|
412  | 
    {fix m
 | 
|
413  | 
from n have "0 < m \<and> m <= n \<and> coprime m n \<longleftrightarrow> coprime m n \<and> m < n"  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
414  | 
apply (cases "m = 0", simp_all)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
415  | 
apply (cases "m = 1", simp_all)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
416  | 
apply (cases "m = n", auto)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
417  | 
done }  | 
| 26126 | 418  | 
hence ?thesis unfolding phi_def by simp}  | 
419  | 
ultimately show ?thesis by auto  | 
|
420  | 
qed  | 
|
421  | 
||
422  | 
lemma phi_finite_lemma[simp]: "finite {m. coprime m n \<and>  m < n}" (is "finite ?S")
 | 
|
423  | 
  by (rule finite_subset[of "?S" "{0..n}"], auto)
 | 
|
424  | 
||
425  | 
lemma phi_another: assumes n: "n\<noteq>1"  | 
|
426  | 
  shows "\<phi> n = card {m. 0 < m \<and> m < n \<and> coprime m n }"
 | 
|
427  | 
proof-  | 
|
| 30488 | 428  | 
  {fix m
 | 
| 26126 | 429  | 
from n have "0 < m \<and> m < n \<and> coprime m n \<longleftrightarrow> coprime m n \<and> m < n"  | 
430  | 
by (cases "m=0", auto)}  | 
|
431  | 
thus ?thesis unfolding phi_alt by auto  | 
|
432  | 
qed  | 
|
433  | 
||
434  | 
lemma phi_limit: "\<phi> n \<le> n"  | 
|
435  | 
proof-  | 
|
436  | 
  have "{ m. coprime m n \<and> m < n} \<subseteq> {0 ..<n}" by auto
 | 
|
437  | 
  with card_mono[of "{0 ..<n}" "{ m. coprime m n \<and> m < n}"]
 | 
|
438  | 
show ?thesis unfolding phi_alt by auto  | 
|
439  | 
qed  | 
|
440  | 
||
441  | 
lemma stupid[simp]: "{m. (0::nat) < m \<and> m < n} = {1..<n}"
 | 
|
442  | 
by auto  | 
|
443  | 
||
| 30488 | 444  | 
lemma phi_limit_strong: assumes n: "n\<noteq>1"  | 
| 26126 | 445  | 
shows "\<phi>(n) \<le> n - 1"  | 
446  | 
proof-  | 
|
447  | 
show ?thesis  | 
|
| 30488 | 448  | 
unfolding phi_another[OF n] finite_number_segment[of n, symmetric]  | 
| 26126 | 449  | 
    by (rule card_mono[of "{m. 0 < m \<and> m < n}" "{m. 0 < m \<and> m < n \<and> coprime m n}"], auto)
 | 
450  | 
qed  | 
|
451  | 
||
452  | 
lemma phi_lowerbound_1_strong: assumes n: "n \<ge> 1"  | 
|
453  | 
shows "\<phi>(n) \<ge> 1"  | 
|
454  | 
proof-  | 
|
455  | 
  let ?S = "{ m. 0 < m \<and> m <= n \<and> coprime m n }"
 | 
|
| 30488 | 456  | 
from card_0_eq[of ?S] n have "\<phi> n \<noteq> 0" unfolding phi_alt  | 
| 26126 | 457  | 
apply auto  | 
458  | 
apply (cases "n=1", simp_all)  | 
|
459  | 
apply (rule exI[where x=1], simp)  | 
|
460  | 
done  | 
|
461  | 
thus ?thesis by arith  | 
|
462  | 
qed  | 
|
463  | 
||
464  | 
lemma phi_lowerbound_1: "2 <= n ==> 1 <= \<phi>(n)"  | 
|
465  | 
using phi_lowerbound_1_strong[of n] by auto  | 
|
466  | 
||
467  | 
lemma phi_lowerbound_2: assumes n: "3 <= n" shows "2 <= \<phi> (n)"  | 
|
468  | 
proof-  | 
|
469  | 
  let ?S = "{ m. 0 < m \<and> m <= n \<and> coprime m n }"
 | 
|
| 30488 | 470  | 
  have inS: "{1, n - 1} \<subseteq> ?S" using n coprime_plus1[of "n - 1"]
 | 
| 26126 | 471  | 
by (auto simp add: coprime_commute)  | 
472  | 
  from n have c2: "card {1, n - 1} = 2" by (auto simp add: card_insert_if)
 | 
|
| 30488 | 473  | 
  from card_mono[of ?S "{1, n - 1}", simplified inS c2] show ?thesis
 | 
| 26126 | 474  | 
unfolding phi_def by auto  | 
475  | 
qed  | 
|
476  | 
||
477  | 
lemma phi_prime: "\<phi> n = n - 1 \<and> n\<noteq>0 \<and> n\<noteq>1 \<longleftrightarrow> prime n"  | 
|
478  | 
proof-  | 
|
479  | 
  {assume "n=0 \<or> n=1" hence ?thesis by (cases "n=1", simp_all)}
 | 
|
480  | 
moreover  | 
|
481  | 
  {assume n: "n\<noteq>0" "n\<noteq>1"
 | 
|
482  | 
    let ?S = "{m. 0 < m \<and> m < n}"
 | 
|
483  | 
have fS: "finite ?S" by simp  | 
|
484  | 
    let ?S' = "{m. 0 < m \<and> m < n \<and> coprime m n}"
 | 
|
485  | 
have fS':"finite ?S'" apply (rule finite_subset[of ?S' ?S]) by auto  | 
|
486  | 
    {assume H: "\<phi> n = n - 1 \<and> n\<noteq>0 \<and> n\<noteq>1"
 | 
|
| 30488 | 487  | 
hence ceq: "card ?S' = card ?S"  | 
| 26126 | 488  | 
using n finite_number_segment[of n] phi_another[OF n(2)] by simp  | 
489  | 
      {fix m assume m: "0 < m" "m < n" "\<not> coprime m n"
 | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
490  | 
hence mS': "m \<notin> ?S'" by auto  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
491  | 
have "insert m ?S' \<le> ?S" using m by auto  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
492  | 
from m have "card (insert m ?S') \<le> card ?S"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
493  | 
by - (rule card_mono[of ?S "insert m ?S'"], auto)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
494  | 
hence False  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
495  | 
unfolding card_insert_disjoint[of "?S'" m, OF fS' mS'] ceq  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
496  | 
by simp }  | 
| 26126 | 497  | 
hence "\<forall>m. 0 <m \<and> m < n \<longrightarrow> coprime m n" by blast  | 
498  | 
hence "prime n" unfolding prime using n by (simp add: coprime_commute)}  | 
|
499  | 
moreover  | 
|
500  | 
    {assume H: "prime n"
 | 
|
| 30488 | 501  | 
hence "?S = ?S'" unfolding prime using n  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
502  | 
by (auto simp add: coprime_commute)  | 
| 26126 | 503  | 
hence "card ?S = card ?S'" by simp  | 
504  | 
hence "\<phi> n = n - 1" unfolding phi_another[OF n(2)] by simp}  | 
|
505  | 
ultimately have ?thesis using n by blast}  | 
|
506  | 
ultimately show ?thesis by (cases "n=0") blast+  | 
|
507  | 
qed  | 
|
508  | 
||
509  | 
(* Multiplicativity property. *)  | 
|
510  | 
||
511  | 
lemma phi_multiplicative: assumes ab: "coprime a b"  | 
|
512  | 
shows "\<phi> (a * b) = \<phi> a * \<phi> b"  | 
|
513  | 
proof-  | 
|
| 30488 | 514  | 
  {assume "a = 0 \<or> b = 0 \<or> a = 1 \<or> b = 1"
 | 
| 26126 | 515  | 
hence ?thesis  | 
516  | 
by (cases "a=0", simp, cases "b=0", simp, cases"a=1", simp_all) }  | 
|
517  | 
moreover  | 
|
518  | 
  {assume a: "a\<noteq>0" "a\<noteq>1" and b: "b\<noteq>0" "b\<noteq>1"
 | 
|
519  | 
hence ab0: "a*b \<noteq> 0" by simp  | 
|
520  | 
    let ?S = "\<lambda>k. {m. coprime m k \<and> m < k}"
 | 
|
521  | 
let ?f = "\<lambda>x. (x mod a, x mod b)"  | 
|
522  | 
have eq: "?f ` (?S (a*b)) = (?S a \<times> ?S b)"  | 
|
523  | 
proof-  | 
|
524  | 
      {fix x assume x:"x \<in> ?S (a*b)"
 | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
525  | 
hence x': "coprime x (a*b)" "x < a*b" by simp_all  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
526  | 
hence xab: "coprime x a" "coprime x b" by (simp_all add: coprime_mul_eq)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
527  | 
from mod_less_divisor a b have xab':"x mod a < a" "x mod b < b" by auto  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
528  | 
from xab xab' have "?f x \<in> (?S a \<times> ?S b)"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
529  | 
by (simp add: coprime_mod[OF a(1)] coprime_mod[OF b(1)])}  | 
| 26126 | 530  | 
moreover  | 
531  | 
      {fix x y assume x: "x \<in> ?S a" and y: "y \<in> ?S b"
 | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
532  | 
hence x': "coprime x a" "x < a" and y': "coprime y b" "y < b" by simp_all  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
533  | 
from chinese_remainder_coprime_unique[OF ab a(1) b(1) x'(1) y'(1)]  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
534  | 
obtain z where z: "coprime z (a * b)" "z < a * b" "[z = x] (mod a)"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
535  | 
"[z = y] (mod b)" by blast  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
536  | 
hence "(x,y) \<in> ?f ` (?S (a*b))"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
537  | 
using y'(2) mod_less_divisor[of b y] x'(2) mod_less_divisor[of a x]  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
538  | 
by (auto simp add: image_iff modeq_def)}  | 
| 26126 | 539  | 
ultimately show ?thesis by auto  | 
540  | 
qed  | 
|
541  | 
have finj: "inj_on ?f (?S (a*b))"  | 
|
542  | 
unfolding inj_on_def  | 
|
543  | 
proof(clarify)  | 
|
| 30488 | 544  | 
fix x y assume H: "coprime x (a * b)" "x < a * b" "coprime y (a * b)"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
545  | 
"y < a * b" "x mod a = y mod a" "x mod b = y mod b"  | 
| 30488 | 546  | 
hence cp: "coprime x a" "coprime x b" "coprime y a" "coprime y b"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
547  | 
by (simp_all add: coprime_mul_eq)  | 
| 26126 | 548  | 
from chinese_remainder_coprime_unique[OF ab a(1) b(1) cp(3,4)] H  | 
549  | 
show "x = y" unfolding modeq_def by blast  | 
|
550  | 
qed  | 
|
551  | 
from card_image[OF finj, unfolded eq] have ?thesis  | 
|
552  | 
unfolding phi_alt by simp }  | 
|
553  | 
ultimately show ?thesis by auto  | 
|
554  | 
qed  | 
|
555  | 
||
556  | 
(* Fermat's Little theorem / Fermat-Euler theorem. *)  | 
|
557  | 
||
558  | 
||
559  | 
lemma nproduct_mod:  | 
|
560  | 
assumes fS: "finite S" and n0: "n \<noteq> 0"  | 
|
561  | 
shows "[setprod (\<lambda>m. a(m) mod n) S = setprod a S] (mod n)"  | 
|
562  | 
proof-  | 
|
563  | 
have th1:"[1 = 1] (mod n)" by (simp add: modeq_def)  | 
|
564  | 
from cong_mult  | 
|
565  | 
have th3:"\<forall>x1 y1 x2 y2.  | 
|
566  | 
[x1 = x2] (mod n) \<and> [y1 = y2] (mod n) \<longrightarrow> [x1 * y1 = x2 * y2] (mod n)"  | 
|
567  | 
by blast  | 
|
568  | 
have th4:"\<forall>x\<in>S. [a x mod n = a x] (mod n)" by (simp add: modeq_def)  | 
|
| 28854 | 569  | 
from fold_image_related[where h="(\<lambda>m. a(m) mod n)" and g=a, OF th1 th3 fS, OF th4] show ?thesis unfolding setprod_def by (simp add: fS)  | 
| 26126 | 570  | 
qed  | 
571  | 
||
572  | 
lemma nproduct_cmul:  | 
|
573  | 
assumes fS:"finite S"  | 
|
| 31021 | 574  | 
  shows "setprod (\<lambda>m. (c::'a::{comm_monoid_mult})* a(m)) S = c ^ (card S) * setprod a S"
 | 
| 26126 | 575  | 
unfolding setprod_timesf setprod_constant[OF fS, of c] ..  | 
576  | 
||
577  | 
lemma coprime_nproduct:  | 
|
578  | 
assumes fS: "finite S" and Sn: "\<forall>x\<in>S. coprime n (a x)"  | 
|
579  | 
shows "coprime n (setprod a S)"  | 
|
| 27368 | 580  | 
using fS unfolding setprod_def by (rule finite_subset_induct)  | 
581  | 
(insert Sn, auto simp add: coprime_mul)  | 
|
| 26126 | 582  | 
|
583  | 
lemma fermat_little: assumes an: "coprime a n"  | 
|
584  | 
shows "[a ^ (\<phi> n) = 1] (mod n)"  | 
|
585  | 
proof-  | 
|
586  | 
  {assume "n=0" hence ?thesis by simp}
 | 
|
587  | 
moreover  | 
|
588  | 
  {assume "n=1" hence ?thesis by (simp add: modeq_def)}
 | 
|
589  | 
moreover  | 
|
590  | 
  {assume nz: "n \<noteq> 0" and n1: "n \<noteq> 1"
 | 
|
591  | 
    let ?S = "{m. coprime m n \<and> m < n}"
 | 
|
592  | 
let ?P = "\<Prod> ?S"  | 
|
593  | 
have fS: "finite ?S" by simp  | 
|
594  | 
have cardfS: "\<phi> n = card ?S" unfolding phi_alt ..  | 
|
595  | 
    {fix m assume m: "m \<in> ?S"
 | 
|
596  | 
hence "coprime m n" by simp  | 
|
| 30488 | 597  | 
with coprime_mul[of n a m] an have "coprime (a*m) n"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
598  | 
by (simp add: coprime_commute)}  | 
| 26126 | 599  | 
hence Sn: "\<forall>m\<in> ?S. coprime (a*m) n " by blast  | 
600  | 
from coprime_nproduct[OF fS, of n "\<lambda>m. m"] have nP:"coprime ?P n"  | 
|
601  | 
by (simp add: coprime_commute)  | 
|
602  | 
have Paphi: "[?P*a^ (\<phi> n) = ?P*1] (mod n)"  | 
|
603  | 
proof-  | 
|
604  | 
let ?h = "\<lambda>m. m mod n"  | 
|
605  | 
      {fix m assume mS: "m\<in> ?S"
 | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
606  | 
hence "?h m \<in> ?S" by simp}  | 
| 26126 | 607  | 
hence hS: "?h ` ?S = ?S"by (auto simp add: image_iff)  | 
608  | 
have "a\<noteq>0" using an n1 nz apply- apply (rule ccontr) by simp  | 
|
609  | 
hence inj: "inj_on (op * a) ?S" unfolding inj_on_def by simp  | 
|
| 30488 | 610  | 
|
| 28854 | 611  | 
      have eq0: "fold_image op * (?h \<circ> op * a) 1 {m. coprime m n \<and> m < n} =
 | 
612  | 
     fold_image op * (\<lambda>m. m) 1 {m. coprime m n \<and> m < n}"
 | 
|
613  | 
proof (rule fold_image_eq_general[where h="?h o (op * a)"])  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
614  | 
show "finite ?S" using fS .  | 
| 26126 | 615  | 
next  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
616  | 
        {fix y assume yS: "y \<in> ?S" hence y: "coprime y n" "y < n" by simp_all
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
617  | 
from cong_solve_unique[OF an nz, of y]  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
618  | 
obtain x where x:"x < n" "[a * x = y] (mod n)" "\<forall>z. z < n \<and> [a * z = y] (mod n) \<longrightarrow> z=x" by blast  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
619  | 
from cong_coprime[OF x(2)] y(1)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
620  | 
have xm: "coprime x n" by (simp add: coprime_mul_eq coprime_commute)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
621  | 
          {fix z assume "z \<in> ?S" "(?h \<circ> op * a) z = y"
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
622  | 
hence z: "coprime z n" "z < n" "(?h \<circ> op * a) z = y" by simp_all  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
623  | 
from x(3)[rule_format, of z] z(2,3) have "z=x"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
624  | 
unfolding modeq_def mod_less[OF y(2)] by simp}  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
625  | 
with xm x(1,2) have "\<exists>!x. x \<in> ?S \<and> (?h \<circ> op * a) x = y"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
626  | 
unfolding modeq_def mod_less[OF y(2)] by auto }  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
627  | 
        thus "\<forall>y\<in>{m. coprime m n \<and> m < n}.
 | 
| 26126 | 628  | 
       \<exists>!x. x \<in> {m. coprime m n \<and> m < n} \<and> ((\<lambda>m. m mod n) \<circ> op * a) x = y" by blast
 | 
629  | 
next  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
630  | 
        {fix x assume xS: "x\<in> ?S"
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
631  | 
hence x: "coprime x n" "x < n" by simp_all  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
632  | 
with an have "coprime (a*x) n"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
633  | 
by (simp add: coprime_mul_eq[of n a x] coprime_commute)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
634  | 
hence "?h (a*x) \<in> ?S" using nz  | 
| 41541 | 635  | 
by (simp add: coprime_mod[OF nz])}  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
636  | 
        thus " \<forall>x\<in>{m. coprime m n \<and> m < n}.
 | 
| 26126 | 637  | 
       ((\<lambda>m. m mod n) \<circ> op * a) x \<in> {m. coprime m n \<and> m < n} \<and>
 | 
638  | 
((\<lambda>m. m mod n) \<circ> op * a) x = ((\<lambda>m. m mod n) \<circ> op * a) x" by simp  | 
|
639  | 
qed  | 
|
640  | 
from nproduct_mod[OF fS nz, of "op * a"]  | 
|
641  | 
have "[(setprod (op *a) ?S) = (setprod (?h o (op * a)) ?S)] (mod n)"  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
642  | 
unfolding o_def  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
643  | 
by (simp add: cong_commute)  | 
| 26126 | 644  | 
also have "[setprod (?h o (op * a)) ?S = ?P ] (mod n)"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
645  | 
using eq0 fS an by (simp add: setprod_def modeq_def o_def)  | 
| 26126 | 646  | 
finally show "[?P*a^ (\<phi> n) = ?P*1] (mod n)"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
647  | 
unfolding cardfS mult_commute[of ?P "a^ (card ?S)"]  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
648  | 
nproduct_cmul[OF fS, symmetric] mult_1_right by simp  | 
| 26126 | 649  | 
qed  | 
650  | 
from cong_mult_lcancel[OF nP Paphi] have ?thesis . }  | 
|
651  | 
ultimately show ?thesis by blast  | 
|
652  | 
qed  | 
|
653  | 
||
654  | 
lemma fermat_little_prime: assumes p: "prime p" and ap: "coprime a p"  | 
|
655  | 
shows "[a^ (p - 1) = 1] (mod p)"  | 
|
656  | 
using fermat_little[OF ap] p[unfolded phi_prime[symmetric]]  | 
|
657  | 
by simp  | 
|
658  | 
||
659  | 
||
660  | 
(* Lucas's theorem. *)  | 
|
661  | 
||
662  | 
lemma lucas_coprime_lemma:  | 
|
663  | 
assumes m: "m\<noteq>0" and am: "[a^m = 1] (mod n)"  | 
|
664  | 
shows "coprime a n"  | 
|
665  | 
proof-  | 
|
666  | 
  {assume "n=1" hence ?thesis by simp}
 | 
|
667  | 
moreover  | 
|
668  | 
  {assume "n = 0" hence ?thesis using am m exp_eq_1[of a m] by simp}
 | 
|
669  | 
moreover  | 
|
670  | 
  {assume n: "n\<noteq>0" "n\<noteq>1"
 | 
|
671  | 
from m obtain m' where m': "m = Suc m'" by (cases m, blast+)  | 
|
672  | 
    {fix d
 | 
|
673  | 
assume d: "d dvd a" "d dvd n"  | 
|
| 30488 | 674  | 
from n have n1: "1 < n" by arith  | 
| 26126 | 675  | 
from am mod_less[OF n1] have am1: "a^m mod n = 1" unfolding modeq_def by simp  | 
676  | 
from dvd_mult2[OF d(1), of "a^m'"] have dam:"d dvd a^m" by (simp add: m')  | 
|
677  | 
from dvd_mod_iff[OF d(2), of "a^m"] dam am1  | 
|
678  | 
have "d = 1" by simp }  | 
|
679  | 
hence ?thesis unfolding coprime by auto  | 
|
680  | 
}  | 
|
| 30488 | 681  | 
ultimately show ?thesis by blast  | 
| 26126 | 682  | 
qed  | 
683  | 
||
684  | 
lemma lucas_weak:  | 
|
| 30488 | 685  | 
assumes n: "n \<ge> 2" and an:"[a^(n - 1) = 1] (mod n)"  | 
| 26126 | 686  | 
and nm: "\<forall>m. 0 <m \<and> m < n - 1 \<longrightarrow> \<not> [a^m = 1] (mod n)"  | 
687  | 
shows "prime n"  | 
|
688  | 
proof-  | 
|
689  | 
from n have n1: "n \<noteq> 1" "n\<noteq>0" "n - 1 \<noteq> 0" "n - 1 > 0" "n - 1 < n" by arith+  | 
|
690  | 
from lucas_coprime_lemma[OF n1(3) an] have can: "coprime a n" .  | 
|
691  | 
from fermat_little[OF can] have afn: "[a ^ \<phi> n = 1] (mod n)" .  | 
|
692  | 
  {assume "\<phi> n \<noteq> n - 1"
 | 
|
693  | 
with phi_limit_strong[OF n1(1)] phi_lowerbound_1[OF n]  | 
|
694  | 
have c:"\<phi> n > 0 \<and> \<phi> n < n - 1" by arith  | 
|
695  | 
from nm[rule_format, OF c] afn have False ..}  | 
|
696  | 
hence "\<phi> n = n - 1" by blast  | 
|
697  | 
with phi_prime[of n] n1(1,2) show ?thesis by simp  | 
|
698  | 
qed  | 
|
699  | 
||
| 30488 | 700  | 
lemma nat_exists_least_iff: "(\<exists>(n::nat). P n) \<longleftrightarrow> (\<exists>n. P n \<and> (\<forall>m < n. \<not> P m))"  | 
| 26126 | 701  | 
(is "?lhs \<longleftrightarrow> ?rhs")  | 
702  | 
proof  | 
|
703  | 
assume ?rhs thus ?lhs by blast  | 
|
704  | 
next  | 
|
705  | 
assume H: ?lhs then obtain n where n: "P n" by blast  | 
|
706  | 
let ?x = "Least P"  | 
|
707  | 
  {fix m assume m: "m < ?x"
 | 
|
708  | 
from not_less_Least[OF m] have "\<not> P m" .}  | 
|
709  | 
with LeastI_ex[OF H] show ?rhs by blast  | 
|
710  | 
qed  | 
|
711  | 
||
| 30488 | 712  | 
lemma nat_exists_least_iff': "(\<exists>(n::nat). P n) \<longleftrightarrow> (P (Least P) \<and> (\<forall>m < (Least P). \<not> P m))"  | 
| 26126 | 713  | 
(is "?lhs \<longleftrightarrow> ?rhs")  | 
714  | 
proof-  | 
|
715  | 
  {assume ?rhs hence ?lhs by blast}
 | 
|
| 30488 | 716  | 
moreover  | 
| 26126 | 717  | 
  { assume H: ?lhs then obtain n where n: "P n" by blast
 | 
718  | 
let ?x = "Least P"  | 
|
719  | 
    {fix m assume m: "m < ?x"
 | 
|
720  | 
from not_less_Least[OF m] have "\<not> P m" .}  | 
|
721  | 
with LeastI_ex[OF H] have ?rhs by blast}  | 
|
722  | 
ultimately show ?thesis by blast  | 
|
723  | 
qed  | 
|
| 30488 | 724  | 
|
| 26126 | 725  | 
lemma power_mod: "((x::nat) mod m)^n mod m = x^n mod m"  | 
726  | 
proof(induct n)  | 
|
727  | 
case 0 thus ?case by simp  | 
|
728  | 
next  | 
|
| 30488 | 729  | 
case (Suc n)  | 
730  | 
have "(x mod m)^(Suc n) mod m = ((x mod m) * (((x mod m) ^ n) mod m)) mod m"  | 
|
| 30224 | 731  | 
by (simp add: mod_mult_right_eq[symmetric])  | 
| 26126 | 732  | 
also have "\<dots> = ((x mod m) * (x^n mod m)) mod m" using Suc.hyps by simp  | 
733  | 
also have "\<dots> = x^(Suc n) mod m"  | 
|
| 30224 | 734  | 
by (simp add: mod_mult_left_eq[symmetric] mod_mult_right_eq[symmetric])  | 
| 26126 | 735  | 
finally show ?case .  | 
736  | 
qed  | 
|
737  | 
||
738  | 
lemma lucas:  | 
|
| 30488 | 739  | 
assumes n2: "n \<ge> 2" and an1: "[a^(n - 1) = 1] (mod n)"  | 
| 26126 | 740  | 
and pn: "\<forall>p. prime p \<and> p dvd n - 1 \<longrightarrow> \<not> [a^((n - 1) div p) = 1] (mod n)"  | 
741  | 
shows "prime n"  | 
|
742  | 
proof-  | 
|
743  | 
from n2 have n01: "n\<noteq>0" "n\<noteq>1" "n - 1 \<noteq> 0" by arith+  | 
|
744  | 
from mod_less_divisor[of n 1] n01 have onen: "1 mod n = 1" by simp  | 
|
| 30488 | 745  | 
from lucas_coprime_lemma[OF n01(3) an1] cong_coprime[OF an1]  | 
| 26126 | 746  | 
have an: "coprime a n" "coprime (a^(n - 1)) n" by (simp_all add: coprime_commute)  | 
747  | 
  {assume H0: "\<exists>m. 0 < m \<and> m < n - 1 \<and> [a ^ m = 1] (mod n)" (is "EX m. ?P m")
 | 
|
| 30488 | 748  | 
from H0[unfolded nat_exists_least_iff[of ?P]] obtain m where  | 
| 26126 | 749  | 
m: "0 < m" "m < n - 1" "[a ^ m = 1] (mod n)" "\<forall>k <m. \<not>?P k" by blast  | 
| 30488 | 750  | 
    {assume nm1: "(n - 1) mod m > 0"
 | 
751  | 
from mod_less_divisor[OF m(1)] have th0:"(n - 1) mod m < m" by blast  | 
|
| 26126 | 752  | 
let ?y = "a^ ((n - 1) div m * m)"  | 
753  | 
note mdeq = mod_div_equality[of "(n - 1)" m]  | 
|
| 30488 | 754  | 
from coprime_exp[OF an(1)[unfolded coprime_commute[of a n]],  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
755  | 
of "(n - 1) div m * m"]  | 
| 30488 | 756  | 
have yn: "coprime ?y n" by (simp add: coprime_commute)  | 
757  | 
have "?y mod n = (a^m)^((n - 1) div m) mod n"  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
758  | 
by (simp add: algebra_simps power_mult)  | 
| 30488 | 759  | 
also have "\<dots> = (a^m mod n)^((n - 1) div m) mod n"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
760  | 
using power_mod[of "a^m" n "(n - 1) div m"] by simp  | 
| 30488 | 761  | 
also have "\<dots> = 1" using m(3)[unfolded modeq_def onen] onen  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
762  | 
by (simp add: power_Suc0)  | 
| 30488 | 763  | 
finally have th3: "?y mod n = 1" .  | 
764  | 
have th2: "[?y * a ^ ((n - 1) mod m) = ?y* 1] (mod n)"  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
765  | 
using an1[unfolded modeq_def onen] onen  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
766  | 
mod_div_equality[of "(n - 1)" m, symmetric]  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
767  | 
by (simp add:power_add[symmetric] modeq_def th3 del: One_nat_def)  | 
| 26126 | 768  | 
from cong_mult_lcancel[of ?y n "a^((n - 1) mod m)" 1, OF yn th2]  | 
| 30488 | 769  | 
have th1: "[a ^ ((n - 1) mod m) = 1] (mod n)" .  | 
770  | 
from m(4)[rule_format, OF th0] nm1  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
771  | 
less_trans[OF mod_less_divisor[OF m(1), of "n - 1"] m(2)] th1  | 
| 26126 | 772  | 
have False by blast }  | 
773  | 
hence "(n - 1) mod m = 0" by auto  | 
|
774  | 
then have mn: "m dvd n - 1" by presburger  | 
|
775  | 
then obtain r where r: "n - 1 = m*r" unfolding dvd_def by blast  | 
|
776  | 
from n01 r m(2) have r01: "r\<noteq>0" "r\<noteq>1" by - (rule ccontr, simp)+  | 
|
777  | 
from prime_factor[OF r01(2)] obtain p where p: "prime p" "p dvd r" by blast  | 
|
778  | 
hence th: "prime p \<and> p dvd n - 1" unfolding r by (auto intro: dvd_mult)  | 
|
779  | 
have "(a ^ ((n - 1) div p)) mod n = (a^(m*r div p)) mod n" using r  | 
|
780  | 
by (simp add: power_mult)  | 
|
781  | 
also have "\<dots> = (a^(m*(r div p))) mod n" using div_mult1_eq[of m r p] p(2)[unfolded dvd_eq_mod_eq_0] by simp  | 
|
782  | 
also have "\<dots> = ((a^m)^(r div p)) mod n" by (simp add: power_mult)  | 
|
783  | 
also have "\<dots> = ((a^m mod n)^(r div p)) mod n" using power_mod[of "a^m" "n" "r div p" ] ..  | 
|
| 26158 | 784  | 
also have "\<dots> = 1" using m(3) onen by (simp add: modeq_def power_Suc0)  | 
| 30488 | 785  | 
finally have "[(a ^ ((n - 1) div p))= 1] (mod n)"  | 
| 26126 | 786  | 
using onen by (simp add: modeq_def)  | 
787  | 
with pn[rule_format, OF th] have False by blast}  | 
|
788  | 
hence th: "\<forall>m. 0 < m \<and> m < n - 1 \<longrightarrow> \<not> [a ^ m = 1] (mod n)" by blast  | 
|
789  | 
from lucas_weak[OF n2 an1 th] show ?thesis .  | 
|
790  | 
qed  | 
|
791  | 
||
792  | 
(* Definition of the order of a number mod n (0 in non-coprime case). *)  | 
|
793  | 
||
794  | 
definition "ord n a = (if coprime n a then Least (\<lambda>d. d > 0 \<and> [a ^d = 1] (mod n)) else 0)"  | 
|
795  | 
||
796  | 
(* This has the expected properties. *)  | 
|
797  | 
||
798  | 
lemma coprime_ord:  | 
|
| 30488 | 799  | 
assumes na: "coprime n a"  | 
| 26126 | 800  | 
shows "ord n a > 0 \<and> [a ^(ord n a) = 1] (mod n) \<and> (\<forall>m. 0 < m \<and> m < ord n a \<longrightarrow> \<not> [a^ m = 1] (mod n))"  | 
801  | 
proof-  | 
|
802  | 
let ?P = "\<lambda>d. 0 < d \<and> [a ^ d = 1] (mod n)"  | 
|
803  | 
from euclid[of a] obtain p where p: "prime p" "a < p" by blast  | 
|
804  | 
from na have o: "ord n a = Least ?P" by (simp add: ord_def)  | 
|
805  | 
  {assume "n=0 \<or> n=1" with na have "\<exists>m>0. ?P m" apply auto apply (rule exI[where x=1]) by (simp  add: modeq_def)}
 | 
|
806  | 
moreover  | 
|
| 30488 | 807  | 
  {assume "n\<noteq>0 \<and> n\<noteq>1" hence n2:"n \<ge> 2" by arith
 | 
| 26126 | 808  | 
from na have na': "coprime a n" by (simp add: coprime_commute)  | 
809  | 
from phi_lowerbound_1[OF n2] fermat_little[OF na']  | 
|
810  | 
have ex: "\<exists>m>0. ?P m" by - (rule exI[where x="\<phi> n"], auto) }  | 
|
811  | 
ultimately have ex: "\<exists>m>0. ?P m" by blast  | 
|
| 30488 | 812  | 
from nat_exists_least_iff'[of ?P] ex na show ?thesis  | 
| 26126 | 813  | 
unfolding o[symmetric] by auto  | 
814  | 
qed  | 
|
815  | 
(* With the special value 0 for non-coprime case, it's more convenient. *)  | 
|
816  | 
lemma ord_works:  | 
|
817  | 
"[a ^ (ord n a) = 1] (mod n) \<and> (\<forall>m. 0 < m \<and> m < ord n a \<longrightarrow> ~[a^ m = 1] (mod n))"  | 
|
818  | 
apply (cases "coprime n a")  | 
|
819  | 
using coprime_ord[of n a]  | 
|
820  | 
by (blast, simp add: ord_def modeq_def)  | 
|
821  | 
||
| 30488 | 822  | 
lemma ord: "[a^(ord n a) = 1] (mod n)" using ord_works by blast  | 
823  | 
lemma ord_minimal: "0 < m \<Longrightarrow> m < ord n a \<Longrightarrow> ~[a^m = 1] (mod n)"  | 
|
| 26126 | 824  | 
using ord_works by blast  | 
825  | 
lemma ord_eq_0: "ord n a = 0 \<longleftrightarrow> ~coprime n a"  | 
|
| 41541 | 826  | 
by (cases "coprime n a", simp add: coprime_ord, simp add: ord_def)  | 
| 26126 | 827  | 
|
828  | 
lemma ord_divides:  | 
|
829  | 
"[a ^ d = 1] (mod n) \<longleftrightarrow> ord n a dvd d" (is "?lhs \<longleftrightarrow> ?rhs")  | 
|
830  | 
proof  | 
|
831  | 
assume rh: ?rhs  | 
|
832  | 
then obtain k where "d = ord n a * k" unfolding dvd_def by blast  | 
|
833  | 
hence "[a ^ d = (a ^ (ord n a) mod n)^k] (mod n)"  | 
|
834  | 
by (simp add : modeq_def power_mult power_mod)  | 
|
| 30488 | 835  | 
also have "[(a ^ (ord n a) mod n)^k = 1] (mod n)"  | 
836  | 
using ord[of a n, unfolded modeq_def]  | 
|
| 26158 | 837  | 
by (simp add: modeq_def power_mod power_Suc0)  | 
| 26126 | 838  | 
finally show ?lhs .  | 
| 30488 | 839  | 
next  | 
| 26126 | 840  | 
assume lh: ?lhs  | 
841  | 
  { assume H: "\<not> coprime n a"
 | 
|
842  | 
hence o: "ord n a = 0" by (simp add: ord_def)  | 
|
843  | 
    {assume d: "d=0" with o H have ?rhs by (simp add: modeq_def)}
 | 
|
844  | 
moreover  | 
|
845  | 
    {assume d0: "d\<noteq>0" then obtain d' where d': "d = Suc d'" by (cases d, auto)
 | 
|
| 30488 | 846  | 
from H[unfolded coprime]  | 
847  | 
obtain p where p: "p dvd n" "p dvd a" "p \<noteq> 1" by auto  | 
|
848  | 
from lh[unfolded nat_mod]  | 
|
| 26126 | 849  | 
obtain q1 q2 where q12:"a ^ d + n * q1 = 1 + n * q2" by blast  | 
850  | 
hence "a ^ d + n * q1 - n * q2 = 1" by simp  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31197 
diff
changeset
 | 
851  | 
with dvd_diff_nat [OF dvd_add [OF divides_rexp[OF p(2), of d'] dvd_mult2[OF p(1), of q1]] dvd_mult2[OF p(1), of q2]] d' have "p dvd 1" by simp  | 
| 26126 | 852  | 
with p(3) have False by simp  | 
853  | 
hence ?rhs ..}  | 
|
854  | 
ultimately have ?rhs by blast}  | 
|
855  | 
moreover  | 
|
856  | 
  {assume H: "coprime n a"
 | 
|
857  | 
let ?o = "ord n a"  | 
|
858  | 
let ?q = "d div ord n a"  | 
|
859  | 
let ?r = "d mod ord n a"  | 
|
| 30488 | 860  | 
from cong_exp[OF ord[of a n], of ?q]  | 
| 26158 | 861  | 
have eqo: "[(a^?o)^?q = 1] (mod n)" by (simp add: modeq_def power_Suc0)  | 
| 26126 | 862  | 
from H have onz: "?o \<noteq> 0" by (simp add: ord_eq_0)  | 
863  | 
hence op: "?o > 0" by simp  | 
|
864  | 
from mod_div_equality[of d "ord n a"] lh  | 
|
865  | 
have "[a^(?o*?q + ?r) = 1] (mod n)" by (simp add: modeq_def mult_commute)  | 
|
| 30488 | 866  | 
hence "[(a^?o)^?q * (a^?r) = 1] (mod n)"  | 
| 26126 | 867  | 
by (simp add: modeq_def power_mult[symmetric] power_add[symmetric])  | 
868  | 
hence th: "[a^?r = 1] (mod n)"  | 
|
| 30224 | 869  | 
using eqo mod_mult_left_eq[of "(a^?o)^?q" "a^?r" n]  | 
| 26126 | 870  | 
apply (simp add: modeq_def del: One_nat_def)  | 
| 30224 | 871  | 
by (simp add: mod_mult_left_eq[symmetric])  | 
| 26126 | 872  | 
    {assume r: "?r = 0" hence ?rhs by (simp add: dvd_eq_mod_eq_0)}
 | 
873  | 
moreover  | 
|
| 30488 | 874  | 
    {assume r: "?r \<noteq> 0"
 | 
| 26126 | 875  | 
with mod_less_divisor[OF op, of d] have r0o:"?r >0 \<and> ?r < ?o" by simp  | 
| 30488 | 876  | 
from conjunct2[OF ord_works[of a n], rule_format, OF r0o] th  | 
| 26126 | 877  | 
have ?rhs by blast}  | 
878  | 
ultimately have ?rhs by blast}  | 
|
879  | 
ultimately show ?rhs by blast  | 
|
880  | 
qed  | 
|
881  | 
||
882  | 
lemma order_divides_phi: "coprime n a \<Longrightarrow> ord n a dvd \<phi> n"  | 
|
883  | 
using ord_divides fermat_little coprime_commute by simp  | 
|
| 30488 | 884  | 
lemma order_divides_expdiff:  | 
| 26126 | 885  | 
assumes na: "coprime n a"  | 
886  | 
shows "[a^d = a^e] (mod n) \<longleftrightarrow> [d = e] (mod (ord n a))"  | 
|
887  | 
proof-  | 
|
| 30488 | 888  | 
  {fix n a d e
 | 
| 26126 | 889  | 
assume na: "coprime n a" and ed: "(e::nat) \<le> d"  | 
890  | 
hence "\<exists>c. d = e + c" by arith  | 
|
891  | 
then obtain c where c: "d = e + c" by arith  | 
|
892  | 
from na have an: "coprime a n" by (simp add: coprime_commute)  | 
|
| 30488 | 893  | 
from coprime_exp[OF na, of e]  | 
| 26126 | 894  | 
have aen: "coprime (a^e) n" by (simp add: coprime_commute)  | 
| 30488 | 895  | 
from coprime_exp[OF na, of c]  | 
| 26126 | 896  | 
have acn: "coprime (a^c) n" by (simp add: coprime_commute)  | 
897  | 
have "[a^d = a^e] (mod n) \<longleftrightarrow> [a^(e + c) = a^(e + 0)] (mod n)"  | 
|
898  | 
using c by simp  | 
|
899  | 
also have "\<dots> \<longleftrightarrow> [a^e* a^c = a^e *a^0] (mod n)" by (simp add: power_add)  | 
|
900  | 
also have "\<dots> \<longleftrightarrow> [a ^ c = 1] (mod n)"  | 
|
901  | 
using cong_mult_lcancel_eq[OF aen, of "a^c" "a^0"] by simp  | 
|
902  | 
also have "\<dots> \<longleftrightarrow> ord n a dvd c" by (simp only: ord_divides)  | 
|
903  | 
also have "\<dots> \<longleftrightarrow> [e + c = e + 0] (mod ord n a)"  | 
|
904  | 
using cong_add_lcancel_eq[of e c 0 "ord n a", simplified cong_0_divides]  | 
|
905  | 
by simp  | 
|
906  | 
finally have "[a^d = a^e] (mod n) \<longleftrightarrow> [d = e] (mod (ord n a))"  | 
|
907  | 
using c by simp }  | 
|
908  | 
note th = this  | 
|
909  | 
have "e \<le> d \<or> d \<le> e" by arith  | 
|
910  | 
moreover  | 
|
911  | 
  {assume ed: "e \<le> d" from th[OF na ed] have ?thesis .}
 | 
|
912  | 
moreover  | 
|
913  | 
  {assume de: "d \<le> e"
 | 
|
914  | 
from th[OF na de] have ?thesis by (simp add: cong_commute) }  | 
|
915  | 
ultimately show ?thesis by blast  | 
|
916  | 
qed  | 
|
917  | 
||
918  | 
(* Another trivial primality characterization. *)  | 
|
919  | 
||
920  | 
lemma prime_prime_factor:  | 
|
921  | 
"prime n \<longleftrightarrow> n \<noteq> 1\<and> (\<forall>p. prime p \<and> p dvd n \<longrightarrow> p = n)"  | 
|
922  | 
proof-  | 
|
923  | 
  {assume n: "n=0 \<or> n=1" hence ?thesis using prime_0 two_is_prime by auto}
 | 
|
924  | 
moreover  | 
|
925  | 
  {assume n: "n\<noteq>0" "n\<noteq>1"
 | 
|
926  | 
    {assume pn: "prime n"
 | 
|
| 30488 | 927  | 
|
| 26126 | 928  | 
from pn[unfolded prime_def] have "\<forall>p. prime p \<and> p dvd n \<longrightarrow> p = n"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
929  | 
using n  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
930  | 
apply (cases "n = 0 \<or> n=1",simp)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
931  | 
by (clarsimp, erule_tac x="p" in allE, auto)}  | 
| 26126 | 932  | 
moreover  | 
933  | 
    {assume H: "\<forall>p. prime p \<and> p dvd n \<longrightarrow> p = n"
 | 
|
934  | 
from n have n1: "n > 1" by arith  | 
|
935  | 
      {fix m assume m: "m dvd n" "m\<noteq>1"
 | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
936  | 
from prime_factor[OF m(2)] obtain p where  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
937  | 
p: "prime p" "p dvd m" by blast  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
938  | 
from dvd_trans[OF p(2) m(1)] p(1) H have "p = n" by blast  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
939  | 
with p(2) have "n dvd m" by simp  | 
| 33657 | 940  | 
hence "m=n" using dvd_antisym[OF m(1)] by simp }  | 
| 26126 | 941  | 
with n1 have "prime n" unfolding prime_def by auto }  | 
| 30488 | 942  | 
ultimately have ?thesis using n by blast}  | 
943  | 
ultimately show ?thesis by auto  | 
|
| 26126 | 944  | 
qed  | 
945  | 
||
946  | 
lemma prime_divisor_sqrt:  | 
|
947  | 
"prime n \<longleftrightarrow> n \<noteq> 1 \<and> (\<forall>d. d dvd n \<and> d^2 \<le> n \<longrightarrow> d = 1)"  | 
|
948  | 
proof-  | 
|
| 30488 | 949  | 
  {assume "n=0 \<or> n=1" hence ?thesis using prime_0 prime_1
 | 
| 26126 | 950  | 
by (auto simp add: nat_power_eq_0_iff)}  | 
951  | 
moreover  | 
|
952  | 
  {assume n: "n\<noteq>0" "n\<noteq>1"
 | 
|
953  | 
hence np: "n > 1" by arith  | 
|
954  | 
    {fix d assume d: "d dvd n" "d^2 \<le> n" and H: "\<forall>m. m dvd n \<longrightarrow> m=1 \<or> m=n"
 | 
|
955  | 
from H d have d1n: "d = 1 \<or> d=n" by blast  | 
|
956  | 
      {assume dn: "d=n"
 | 
|
| 41541 | 957  | 
have "n^2 > n*1" using n by (simp add: power2_eq_square)  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
958  | 
with dn d(2) have "d=1" by simp}  | 
| 26126 | 959  | 
with d1n have "d = 1" by blast }  | 
960  | 
moreover  | 
|
961  | 
    {fix d assume d: "d dvd n" and H: "\<forall>d'. d' dvd n \<and> d'^2 \<le> n \<longrightarrow> d' = 1"
 | 
|
962  | 
from d n have "d \<noteq> 0" apply - apply (rule ccontr) by simp  | 
|
963  | 
hence dp: "d > 0" by simp  | 
|
964  | 
from d[unfolded dvd_def] obtain e where e: "n= d*e" by blast  | 
|
965  | 
from n dp e have ep:"e > 0" by simp  | 
|
966  | 
have "d^2 \<le> n \<or> e^2 \<le> n" using dp ep  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
967  | 
by (auto simp add: e power2_eq_square mult_le_cancel_left)  | 
| 26126 | 968  | 
moreover  | 
969  | 
      {assume h: "d^2 \<le> n"
 | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
970  | 
from H[rule_format, of d] h d have "d = 1" by blast}  | 
| 26126 | 971  | 
moreover  | 
972  | 
      {assume h: "e^2 \<le> n"
 | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
973  | 
from e have "e dvd n" unfolding dvd_def by (simp add: mult_commute)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
974  | 
with H[rule_format, of e] h have "e=1" by simp  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
975  | 
with e have "d = n" by simp}  | 
| 26126 | 976  | 
ultimately have "d=1 \<or> d=n" by blast}  | 
977  | 
ultimately have ?thesis unfolding prime_def using np n(2) by blast}  | 
|
978  | 
ultimately show ?thesis by auto  | 
|
979  | 
qed  | 
|
980  | 
lemma prime_prime_factor_sqrt:  | 
|
| 30488 | 981  | 
"prime n \<longleftrightarrow> n \<noteq> 0 \<and> n \<noteq> 1 \<and> \<not> (\<exists>p. prime p \<and> p dvd n \<and> p^2 \<le> n)"  | 
| 26126 | 982  | 
(is "?lhs \<longleftrightarrow>?rhs")  | 
983  | 
proof-  | 
|
984  | 
  {assume "n=0 \<or> n=1" hence ?thesis using prime_0 prime_1 by auto}
 | 
|
985  | 
moreover  | 
|
986  | 
  {assume n: "n\<noteq>0" "n\<noteq>1"
 | 
|
987  | 
    {assume H: ?lhs
 | 
|
| 30488 | 988  | 
from H[unfolded prime_divisor_sqrt] n  | 
| 41541 | 989  | 
have ?rhs  | 
990  | 
apply clarsimp  | 
|
991  | 
apply (erule_tac x="p" in allE)  | 
|
992  | 
apply simp  | 
|
993  | 
done  | 
|
| 26126 | 994  | 
}  | 
995  | 
moreover  | 
|
996  | 
    {assume H: ?rhs
 | 
|
997  | 
      {fix d assume d: "d dvd n" "d^2 \<le> n" "d\<noteq>1"
 | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
998  | 
from prime_factor[OF d(3)]  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
999  | 
obtain p where p: "prime p" "p dvd d" by blast  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
1000  | 
from n have np: "n > 0" by arith  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
1001  | 
from d(1) n have "d \<noteq> 0" by - (rule ccontr, auto)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
1002  | 
hence dp: "d > 0" by arith  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
1003  | 
from mult_mono[OF dvd_imp_le[OF p(2) dp] dvd_imp_le[OF p(2) dp]] d(2)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
1004  | 
have "p^2 \<le> n" unfolding power2_eq_square by arith  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
1005  | 
with H n p(1) dvd_trans[OF p(2) d(1)] have False by blast}  | 
| 26126 | 1006  | 
with n prime_divisor_sqrt have ?lhs by auto}  | 
1007  | 
ultimately have ?thesis by blast }  | 
|
1008  | 
ultimately show ?thesis by (cases "n=0 \<or> n=1", auto)  | 
|
1009  | 
qed  | 
|
1010  | 
(* Pocklington theorem. *)  | 
|
1011  | 
||
1012  | 
lemma pocklington_lemma:  | 
|
1013  | 
assumes n: "n \<ge> 2" and nqr: "n - 1 = q*r" and an: "[a^ (n - 1) = 1] (mod n)"  | 
|
1014  | 
and aq:"\<forall>p. prime p \<and> p dvd q \<longrightarrow> coprime (a^ ((n - 1) div p) - 1) n"  | 
|
1015  | 
and pp: "prime p" and pn: "p dvd n"  | 
|
1016  | 
shows "[p = 1] (mod q)"  | 
|
1017  | 
proof-  | 
|
1018  | 
from pp prime_0 prime_1 have p01: "p \<noteq> 0" "p \<noteq> 1" by - (rule ccontr, simp)+  | 
|
| 30488 | 1019  | 
from cong_1_divides[OF an, unfolded nqr, unfolded dvd_def]  | 
| 26126 | 1020  | 
obtain k where k: "a ^ (q * r) - 1 = n*k" by blast  | 
1021  | 
from pn[unfolded dvd_def] obtain l where l: "n = p*l" by blast  | 
|
1022  | 
  {assume a0: "a = 0"
 | 
|
1023  | 
hence "a^ (n - 1) = 0" using n by (simp add: power_0_left)  | 
|
1024  | 
with n an mod_less[of 1 n] have False by (simp add: power_0_left modeq_def)}  | 
|
1025  | 
hence a0: "a\<noteq>0" ..  | 
|
| 41541 | 1026  | 
from n nqr have aqr0: "a ^ (q * r) \<noteq> 0" using a0 by simp  | 
| 26126 | 1027  | 
hence "(a ^ (q * r) - 1) + 1 = a ^ (q * r)" by simp  | 
1028  | 
with k l have "a ^ (q * r) = p*l*k + 1" by simp  | 
|
1029  | 
hence "a ^ (r * q) + p * 0 = 1 + p * (l*k)" by (simp add: mult_ac)  | 
|
1030  | 
hence odq: "ord p (a^r) dvd q"  | 
|
1031  | 
unfolding ord_divides[symmetric] power_mult[symmetric] nat_mod by blast  | 
|
1032  | 
from odq[unfolded dvd_def] obtain d where d: "q = ord p (a^r) * d" by blast  | 
|
| 30488 | 1033  | 
  {assume d1: "d \<noteq> 1"
 | 
| 26126 | 1034  | 
from prime_factor[OF d1] obtain P where P: "prime P" "P dvd d" by blast  | 
1035  | 
from d dvd_mult[OF P(2), of "ord p (a^r)"] have Pq: "P dvd q" by simp  | 
|
1036  | 
from aq P(1) Pq have caP:"coprime (a^ ((n - 1) div P) - 1) n" by blast  | 
|
1037  | 
from Pq obtain s where s: "q = P*s" unfolding dvd_def by blast  | 
|
1038  | 
have P0: "P \<noteq> 0" using P(1) prime_0 by - (rule ccontr, simp)  | 
|
1039  | 
from P(2) obtain t where t: "d = P*t" unfolding dvd_def by blast  | 
|
1040  | 
from d s t P0 have s': "ord p (a^r) * t = s" by algebra  | 
|
1041  | 
have "ord p (a^r) * t*r = r * ord p (a^r) * t" by algebra  | 
|
1042  | 
hence exps: "a^(ord p (a^r) * t*r) = ((a ^ r) ^ ord p (a^r)) ^ t"  | 
|
1043  | 
by (simp only: power_mult)  | 
|
| 30488 | 1044  | 
have "[((a ^ r) ^ ord p (a^r)) ^ t= 1^t] (mod p)"  | 
| 26126 | 1045  | 
by (rule cong_exp, rule ord)  | 
| 30488 | 1046  | 
then have th: "[((a ^ r) ^ ord p (a^r)) ^ t= 1] (mod p)"  | 
| 26158 | 1047  | 
by (simp add: power_Suc0)  | 
| 26126 | 1048  | 
from cong_1_divides[OF th] exps have pd0: "p dvd a^(ord p (a^r) * t*r) - 1" by simp  | 
1049  | 
from nqr s s' have "(n - 1) div P = ord p (a^r) * t*r" using P0 by simp  | 
|
1050  | 
with caP have "coprime (a^(ord p (a^r) * t*r) - 1) n" by simp  | 
|
1051  | 
with p01 pn pd0 have False unfolding coprime by auto}  | 
|
| 30488 | 1052  | 
hence d1: "d = 1" by blast  | 
1053  | 
hence o: "ord p (a^r) = q" using d by simp  | 
|
| 26126 | 1054  | 
from pp phi_prime[of p] have phip: " \<phi> p = p - 1" by simp  | 
1055  | 
  {fix d assume d: "d dvd p" "d dvd a" "d \<noteq> 1"
 | 
|
1056  | 
from pp[unfolded prime_def] d have dp: "d = p" by blast  | 
|
1057  | 
from n have n12:"Suc (n - 2) = n - 1" by arith  | 
|
1058  | 
with divides_rexp[OF d(2)[unfolded dp], of "n - 2"]  | 
|
1059  | 
have th0: "p dvd a ^ (n - 1)" by simp  | 
|
1060  | 
from n have n0: "n \<noteq> 0" by simp  | 
|
| 30488 | 1061  | 
from d(2) an n12[symmetric] have a0: "a \<noteq> 0"  | 
| 26126 | 1062  | 
by - (rule ccontr, simp add: modeq_def)  | 
| 41541 | 1063  | 
have th1: "a^ (n - 1) \<noteq> 0" using n d(2) dp a0 by auto  | 
| 30488 | 1064  | 
from coprime_minus1[OF th1, unfolded coprime]  | 
| 26126 | 1065  | 
dvd_trans[OF pn cong_1_divides[OF an]] th0 d(3) dp  | 
1066  | 
have False by auto}  | 
|
| 30488 | 1067  | 
hence cpa: "coprime p a" using coprime by auto  | 
1068  | 
from coprime_exp[OF cpa, of r] coprime_commute  | 
|
| 26126 | 1069  | 
have arp: "coprime (a^r) p" by blast  | 
1070  | 
from fermat_little[OF arp, simplified ord_divides] o phip  | 
|
1071  | 
have "q dvd (p - 1)" by simp  | 
|
1072  | 
then obtain d where d:"p - 1 = q * d" unfolding dvd_def by blast  | 
|
1073  | 
from prime_0 pp have p0:"p \<noteq> 0" by - (rule ccontr, auto)  | 
|
1074  | 
from p0 d have "p + q * 0 = 1 + q * d" by simp  | 
|
1075  | 
with nat_mod[of p 1 q, symmetric]  | 
|
1076  | 
show ?thesis by blast  | 
|
1077  | 
qed  | 
|
1078  | 
||
1079  | 
lemma pocklington:  | 
|
1080  | 
assumes n: "n \<ge> 2" and nqr: "n - 1 = q*r" and sqr: "n \<le> q^2"  | 
|
1081  | 
and an: "[a^ (n - 1) = 1] (mod n)"  | 
|
1082  | 
and aq:"\<forall>p. prime p \<and> p dvd q \<longrightarrow> coprime (a^ ((n - 1) div p) - 1) n"  | 
|
1083  | 
shows "prime n"  | 
|
1084  | 
unfolding prime_prime_factor_sqrt[of n]  | 
|
1085  | 
proof-  | 
|
1086  | 
let ?ths = "n \<noteq> 0 \<and> n \<noteq> 1 \<and> \<not> (\<exists>p. prime p \<and> p dvd n \<and> p\<twosuperior> \<le> n)"  | 
|
1087  | 
from n have n01: "n\<noteq>0" "n\<noteq>1" by arith+  | 
|
1088  | 
  {fix p assume p: "prime p" "p dvd n" "p^2 \<le> n"
 | 
|
1089  | 
from p(3) sqr have "p^(Suc 1) \<le> q^(Suc 1)" by (simp add: power2_eq_square)  | 
|
1090  | 
hence pq: "p \<le> q" unfolding exp_mono_le .  | 
|
1091  | 
from pocklington_lemma[OF n nqr an aq p(1,2)] cong_1_divides  | 
|
1092  | 
have th: "q dvd p - 1" by blast  | 
|
1093  | 
have "p - 1 \<noteq> 0"using prime_ge_2[OF p(1)] by arith  | 
|
1094  | 
with divides_ge[OF th] pq have False by arith }  | 
|
1095  | 
with n01 show ?ths by blast  | 
|
1096  | 
qed  | 
|
1097  | 
||
1098  | 
(* Variant for application, to separate the exponentiation. *)  | 
|
1099  | 
lemma pocklington_alt:  | 
|
1100  | 
assumes n: "n \<ge> 2" and nqr: "n - 1 = q*r" and sqr: "n \<le> q^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)"  | 
|
1103  | 
shows "prime n"  | 
|
1104  | 
proof-  | 
|
1105  | 
  {fix p assume p: "prime p" "p dvd q"
 | 
|
| 30488 | 1106  | 
from aq[rule_format] p obtain b where  | 
| 26126 | 1107  | 
b: "[a^((n - 1) div p) = b] (mod n)" "coprime (b - 1) n" by blast  | 
1108  | 
    {assume a0: "a=0"
 | 
|
1109  | 
from n an have "[0 = 1] (mod n)" unfolding a0 power_0_left by auto  | 
|
1110  | 
hence False using n by (simp add: modeq_def dvd_eq_mod_eq_0[symmetric])}  | 
|
1111  | 
hence a0: "a\<noteq> 0" ..  | 
|
1112  | 
hence a1: "a \<ge> 1" by arith  | 
|
1113  | 
from one_le_power[OF a1] have ath: "1 \<le> a ^ ((n - 1) div p)" .  | 
|
1114  | 
    {assume b0: "b = 0"
 | 
|
| 30488 | 1115  | 
from p(2) nqr have "(n - 1) mod p = 0"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
1116  | 
apply (simp only: dvd_eq_mod_eq_0[symmetric]) by (rule dvd_mult2, simp)  | 
| 30488 | 1117  | 
with mod_div_equality[of "n - 1" p]  | 
1118  | 
have "(n - 1) div p * p= n - 1" by auto  | 
|
| 26126 | 1119  | 
hence eq: "(a^((n - 1) div p))^p = a^(n - 1)"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
1120  | 
by (simp only: power_mult[symmetric])  | 
| 26126 | 1121  | 
from prime_ge_2[OF p(1)] have pS: "Suc (p - 1) = p" by arith  | 
1122  | 
from b(1) have d: "n dvd a^((n - 1) div p)" unfolding b0 cong_0_divides .  | 
|
1123  | 
from divides_rexp[OF d, of "p - 1"] pS eq cong_divides[OF an] n  | 
|
1124  | 
have False by simp}  | 
|
| 30488 | 1125  | 
then have b0: "b \<noteq> 0" ..  | 
1126  | 
hence b1: "b \<ge> 1" by arith  | 
|
| 26126 | 1127  | 
from cong_coprime[OF cong_sub[OF b(1) cong_refl[of 1] ath b1]] b(2) nqr  | 
1128  | 
have "coprime (a ^ ((n - 1) div p) - 1) n" by (simp add: coprime_commute)}  | 
|
| 30488 | 1129  | 
hence th: "\<forall>p. prime p \<and> p dvd q \<longrightarrow> coprime (a ^ ((n - 1) div p) - 1) n "  | 
| 26126 | 1130  | 
by blast  | 
1131  | 
from pocklington[OF n nqr sqr an th] show ?thesis .  | 
|
1132  | 
qed  | 
|
1133  | 
||
1134  | 
(* Prime factorizations. *)  | 
|
1135  | 
||
1136  | 
definition "primefact ps n = (foldr op * ps 1 = n \<and> (\<forall>p\<in> set ps. prime p))"  | 
|
1137  | 
||
1138  | 
lemma primefact: assumes n: "n \<noteq> 0"  | 
|
1139  | 
shows "\<exists>ps. primefact ps n"  | 
|
1140  | 
using n  | 
|
1141  | 
proof(induct n rule: nat_less_induct)  | 
|
1142  | 
fix n assume H: "\<forall>m<n. m \<noteq> 0 \<longrightarrow> (\<exists>ps. primefact ps m)" and n: "n\<noteq>0"  | 
|
1143  | 
let ?ths = "\<exists>ps. primefact ps n"  | 
|
| 30488 | 1144  | 
  {assume "n = 1"
 | 
| 26126 | 1145  | 
hence "primefact [] n" by (simp add: primefact_def)  | 
1146  | 
hence ?ths by blast }  | 
|
1147  | 
moreover  | 
|
1148  | 
  {assume n1: "n \<noteq> 1"
 | 
|
1149  | 
with n have n2: "n \<ge> 2" by arith  | 
|
1150  | 
from prime_factor[OF n1] obtain p where p: "prime p" "p dvd n" by blast  | 
|
1151  | 
from p(2) obtain m where m: "n = p*m" unfolding dvd_def by blast  | 
|
1152  | 
from n m have m0: "m > 0" "m\<noteq>0" by auto  | 
|
1153  | 
from prime_ge_2[OF p(1)] have "1 < p" by arith  | 
|
1154  | 
with m0 m have mn: "m < n" by auto  | 
|
1155  | 
from H[rule_format, OF mn m0(2)] obtain ps where ps: "primefact ps m" ..  | 
|
1156  | 
from ps m p(1) have "primefact (p#ps) n" by (simp add: primefact_def)  | 
|
1157  | 
hence ?ths by blast}  | 
|
1158  | 
ultimately show ?ths by blast  | 
|
1159  | 
qed  | 
|
1160  | 
||
| 30488 | 1161  | 
lemma primefact_contains:  | 
| 26126 | 1162  | 
assumes pf: "primefact ps n" and p: "prime p" and pn: "p dvd n"  | 
1163  | 
shows "p \<in> set ps"  | 
|
1164  | 
using pf p pn  | 
|
1165  | 
proof(induct ps arbitrary: p n)  | 
|
1166  | 
case Nil thus ?case by (auto simp add: primefact_def)  | 
|
1167  | 
next  | 
|
1168  | 
case (Cons q qs p n)  | 
|
| 30488 | 1169  | 
from Cons.prems[unfolded primefact_def]  | 
| 26126 | 1170  | 
have q: "prime q" "q * foldr op * qs 1 = n" "\<forall>p \<in>set qs. prime p" and p: "prime p" "p dvd q * foldr op * qs 1" by simp_all  | 
1171  | 
  {assume "p dvd q"
 | 
|
1172  | 
with p(1) q(1) have "p = q" unfolding prime_def by auto  | 
|
1173  | 
hence ?case by simp}  | 
|
1174  | 
moreover  | 
|
1175  | 
  { assume h: "p dvd foldr op * qs 1"
 | 
|
| 30488 | 1176  | 
from q(3) have pqs: "primefact qs (foldr op * qs 1)"  | 
| 26126 | 1177  | 
by (simp add: primefact_def)  | 
1178  | 
from Cons.hyps[OF pqs p(1) h] have ?case by simp}  | 
|
1179  | 
ultimately show ?case using prime_divprod[OF p] by blast  | 
|
1180  | 
qed  | 
|
1181  | 
||
| 37602 | 1182  | 
lemma primefact_variant: "primefact ps n \<longleftrightarrow> foldr op * ps 1 = n \<and> list_all prime ps"  | 
1183  | 
by (auto simp add: primefact_def list_all_iff)  | 
|
| 26126 | 1184  | 
|
1185  | 
(* Variant of Lucas theorem. *)  | 
|
1186  | 
||
1187  | 
lemma lucas_primefact:  | 
|
| 30488 | 1188  | 
assumes n: "n \<ge> 2" and an: "[a^(n - 1) = 1] (mod n)"  | 
1189  | 
and psn: "foldr op * ps 1 = n - 1"  | 
|
| 26126 | 1190  | 
and psp: "list_all (\<lambda>p. prime p \<and> \<not> [a^((n - 1) div p) = 1] (mod n)) ps"  | 
1191  | 
shows "prime n"  | 
|
1192  | 
proof-  | 
|
1193  | 
  {fix p assume p: "prime p" "p dvd n - 1" "[a ^ ((n - 1) div p) = 1] (mod n)"
 | 
|
| 30488 | 1194  | 
from psn psp have psn1: "primefact ps (n - 1)"  | 
| 26126 | 1195  | 
by (auto simp add: list_all_iff primefact_variant)  | 
1196  | 
from p(3) primefact_contains[OF psn1 p(1,2)] psp  | 
|
1197  | 
have False by (induct ps, auto)}  | 
|
1198  | 
with lucas[OF n an] show ?thesis by blast  | 
|
1199  | 
qed  | 
|
1200  | 
||
1201  | 
(* Variant of Pocklington theorem. *)  | 
|
1202  | 
||
1203  | 
lemma mod_le: assumes n: "n \<noteq> (0::nat)" shows "m mod n \<le> m"  | 
|
1204  | 
proof-  | 
|
1205  | 
from mod_div_equality[of m n]  | 
|
| 30488 | 1206  | 
have "\<exists>x. x + m mod n = m" by blast  | 
| 26126 | 1207  | 
then show ?thesis by auto  | 
1208  | 
qed  | 
|
| 30488 | 1209  | 
|
| 26126 | 1210  | 
|
1211  | 
lemma pocklington_primefact:  | 
|
1212  | 
assumes n: "n \<ge> 2" and qrn: "q*r = n - 1" and nq2: "n \<le> q^2"  | 
|
| 30488 | 1213  | 
and arnb: "(a^r) mod n = b" and psq: "foldr op * ps 1 = q"  | 
| 26126 | 1214  | 
and bqn: "(b^q) mod n = 1"  | 
1215  | 
and psp: "list_all (\<lambda>p. prime p \<and> coprime ((b^(q div p)) mod n - 1) n) ps"  | 
|
1216  | 
shows "prime n"  | 
|
1217  | 
proof-  | 
|
1218  | 
from bqn psp qrn  | 
|
1219  | 
have bqn: "a ^ (n - 1) mod n = 1"  | 
|
| 30488 | 1220  | 
and psp: "list_all (\<lambda>p. prime p \<and> coprime (a^(r *(q div p)) mod n - 1) n) ps" unfolding arnb[symmetric] power_mod  | 
| 29667 | 1221  | 
by (simp_all add: power_mult[symmetric] algebra_simps)  | 
| 26126 | 1222  | 
from n have n0: "n > 0" by arith  | 
1223  | 
from mod_div_equality[of "a^(n - 1)" n]  | 
|
1224  | 
mod_less_divisor[OF n0, of "a^(n - 1)"]  | 
|
| 30488 | 1225  | 
have an1: "[a ^ (n - 1) = 1] (mod n)"  | 
| 26126 | 1226  | 
unfolding nat_mod bqn  | 
1227  | 
apply -  | 
|
1228  | 
apply (rule exI[where x="0"])  | 
|
1229  | 
apply (rule exI[where x="a^(n - 1) div n"])  | 
|
| 29667 | 1230  | 
by (simp add: algebra_simps)  | 
| 26126 | 1231  | 
  {fix p assume p: "prime p" "p dvd q"
 | 
1232  | 
from psp psq have pfpsq: "primefact ps q"  | 
|
1233  | 
by (auto simp add: primefact_variant list_all_iff)  | 
|
| 30488 | 1234  | 
from psp primefact_contains[OF pfpsq p]  | 
| 26126 | 1235  | 
have p': "coprime (a ^ (r * (q div p)) mod n - 1) n"  | 
1236  | 
by (simp add: list_all_iff)  | 
|
1237  | 
from prime_ge_2[OF p(1)] have p01: "p \<noteq> 0" "p \<noteq> 1" "p =Suc(p - 1)" by arith+  | 
|
| 30488 | 1238  | 
from div_mult1_eq[of r q p] p(2)  | 
| 26126 | 1239  | 
have eq1: "r* (q div p) = (n - 1) div p"  | 
1240  | 
unfolding qrn[symmetric] dvd_eq_mod_eq_0 by (simp add: mult_commute)  | 
|
1241  | 
have ath: "\<And>a (b::nat). a <= b \<Longrightarrow> a \<noteq> 0 ==> 1 <= a \<and> 1 <= b" by arith  | 
|
1242  | 
from n0 have n00: "n \<noteq> 0" by arith  | 
|
1243  | 
from mod_le[OF n00]  | 
|
1244  | 
have th10: "a ^ ((n - 1) div p) mod n \<le> a ^ ((n - 1) div p)" .  | 
|
1245  | 
    {assume "a ^ ((n - 1) div p) mod n = 0"
 | 
|
1246  | 
then obtain s where s: "a ^ ((n - 1) div p) = n*s"  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
1247  | 
unfolding mod_eq_0_iff by blast  | 
| 26126 | 1248  | 
hence eq0: "(a^((n - 1) div p))^p = (n*s)^p" by simp  | 
1249  | 
from qrn[symmetric] have qn1: "q dvd n - 1" unfolding dvd_def by auto  | 
|
1250  | 
from dvd_trans[OF p(2) qn1] div_mod_equality'[of "n - 1" p]  | 
|
| 30488 | 1251  | 
have npp: "(n - 1) div p * p = n - 1" by (simp add: dvd_eq_mod_eq_0)  | 
| 26126 | 1252  | 
with eq0 have "a^ (n - 1) = (n*s)^p"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32479 
diff
changeset
 | 
1253  | 
by (simp add: power_mult[symmetric])  | 
| 26126 | 1254  | 
hence "1 = (n*s)^(Suc (p - 1)) mod n" using bqn p01 by simp  | 
| 
28668
 
e79e196039a1
fixed and reactivated HOL/Library/Pocklington.thy -- by Mark Hillebrand;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
1255  | 
also have "\<dots> = 0" by (simp add: mult_assoc)  | 
| 26126 | 1256  | 
finally have False by simp }  | 
| 30488 | 1257  | 
then have th11: "a ^ ((n - 1) div p) mod n \<noteq> 0" by auto  | 
1258  | 
have th1: "[a ^ ((n - 1) div p) mod n = a ^ ((n - 1) div p)] (mod n)"  | 
|
1259  | 
unfolding modeq_def by simp  | 
|
| 26126 | 1260  | 
from cong_sub[OF th1 cong_refl[of 1]] ath[OF th10 th11]  | 
1261  | 
have th: "[a ^ ((n - 1) div p) mod n - 1 = a ^ ((n - 1) div p) - 1] (mod n)"  | 
|
| 30488 | 1262  | 
by blast  | 
1263  | 
from cong_coprime[OF th] p'[unfolded eq1]  | 
|
| 26126 | 1264  | 
have "coprime (a ^ ((n - 1) div p) - 1) n" by (simp add: coprime_commute) }  | 
1265  | 
with pocklington[OF n qrn[symmetric] nq2 an1]  | 
|
| 30488 | 1266  | 
show ?thesis by blast  | 
| 26126 | 1267  | 
qed  | 
1268  | 
||
1269  | 
end  |