| author | wenzelm | 
| Fri, 15 Dec 2006 00:08:16 +0100 | |
| changeset 21861 | a972053ed147 | 
| parent 21404 | eb85850d3eb7 | 
| child 23894 | 1a4167d761ac | 
| permissions | -rw-r--r-- | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 1 | (* Title: HOL/NumberTheory/WilsonRuss.thy | 
| 9508 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 2 | ID: $Id$ | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 3 | Author: Thomas M. Rasmussen | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 4 | Copyright 2000 University of Cambridge | 
| 9508 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 5 | *) | 
| 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 6 | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 7 | header {* Wilson's Theorem according to Russinoff *}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 8 | |
| 16417 | 9 | theory WilsonRuss imports EulerFermat begin | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 10 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 11 | text {*
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 12 | Wilson's Theorem following quite closely Russinoff's approach | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 13 | using Boyer-Moore (using finite sets instead of lists, though). | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 14 | *} | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 15 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 16 | subsection {* Definitions and lemmas *}
 | 
| 9508 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 17 | |
| 19670 | 18 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19670diff
changeset | 19 | inv :: "int => int => int" where | 
| 19670 | 20 | "inv p a = (a^(nat (p - 2))) mod p" | 
| 21 | ||
| 9508 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 22 | consts | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 23 | wset :: "int * int => int set" | 
| 9508 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 24 | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 25 | recdef wset | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 26 | "measure ((\<lambda>(a, p). nat a) :: int * int => nat)" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 27 | "wset (a, p) = | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11704diff
changeset | 28 | (if 1 < a then | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11704diff
changeset | 29 | let ws = wset (a - 1, p) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 30 |       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: 
9508diff
changeset | 31 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 32 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 33 | text {* \medskip @{term [source] inv} *}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 34 | |
| 13524 | 35 | lemma inv_is_inv_aux: "1 < m ==> Suc (nat (m - 2)) = nat (m - 1)" | 
| 13833 | 36 | by (subst int_int_eq [symmetric], auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 37 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 38 | lemma inv_is_inv: | 
| 16663 | 39 | "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: 
9508diff
changeset | 40 | apply (unfold inv_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 41 | apply (subst zcong_zmod) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 42 | apply (subst zmod_zmult1_eq [symmetric]) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 43 | apply (subst zcong_zmod [symmetric]) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 44 | apply (subst power_Suc [symmetric]) | 
| 13524 | 45 | apply (subst inv_is_inv_aux) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 46 | apply (erule_tac [2] Little_Fermat) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 47 | apply (erule_tac [2] zdvd_not_zless) | 
| 13833 | 48 | apply (unfold zprime_def, auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 49 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 50 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 51 | lemma inv_distinct: | 
| 16663 | 52 | "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: 
9508diff
changeset | 53 | apply safe | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 54 | apply (cut_tac a = a and p = p in zcong_square) | 
| 13833 | 55 | 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: 
11704diff
changeset | 56 | apply (subgoal_tac "a = 1") | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 57 | 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: 
11704diff
changeset | 58 | apply (subgoal_tac [7] "a = p - 1") | 
| 13833 | 59 | 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: 
9508diff
changeset | 60 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 61 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 62 | lemma inv_not_0: | 
| 16663 | 63 | "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: 
9508diff
changeset | 64 | apply safe | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 65 | apply (cut_tac a = a and p = p in inv_is_inv) | 
| 13833 | 66 | apply (unfold zcong_def, auto) | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11704diff
changeset | 67 | apply (subgoal_tac "\<not> p dvd 1") | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 68 | apply (rule_tac [2] zdvd_not_zless) | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11704diff
changeset | 69 | apply (subgoal_tac "p dvd 1") | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 70 | prefer 2 | 
| 13833 | 71 | apply (subst zdvd_zminus_iff [symmetric], auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 72 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 73 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 74 | lemma inv_not_1: | 
| 16663 | 75 | "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: 
9508diff
changeset | 76 | apply safe | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 77 | 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: 
9508diff
changeset | 78 | prefer 4 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 79 | apply simp | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11704diff
changeset | 80 | apply (subgoal_tac "a = 1") | 
| 13833 | 81 | apply (rule_tac [2] zcong_zless_imp_eq, auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 82 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 83 | |
| 19670 | 84 | lemma inv_not_p_minus_1_aux: | 
| 85 | "[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: 
9508diff
changeset | 86 | apply (unfold zcong_def) | 
| 14738 | 87 | apply (simp add: OrderedGroup.diff_diff_eq diff_diff_eq2 zdiff_zmult_distrib2) | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11704diff
changeset | 88 | apply (rule_tac s = "p dvd -((a + 1) + (p * -a))" in trans) | 
| 14271 | 89 | apply (simp add: mult_commute) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 90 | apply (subst zdvd_zminus_iff) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 91 | apply (subst zdvd_reduce) | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11704diff
changeset | 92 | apply (rule_tac s = "p dvd (a + 1) + (p * -1)" in trans) | 
| 13833 | 93 | apply (subst zdvd_reduce, auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 94 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 95 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 96 | lemma inv_not_p_minus_1: | 
| 16663 | 97 | "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: 
9508diff
changeset | 98 | apply safe | 
| 13833 | 99 | apply (cut_tac a = a and p = p in inv_is_inv, auto) | 
| 13524 | 100 | 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: 
11704diff
changeset | 101 | apply (subgoal_tac "a = p - 1") | 
| 13833 | 102 | apply (rule_tac [2] zcong_zless_imp_eq, auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 103 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 104 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 105 | lemma inv_g_1: | 
| 16663 | 106 | "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: 
11704diff
changeset | 107 | 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: 
11704diff
changeset | 108 | 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: 
11704diff
changeset | 109 | apply (subgoal_tac "inv p a \<noteq> 0") | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 110 | apply (subst order_less_le) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 111 | apply (subst zle_add1_eq_le [symmetric]) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 112 | apply (subst order_less_le) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 113 | apply (rule_tac [2] inv_not_0) | 
| 13833 | 114 | apply (rule_tac [5] inv_not_1, auto) | 
| 115 | apply (unfold inv_def zprime_def, simp) | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 116 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 117 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 118 | lemma inv_less_p_minus_1: | 
| 16663 | 119 | "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: 
9508diff
changeset | 120 | apply (case_tac "inv p a < p") | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 121 | apply (subst order_less_le) | 
| 13833 | 122 | apply (simp add: inv_not_p_minus_1, auto) | 
| 123 | apply (unfold inv_def zprime_def, simp) | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 124 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 125 | |
| 13524 | 126 | 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: 
11704diff
changeset | 127 | nat (p - 2) * nat (p - 2) = Suc (nat (p - 1) * nat (p - 3))" | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 128 | apply (subst int_int_eq [symmetric]) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 129 | apply (simp add: zmult_int [symmetric]) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 130 | apply (simp add: zdiff_zmult_distrib zdiff_zmult_distrib2) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 131 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 132 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 133 | lemma zcong_zpower_zmult: | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11704diff
changeset | 134 | "[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: 
9508diff
changeset | 135 | apply (induct z) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 136 | apply (auto simp add: zpower_zadd_distrib) | 
| 15236 
f289e8ba2bb3
Proofs needed to be updated because induction now preserves name of
 nipkow parents: 
15197diff
changeset | 137 | apply (subgoal_tac "zcong (x^y * x^(y * z)) (1 * 1) p") | 
| 13833 | 138 | apply (rule_tac [2] zcong_zmult, simp_all) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 139 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 140 | |
| 16663 | 141 | lemma inv_inv: "zprime p \<Longrightarrow> | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11704diff
changeset | 142 | 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: 
9508diff
changeset | 143 | apply (unfold inv_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 144 | apply (subst zpower_zmod) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 145 | apply (subst zpower_zpower) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 146 | apply (rule zcong_zless_imp_eq) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 147 | prefer 5 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 148 | apply (subst zcong_zmod) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 149 | apply (subst mod_mod_trivial) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 150 | apply (subst zcong_zmod [symmetric]) | 
| 13524 | 151 | apply (subst inv_inv_aux) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 152 | apply (subgoal_tac [2] | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11704diff
changeset | 153 | "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: 
9508diff
changeset | 154 | apply (rule_tac [3] zcong_zmult) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 155 | apply (rule_tac [4] zcong_zpower_zmult) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 156 | apply (erule_tac [4] Little_Fermat) | 
| 13833 | 157 | apply (rule_tac [4] zdvd_not_zless, simp_all) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 158 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 159 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 160 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 161 | text {* \medskip @{term wset} *}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 162 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 163 | declare wset.simps [simp del] | 
| 9508 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 164 | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 165 | lemma wset_induct: | 
| 18369 | 166 |   assumes "!!a p. P {} a p"
 | 
| 19670 | 167 | and "!!a p. 1 < (a::int) \<Longrightarrow> | 
| 168 | P (wset (a - 1, p)) (a - 1) p ==> P (wset (a, p)) a p" | |
| 18369 | 169 | shows "P (wset (u, v)) u v" | 
| 170 | apply (rule wset.induct, safe) | |
| 171 | prefer 2 | |
| 172 | apply (case_tac "1 < a") | |
| 173 | apply (rule prems) | |
| 174 | apply simp_all | |
| 175 | apply (simp_all add: wset.simps prems) | |
| 176 | done | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 177 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 178 | lemma wset_mem_imp_or [rule_format]: | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11704diff
changeset | 179 | "1 < a \<Longrightarrow> b \<notin> wset (a - 1, p) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 180 | ==> b \<in> wset (a, p) --> b = a \<or> b = inv p a" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 181 | apply (subst wset.simps) | 
| 13833 | 182 | apply (unfold Let_def, simp) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 183 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 184 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11704diff
changeset | 185 | 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: 
9508diff
changeset | 186 | apply (subst wset.simps) | 
| 13833 | 187 | apply (unfold Let_def, simp) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 188 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 189 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11704diff
changeset | 190 | 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: 
9508diff
changeset | 191 | apply (subst wset.simps) | 
| 13833 | 192 | apply (unfold Let_def, auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 193 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 194 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 195 | lemma wset_g_1 [rule_format]: | 
| 16663 | 196 | "zprime p --> a < p - 1 --> b \<in> wset (a, p) --> 1 < b" | 
| 13833 | 197 | apply (induct a p rule: wset_induct, auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 198 | apply (case_tac "b = a") | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 199 | apply (case_tac [2] "b = inv p a") | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 200 | 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: 
9508diff
changeset | 201 | apply (rule_tac [4] wset_mem_imp_or) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 202 | prefer 2 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 203 | apply simp | 
| 13833 | 204 | apply (rule inv_g_1, auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 205 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 206 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 207 | lemma wset_less [rule_format]: | 
| 16663 | 208 | "zprime p --> a < p - 1 --> b \<in> wset (a, p) --> b < p - 1" | 
| 13833 | 209 | apply (induct a p rule: wset_induct, auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 210 | apply (case_tac "b = a") | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 211 | apply (case_tac [2] "b = inv p a") | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 212 | 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: 
9508diff
changeset | 213 | apply (rule_tac [4] wset_mem_imp_or) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 214 | prefer 2 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 215 | apply simp | 
| 13833 | 216 | apply (rule inv_less_p_minus_1, auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 217 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 218 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 219 | lemma wset_mem [rule_format]: | 
| 16663 | 220 | "zprime p --> | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11704diff
changeset | 221 | a < p - 1 --> 1 < b --> b \<le> a --> b \<in> wset (a, p)" | 
| 13833 | 222 | apply (induct a p rule: wset.induct, auto) | 
| 15197 | 223 | apply (rule_tac wset_subset) | 
| 224 | apply (simp (no_asm_simp)) | |
| 225 | apply auto | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 226 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 227 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 228 | lemma wset_mem_inv_mem [rule_format]: | 
| 16663 | 229 | "zprime p --> 5 \<le> p --> a < p - 1 --> b \<in> wset (a, p) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 230 | --> inv p b \<in> wset (a, p)" | 
| 13833 | 231 | apply (induct a p rule: wset_induct, auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 232 | apply (case_tac "b = a") | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 233 | apply (subst wset.simps) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 234 | apply (unfold Let_def) | 
| 13833 | 235 | apply (rule_tac [3] wset_subset, auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 236 | apply (case_tac "b = inv p a") | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 237 | apply (simp (no_asm_simp)) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 238 | apply (subst inv_inv) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 239 | apply (subgoal_tac [6] "b = a \<or> b = inv p a") | 
| 13833 | 240 | apply (rule_tac [7] wset_mem_imp_or, auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 241 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 242 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 243 | lemma wset_inv_mem_mem: | 
| 16663 | 244 | "zprime p \<Longrightarrow> 5 \<le> p \<Longrightarrow> a < p - 1 \<Longrightarrow> 1 < b \<Longrightarrow> b < p - 1 | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 245 | \<Longrightarrow> inv p b \<in> wset (a, p) \<Longrightarrow> b \<in> wset (a, p)" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 246 | 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: 
9508diff
changeset | 247 | apply (rule_tac [2] wset_mem_inv_mem) | 
| 13833 | 248 | apply (rule inv_inv, simp_all) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 249 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 250 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 251 | lemma wset_fin: "finite (wset (a, p))" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 252 | apply (induct a p rule: wset_induct) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 253 | prefer 2 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 254 | apply (subst wset.simps) | 
| 13833 | 255 | apply (unfold Let_def, auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 256 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 257 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 258 | lemma wset_zcong_prod_1 [rule_format]: | 
| 16663 | 259 | "zprime p --> | 
| 15392 | 260 | 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: 
9508diff
changeset | 261 | apply (induct a p rule: wset_induct) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 262 | prefer 2 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 263 | apply (subst wset.simps) | 
| 13833 | 264 | apply (unfold Let_def, auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 265 | apply (subst setprod_insert) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 266 |     apply (tactic {* stac (thm "setprod_insert") 3 *})
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 267 | apply (subgoal_tac [5] | 
| 15392 | 268 | "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: 
9508diff
changeset | 269 | prefer 5 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 270 | apply (simp add: zmult_assoc) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 271 | apply (rule_tac [5] zcong_zmult) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 272 | apply (rule_tac [5] inv_is_inv) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 273 | apply (tactic "Clarify_tac 4") | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11704diff
changeset | 274 | 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: 
9508diff
changeset | 275 | apply (rule_tac [5] wset_inv_mem_mem) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 276 | apply (simp_all add: wset_fin) | 
| 13833 | 277 | apply (rule inv_distinct, auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 278 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 279 | |
| 16663 | 280 | 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: 
9508diff
changeset | 281 | apply safe | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 282 | apply (erule wset_mem) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 283 | apply (rule_tac [2] d22set_g_1) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 284 | apply (rule_tac [3] d22set_le) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 285 | apply (rule_tac [4] d22set_mem) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 286 | apply (erule_tac [4] wset_g_1) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 287 | prefer 6 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 288 | 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: 
11704diff
changeset | 289 | apply (subgoal_tac "p - 2 + 1 = p - 1") | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 290 | apply (simp (no_asm_simp)) | 
| 13833 | 291 | apply (erule wset_less, auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 292 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 293 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 294 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 295 | subsection {* Wilson *}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 296 | |
| 16663 | 297 | 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: 
9508diff
changeset | 298 | apply (unfold zprime_def dvd_def) | 
| 13833 | 299 | apply (case_tac "p = 4", auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 300 | apply (rule notE) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 301 | prefer 2 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 302 | apply assumption | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 303 | apply (simp (no_asm)) | 
| 13833 | 304 | apply (rule_tac x = 2 in exI) | 
| 305 | apply (safe, arith) | |
| 306 | apply (rule_tac x = 2 in exI, auto) | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 307 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 308 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 309 | theorem Wilson_Russ: | 
| 16663 | 310 | "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: 
11704diff
changeset | 311 | 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: 
9508diff
changeset | 312 | apply (rule_tac [2] zcong_zmult) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 313 | apply (simp only: zprime_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 314 | apply (subst zfact.simps) | 
| 13833 | 315 | 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: 
9508diff
changeset | 316 | apply (simp only: zcong_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 317 | apply (simp (no_asm_simp)) | 
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 318 | apply (case_tac "p = 2") | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 319 | apply (simp add: zfact.simps) | 
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 320 | apply (case_tac "p = 3") | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 321 | apply (simp add: zfact.simps) | 
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 322 | apply (subgoal_tac "5 \<le> p") | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 323 | apply (erule_tac [2] prime_g_5) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 324 | apply (subst d22set_prod_zfact [symmetric]) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 325 | apply (subst d22set_eq_wset) | 
| 13833 | 326 | apply (rule_tac [2] wset_zcong_prod_1, auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 327 | done | 
| 9508 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 328 | |
| 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 329 | end |