author | wenzelm |
Mon, 15 Feb 2016 14:55:44 +0100 | |
changeset 62337 | d3996d5873dd |
parent 61694 | 6571c78c9667 |
child 64272 | f76b6dda2e56 |
permissions | -rw-r--r-- |
38159 | 1 |
(* Title: HOL/Old_Number_Theory/WilsonRuss.thy |
2 |
Author: Thomas M. Rasmussen |
|
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
3 |
Copyright 2000 University of Cambridge |
9508
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
paulson
parents:
diff
changeset
|
4 |
*) |
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
paulson
parents:
diff
changeset
|
5 |
|
61382 | 6 |
section \<open>Wilson's Theorem according to Russinoff\<close> |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
7 |
|
38159 | 8 |
theory WilsonRuss |
9 |
imports EulerFermat |
|
10 |
begin |
|
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
11 |
|
61382 | 12 |
text \<open> |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
13 |
Wilson's Theorem following quite closely Russinoff's approach |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
14 |
using Boyer-Moore (using finite sets instead of lists, though). |
61382 | 15 |
\<close> |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
16 |
|
61382 | 17 |
subsection \<open>Definitions and lemmas\<close> |
9508
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
paulson
parents:
diff
changeset
|
18 |
|
38159 | 19 |
definition inv :: "int => int => int" |
20 |
where "inv p a = (a^(nat (p - 2))) mod p" |
|
19670 | 21 |
|
38159 | 22 |
fun wset :: "int \<Rightarrow> int => int set" where |
35440 | 23 |
"wset a p = |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
24 |
(if 1 < a then |
35440 | 25 |
let ws = wset (a - 1) p |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
26 |
in (if a \<in> ws then ws else insert a (insert (inv p a) ws)) else {})" |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
27 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
28 |
|
61382 | 29 |
text \<open>\medskip @{term [source] inv}\<close> |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
30 |
|
13524 | 31 |
lemma inv_is_inv_aux: "1 < m ==> Suc (nat (m - 2)) = nat (m - 1)" |
61694
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61649
diff
changeset
|
32 |
by simp |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
33 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
34 |
lemma inv_is_inv: |
16663 | 35 |
"zprime p \<Longrightarrow> 0 < a \<Longrightarrow> a < p ==> [a * inv p a = 1] (mod p)" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
36 |
apply (unfold inv_def) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
37 |
apply (subst zcong_zmod) |
47163 | 38 |
apply (subst mod_mult_right_eq [symmetric]) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
39 |
apply (subst zcong_zmod [symmetric]) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
40 |
apply (subst power_Suc [symmetric]) |
61694
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61649
diff
changeset
|
41 |
using Little_Fermat inv_is_inv_aux zdvd_not_zless apply auto |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
42 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
43 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
44 |
lemma inv_distinct: |
16663 | 45 |
"zprime p \<Longrightarrow> 1 < a \<Longrightarrow> a < p - 1 ==> a \<noteq> inv p a" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
46 |
apply safe |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
47 |
apply (cut_tac a = a and p = p in zcong_square) |
13833 | 48 |
apply (cut_tac [3] a = a and p = p in inv_is_inv, auto) |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
49 |
apply (subgoal_tac "a = 1") |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
50 |
apply (rule_tac [2] m = p in zcong_zless_imp_eq) |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
51 |
apply (subgoal_tac [7] "a = p - 1") |
13833 | 52 |
apply (rule_tac [8] m = p in zcong_zless_imp_eq, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
53 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
54 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
55 |
lemma inv_not_0: |
16663 | 56 |
"zprime p \<Longrightarrow> 1 < a \<Longrightarrow> a < p - 1 ==> inv p a \<noteq> 0" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
57 |
apply safe |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
58 |
apply (cut_tac a = a and p = p in inv_is_inv) |
13833 | 59 |
apply (unfold zcong_def, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
60 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
61 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
62 |
lemma inv_not_1: |
16663 | 63 |
"zprime p \<Longrightarrow> 1 < a \<Longrightarrow> a < p - 1 ==> inv p a \<noteq> 1" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
64 |
apply safe |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
65 |
apply (cut_tac a = a and p = p in inv_is_inv) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
66 |
prefer 4 |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
67 |
apply simp |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
68 |
apply (subgoal_tac "a = 1") |
13833 | 69 |
apply (rule_tac [2] zcong_zless_imp_eq, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
70 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
71 |
|
19670 | 72 |
lemma inv_not_p_minus_1_aux: |
73 |
"[a * (p - 1) = 1] (mod p) = [a = p - 1] (mod p)" |
|
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
74 |
apply (unfold zcong_def) |
44766 | 75 |
apply (simp add: diff_diff_eq diff_diff_eq2 right_diff_distrib) |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
76 |
apply (rule_tac s = "p dvd -((a + 1) + (p * -a))" in trans) |
35048 | 77 |
apply (simp add: algebra_simps) |
30042 | 78 |
apply (subst dvd_minus_iff) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
79 |
apply (subst zdvd_reduce) |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
80 |
apply (rule_tac s = "p dvd (a + 1) + (p * -1)" in trans) |
13833 | 81 |
apply (subst zdvd_reduce, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
82 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
83 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
84 |
lemma inv_not_p_minus_1: |
16663 | 85 |
"zprime p \<Longrightarrow> 1 < a \<Longrightarrow> a < p - 1 ==> inv p a \<noteq> p - 1" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
86 |
apply safe |
13833 | 87 |
apply (cut_tac a = a and p = p in inv_is_inv, auto) |
13524 | 88 |
apply (simp add: inv_not_p_minus_1_aux) |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
89 |
apply (subgoal_tac "a = p - 1") |
13833 | 90 |
apply (rule_tac [2] zcong_zless_imp_eq, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
91 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
92 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
93 |
lemma inv_g_1: |
16663 | 94 |
"zprime p \<Longrightarrow> 1 < a \<Longrightarrow> a < p - 1 ==> 1 < inv p a" |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
95 |
apply (case_tac "0\<le> inv p a") |
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
96 |
apply (subgoal_tac "inv p a \<noteq> 1") |
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
97 |
apply (subgoal_tac "inv p a \<noteq> 0") |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
98 |
apply (subst order_less_le) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
99 |
apply (subst zle_add1_eq_le [symmetric]) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
100 |
apply (subst order_less_le) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
101 |
apply (rule_tac [2] inv_not_0) |
13833 | 102 |
apply (rule_tac [5] inv_not_1, auto) |
103 |
apply (unfold inv_def zprime_def, simp) |
|
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
104 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
105 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
106 |
lemma inv_less_p_minus_1: |
16663 | 107 |
"zprime p \<Longrightarrow> 1 < a \<Longrightarrow> a < p - 1 ==> inv p a < p - 1" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
108 |
apply (case_tac "inv p a < p") |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
109 |
apply (subst order_less_le) |
13833 | 110 |
apply (simp add: inv_not_p_minus_1, auto) |
111 |
apply (unfold inv_def zprime_def, simp) |
|
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
112 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
113 |
|
13524 | 114 |
lemma inv_inv_aux: "5 \<le> p ==> |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
115 |
nat (p - 2) * nat (p - 2) = Suc (nat (p - 1) * nat (p - 3))" |
61694
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61649
diff
changeset
|
116 |
apply (subst of_nat_eq_iff [where 'a = int, symmetric]) |
44766 | 117 |
apply (simp add: left_diff_distrib right_diff_distrib) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
118 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
119 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
120 |
lemma zcong_zpower_zmult: |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
121 |
"[x^y = 1] (mod p) \<Longrightarrow> [x^(y * z) = 1] (mod p)" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
122 |
apply (induct z) |
44766 | 123 |
apply (auto simp add: power_add) |
15236
f289e8ba2bb3
Proofs needed to be updated because induction now preserves name of
nipkow
parents:
15197
diff
changeset
|
124 |
apply (subgoal_tac "zcong (x^y * x^(y * z)) (1 * 1) p") |
13833 | 125 |
apply (rule_tac [2] zcong_zmult, simp_all) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
126 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
127 |
|
16663 | 128 |
lemma inv_inv: "zprime p \<Longrightarrow> |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
129 |
5 \<le> p \<Longrightarrow> 0 < a \<Longrightarrow> a < p ==> inv p (inv p a) = a" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
130 |
apply (unfold inv_def) |
47164 | 131 |
apply (subst power_mod) |
61649
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61382
diff
changeset
|
132 |
apply (subst power_mult [symmetric]) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
133 |
apply (rule zcong_zless_imp_eq) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
134 |
prefer 5 |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
135 |
apply (subst zcong_zmod) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
136 |
apply (subst mod_mod_trivial) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
137 |
apply (subst zcong_zmod [symmetric]) |
13524 | 138 |
apply (subst inv_inv_aux) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
139 |
apply (subgoal_tac [2] |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32479
diff
changeset
|
140 |
"zcong (a * a^(nat (p - 1) * nat (p - 3))) (a * 1) p") |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
141 |
apply (rule_tac [3] zcong_zmult) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
142 |
apply (rule_tac [4] zcong_zpower_zmult) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
143 |
apply (erule_tac [4] Little_Fermat) |
13833 | 144 |
apply (rule_tac [4] zdvd_not_zless, simp_all) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
145 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
146 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
147 |
|
61382 | 148 |
text \<open>\medskip @{term wset}\<close> |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
149 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
150 |
declare wset.simps [simp del] |
9508
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
paulson
parents:
diff
changeset
|
151 |
|
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
152 |
lemma wset_induct: |
18369 | 153 |
assumes "!!a p. P {} a p" |
19670 | 154 |
and "!!a p. 1 < (a::int) \<Longrightarrow> |
35440 | 155 |
P (wset (a - 1) p) (a - 1) p ==> P (wset a p) a p" |
156 |
shows "P (wset u v) u v" |
|
157 |
apply (rule wset.induct) |
|
158 |
apply (case_tac "1 < a") |
|
159 |
apply (rule assms) |
|
160 |
apply (simp_all add: wset.simps assms) |
|
18369 | 161 |
done |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
162 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
163 |
lemma wset_mem_imp_or [rule_format]: |
35440 | 164 |
"1 < a \<Longrightarrow> b \<notin> wset (a - 1) p |
165 |
==> b \<in> wset a p --> b = a \<or> b = inv p a" |
|
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
166 |
apply (subst wset.simps) |
13833 | 167 |
apply (unfold Let_def, simp) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
168 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
169 |
|
35440 | 170 |
lemma wset_mem_mem [simp]: "1 < a ==> a \<in> wset a p" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
171 |
apply (subst wset.simps) |
13833 | 172 |
apply (unfold Let_def, simp) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
173 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
174 |
|
35440 | 175 |
lemma wset_subset: "1 < a \<Longrightarrow> b \<in> wset (a - 1) p ==> b \<in> wset a p" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
176 |
apply (subst wset.simps) |
13833 | 177 |
apply (unfold Let_def, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
178 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
179 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
180 |
lemma wset_g_1 [rule_format]: |
35440 | 181 |
"zprime p --> a < p - 1 --> b \<in> wset a p --> 1 < b" |
13833 | 182 |
apply (induct a p rule: wset_induct, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
183 |
apply (case_tac "b = a") |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
184 |
apply (case_tac [2] "b = inv p a") |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
185 |
apply (subgoal_tac [3] "b = a \<or> b = inv p a") |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
186 |
apply (rule_tac [4] wset_mem_imp_or) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
187 |
prefer 2 |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
188 |
apply simp |
13833 | 189 |
apply (rule inv_g_1, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
190 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
191 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
192 |
lemma wset_less [rule_format]: |
35440 | 193 |
"zprime p --> a < p - 1 --> b \<in> wset a p --> b < p - 1" |
13833 | 194 |
apply (induct a p rule: wset_induct, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
195 |
apply (case_tac "b = a") |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
196 |
apply (case_tac [2] "b = inv p a") |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
197 |
apply (subgoal_tac [3] "b = a \<or> b = inv p a") |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
198 |
apply (rule_tac [4] wset_mem_imp_or) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
199 |
prefer 2 |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
200 |
apply simp |
13833 | 201 |
apply (rule inv_less_p_minus_1, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
202 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
203 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
204 |
lemma wset_mem [rule_format]: |
16663 | 205 |
"zprime p --> |
35440 | 206 |
a < p - 1 --> 1 < b --> b \<le> a --> b \<in> wset a p" |
13833 | 207 |
apply (induct a p rule: wset.induct, auto) |
15197 | 208 |
apply (rule_tac wset_subset) |
209 |
apply (simp (no_asm_simp)) |
|
210 |
apply auto |
|
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
211 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
212 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
213 |
lemma wset_mem_inv_mem [rule_format]: |
35440 | 214 |
"zprime p --> 5 \<le> p --> a < p - 1 --> b \<in> wset a p |
215 |
--> inv p b \<in> wset a p" |
|
13833 | 216 |
apply (induct a p rule: wset_induct, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
217 |
apply (case_tac "b = a") |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
218 |
apply (subst wset.simps) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
219 |
apply (unfold Let_def) |
13833 | 220 |
apply (rule_tac [3] wset_subset, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
221 |
apply (case_tac "b = inv p a") |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
222 |
apply (simp (no_asm_simp)) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
223 |
apply (subst inv_inv) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
224 |
apply (subgoal_tac [6] "b = a \<or> b = inv p a") |
13833 | 225 |
apply (rule_tac [7] wset_mem_imp_or, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
226 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
227 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
228 |
lemma wset_inv_mem_mem: |
16663 | 229 |
"zprime p \<Longrightarrow> 5 \<le> p \<Longrightarrow> a < p - 1 \<Longrightarrow> 1 < b \<Longrightarrow> b < p - 1 |
35440 | 230 |
\<Longrightarrow> inv p b \<in> wset a p \<Longrightarrow> b \<in> wset a p" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
231 |
apply (rule_tac s = "inv p (inv p b)" and t = b in subst) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
232 |
apply (rule_tac [2] wset_mem_inv_mem) |
13833 | 233 |
apply (rule inv_inv, simp_all) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
234 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
235 |
|
35440 | 236 |
lemma wset_fin: "finite (wset a p)" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
237 |
apply (induct a p rule: wset_induct) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
238 |
prefer 2 |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
239 |
apply (subst wset.simps) |
13833 | 240 |
apply (unfold Let_def, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
241 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
242 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
243 |
lemma wset_zcong_prod_1 [rule_format]: |
16663 | 244 |
"zprime p --> |
35440 | 245 |
5 \<le> p --> a < p - 1 --> [(\<Prod>x\<in>wset a p. x) = 1] (mod p)" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
246 |
apply (induct a p rule: wset_induct) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
247 |
prefer 2 |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
248 |
apply (subst wset.simps) |
35440 | 249 |
apply (auto, unfold Let_def, auto) |
57418 | 250 |
apply (subst setprod.insert) |
61382 | 251 |
apply (tactic \<open>stac @{context} @{thm setprod.insert} 3\<close>) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
252 |
apply (subgoal_tac [5] |
35440 | 253 |
"zcong (a * inv p a * (\<Prod>x\<in>wset (a - 1) p. x)) (1 * 1) p") |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
254 |
prefer 5 |
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57418
diff
changeset
|
255 |
apply (simp add: mult.assoc) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
256 |
apply (rule_tac [5] zcong_zmult) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
257 |
apply (rule_tac [5] inv_is_inv) |
42793 | 258 |
apply (tactic "clarify_tac @{context} 4") |
35440 | 259 |
apply (subgoal_tac [4] "a \<in> wset (a - 1) p") |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
260 |
apply (rule_tac [5] wset_inv_mem_mem) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
261 |
apply (simp_all add: wset_fin) |
13833 | 262 |
apply (rule inv_distinct, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
263 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
264 |
|
35440 | 265 |
lemma d22set_eq_wset: "zprime p ==> d22set (p - 2) = wset (p - 2) p" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
266 |
apply safe |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
267 |
apply (erule wset_mem) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
268 |
apply (rule_tac [2] d22set_g_1) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
269 |
apply (rule_tac [3] d22set_le) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
270 |
apply (rule_tac [4] d22set_mem) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
271 |
apply (erule_tac [4] wset_g_1) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
272 |
prefer 6 |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
273 |
apply (subst zle_add1_eq_le [symmetric]) |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
274 |
apply (subgoal_tac "p - 2 + 1 = p - 1") |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
275 |
apply (simp (no_asm_simp)) |
13833 | 276 |
apply (erule wset_less, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
277 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
278 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
279 |
|
61382 | 280 |
subsection \<open>Wilson\<close> |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
281 |
|
16663 | 282 |
lemma prime_g_5: "zprime p \<Longrightarrow> p \<noteq> 2 \<Longrightarrow> p \<noteq> 3 ==> 5 \<le> p" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
283 |
apply (unfold zprime_def dvd_def) |
13833 | 284 |
apply (case_tac "p = 4", auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
285 |
apply (rule notE) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
286 |
prefer 2 |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
287 |
apply assumption |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
288 |
apply (simp (no_asm)) |
13833 | 289 |
apply (rule_tac x = 2 in exI) |
290 |
apply (safe, arith) |
|
291 |
apply (rule_tac x = 2 in exI, auto) |
|
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
292 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
293 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
294 |
theorem Wilson_Russ: |
16663 | 295 |
"zprime p ==> [zfact (p - 1) = -1] (mod p)" |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
296 |
apply (subgoal_tac "[(p - 1) * zfact (p - 2) = -1 * 1] (mod p)") |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
297 |
apply (rule_tac [2] zcong_zmult) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
298 |
apply (simp only: zprime_def) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
299 |
apply (subst zfact.simps) |
13833 | 300 |
apply (rule_tac t = "p - 1 - 1" and s = "p - 2" in subst, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
301 |
apply (simp only: zcong_def) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
302 |
apply (simp (no_asm_simp)) |
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
303 |
apply (case_tac "p = 2") |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
304 |
apply (simp add: zfact.simps) |
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
305 |
apply (case_tac "p = 3") |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
306 |
apply (simp add: zfact.simps) |
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
307 |
apply (subgoal_tac "5 \<le> p") |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
308 |
apply (erule_tac [2] prime_g_5) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
309 |
apply (subst d22set_prod_zfact [symmetric]) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
310 |
apply (subst d22set_eq_wset) |
13833 | 311 |
apply (rule_tac [2] wset_zcong_prod_1, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
9508
diff
changeset
|
312 |
done |
9508
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
paulson
parents:
diff
changeset
|
313 |
|
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
paulson
parents:
diff
changeset
|
314 |
end |