| author | wenzelm | 
| Thu, 31 May 2001 22:34:58 +0200 | |
| changeset 11357 | 908b761cdfb0 | 
| parent 11049 | 7eef34adb852 | 
| child 11701 | 3d51fbf81c17 | 
| 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/WilsonBij.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 using a more abstract approach *}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 8 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 9 | theory WilsonBij = BijectionRel + IntFact: | 
| 
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 using a more ``abstract'' approach based on | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 13 | bijections between sets. Does not use Fermat's Little Theorem | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 14 | (unlike Russinoff). | 
| 
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 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 17 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 18 | subsection {* Definitions and lemmas *}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 19 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 20 | constdefs | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 21 | reciR :: "int => int => int => bool" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 22 | "reciR p == | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 23 | \<lambda>a b. zcong (a * b) #1 p \<and> #1 < a \<and> a < p - #1 \<and> #1 < b \<and> b < p - #1" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 24 | inv :: "int => int => int" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 25 | "inv p a == | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 26 | if p \<in> zprime \<and> #0 < a \<and> a < p then | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 27 | (SOME x. #0 \<le> x \<and> x < p \<and> zcong (a * x) #1 p) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 28 | else #0" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 29 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 30 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 31 | text {* \medskip Inverse *}
 | 
| 
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 | lemma inv_correct: | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 34 | "p \<in> zprime ==> #0 < a ==> a < p | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 35 | ==> #0 \<le> inv p a \<and> inv p a < p \<and> [a * inv p a = #1] (mod p)" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 36 | apply (unfold inv_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 37 | apply (simp (no_asm_simp)) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 38 | apply (rule zcong_lineq_unique [THEN ex1_implies_ex, THEN someI_ex]) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 39 | apply (erule_tac [2] zless_zprime_imp_zrelprime) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 40 | apply (unfold zprime_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 41 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 42 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 43 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 44 | lemmas inv_ge = inv_correct [THEN conjunct1, standard] | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 45 | lemmas inv_less = inv_correct [THEN conjunct2, THEN conjunct1, standard] | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 46 | lemmas inv_is_inv = inv_correct [THEN conjunct2, THEN conjunct2, standard] | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 47 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 48 | lemma inv_not_0: | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 49 | "p \<in> zprime ==> #1 < a ==> a < p - #1 ==> inv p a \<noteq> #0" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 50 |   -- {* same as @{text WilsonRuss} *}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 51 | apply safe | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 52 | 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 | 53 | apply (unfold zcong_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 54 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 55 | apply (subgoal_tac "\<not> p dvd #1") | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 56 | apply (rule_tac [2] zdvd_not_zless) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 57 | apply (subgoal_tac "p dvd #1") | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 58 | prefer 2 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 59 | apply (subst zdvd_zminus_iff [symmetric]) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 60 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 61 | done | 
| 9508 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 62 | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 63 | lemma inv_not_1: | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 64 | "p \<in> zprime ==> #1 < a ==> a < p - #1 ==> inv p a \<noteq> #1" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 65 |   -- {* same as @{text WilsonRuss} *}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 66 | apply safe | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 67 | 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 | 68 | prefer 4 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 69 | apply simp | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 70 | apply (subgoal_tac "a = #1") | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 71 | apply (rule_tac [2] zcong_zless_imp_eq) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 72 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 73 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 74 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 75 | lemma aux: "[a * (p - #1) = #1] (mod p) = [a = p - #1] (mod p)" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 76 |   -- {* same as @{text WilsonRuss} *}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 77 | apply (unfold zcong_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 78 | apply (simp add: zdiff_zdiff_eq zdiff_zdiff_eq2 zdiff_zmult_distrib2) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 79 | apply (rule_tac s = "p dvd -((a + #1) + (p * -a))" in trans) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 80 | apply (simp add: zmult_commute zminus_zdiff_eq) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 81 | apply (subst zdvd_zminus_iff) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 82 | apply (subst zdvd_reduce) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 83 | apply (rule_tac s = "p dvd (a + #1) + (p * -#1)" in trans) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 84 | apply (subst zdvd_reduce) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 85 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 86 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 87 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 88 | lemma inv_not_p_minus_1: | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 89 | "p \<in> zprime ==> #1 < a ==> a < p - #1 ==> inv p a \<noteq> p - #1" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 90 |   -- {* same as @{text WilsonRuss} *}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 91 | apply safe | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 92 | 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 | 93 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 94 | apply (simp add: aux) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 95 | apply (subgoal_tac "a = p - #1") | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 96 | apply (rule_tac [2] zcong_zless_imp_eq) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 97 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 98 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 99 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 100 | text {*
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 101 |   Below is slightly different as we don't expand @{term [source] inv}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 102 |   but use ``@{text correct}'' theorems.
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 103 | *} | 
| 
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: "p \<in> zprime ==> #1 < a ==> a < p - #1 ==> #1 < inv p a" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 106 | apply (subgoal_tac "inv p a \<noteq> #1") | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 107 | apply (subgoal_tac "inv p a \<noteq> #0") | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 108 | apply (subst order_less_le) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 109 | apply (subst zle_add1_eq_le [symmetric]) | 
| 
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 (rule_tac [2] inv_not_0) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 112 | apply (rule_tac [5] inv_not_1) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 113 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 114 | apply (rule inv_ge) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 115 | apply auto | 
| 
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: | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 119 | "p \<in> zprime ==> #1 < a ==> a < p - #1 ==> inv p a < p - #1" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 120 |   -- {* ditto *}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 121 | apply (subst order_less_le) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 122 | apply (simp add: inv_not_p_minus_1 inv_less) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 123 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 124 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 125 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 126 | text {* \medskip Bijection *}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 127 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 128 | lemma aux1: "#1 < x ==> #0 \<le> (x::int)" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 129 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 130 | done | 
| 9508 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 131 | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 132 | lemma aux2: "#1 < x ==> #0 < (x::int)" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 133 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 134 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 135 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 136 | lemma aux3: "x \<le> p - #2 ==> x < (p::int)" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 137 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 138 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 139 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 140 | lemma aux4: "x \<le> p - #2 ==> x < (p::int)-#1" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 141 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 142 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 143 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 144 | lemma inv_inj: "p \<in> zprime ==> inj_on (inv p) (d22set (p - #2))" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 145 | apply (unfold inj_on_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 146 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 147 | apply (rule zcong_zless_imp_eq) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 148 |       apply (tactic {* stac (thm "zcong_cancel" RS sym) 5 *})
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 149 | apply (rule_tac [7] zcong_trans) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 150 |          apply (tactic {* stac (thm "zcong_sym") 8 *})
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 151 | apply (erule_tac [7] inv_is_inv) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 152 | apply (tactic "Asm_simp_tac 9") | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 153 | apply (erule_tac [9] inv_is_inv) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 154 | apply (rule_tac [6] zless_zprime_imp_zrelprime) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 155 | apply (rule_tac [8] inv_less) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 156 | apply (rule_tac [7] inv_g_1 [THEN aux2]) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 157 | apply (unfold zprime_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 158 | apply (auto intro: d22set_g_1 d22set_le | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 159 | aux1 aux2 aux3 aux4) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 160 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 161 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 162 | lemma inv_d22set_d22set: | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 163 | "p \<in> zprime ==> inv p ` d22set (p - #2) = d22set (p - #2)" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 164 | apply (rule endo_inj_surj) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 165 | apply (rule d22set_fin) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 166 | apply (erule_tac [2] inv_inj) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 167 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 168 | apply (rule d22set_mem) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 169 | apply (erule inv_g_1) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 170 | apply (subgoal_tac [3] "inv p xa < p - #1") | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 171 | apply (erule_tac [4] inv_less_p_minus_1) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 172 | apply (auto intro: d22set_g_1 d22set_le aux4) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 173 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 174 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 175 | lemma d22set_d22set_bij: | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 176 | "p \<in> zprime ==> (d22set (p - #2), d22set (p - #2)) \<in> bijR (reciR p)" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 177 | apply (unfold reciR_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 178 | apply (rule_tac s = "(d22set (p - #2), inv p ` d22set (p - #2))" in subst) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 179 | apply (simp add: inv_d22set_d22set) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 180 | apply (rule inj_func_bijR) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 181 | apply (rule_tac [3] d22set_fin) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 182 | apply (erule_tac [2] inv_inj) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 183 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 184 | apply (erule inv_is_inv) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 185 | apply (erule_tac [5] inv_g_1) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 186 | apply (erule_tac [7] inv_less_p_minus_1) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 187 | apply (auto intro: d22set_g_1 d22set_le aux2 aux3 aux4) | 
| 
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 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 190 | lemma reciP_bijP: "p \<in> zprime ==> bijP (reciR p) (d22set (p - #2))" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 191 | apply (unfold reciR_def bijP_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 192 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 193 | apply (rule d22set_mem) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 194 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 195 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 196 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 197 | lemma reciP_uniq: "p \<in> zprime ==> uniqP (reciR p)" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 198 | apply (unfold reciR_def uniqP_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 199 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 200 | apply (rule zcong_zless_imp_eq) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 201 |        apply (tactic {* stac (thm "zcong_cancel2" RS sym) 5 *})
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 202 | apply (rule_tac [7] zcong_trans) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 203 |           apply (tactic {* stac (thm "zcong_sym") 8 *})
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 204 | apply (rule_tac [6] zless_zprime_imp_zrelprime) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 205 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 206 | apply (rule zcong_zless_imp_eq) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 207 |       apply (tactic {* stac (thm "zcong_cancel" RS sym) 5 *})
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 208 | apply (rule_tac [7] zcong_trans) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 209 |          apply (tactic {* stac (thm "zcong_sym") 8 *})
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 210 | apply (rule_tac [6] zless_zprime_imp_zrelprime) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 211 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 212 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 213 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 214 | lemma reciP_sym: "p \<in> zprime ==> symP (reciR p)" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 215 | apply (unfold reciR_def symP_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 216 | apply (simp add: zmult_commute) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 217 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 218 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 219 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 220 | lemma bijER_d22set: "p \<in> zprime ==> d22set (p - #2) \<in> bijER (reciR p)" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 221 | apply (rule bijR_bijER) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 222 | apply (erule d22set_d22set_bij) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 223 | apply (erule reciP_bijP) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 224 | apply (erule reciP_uniq) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 225 | apply (erule reciP_sym) | 
| 
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 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 229 | subsection {* Wilson *}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 230 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 231 | lemma bijER_zcong_prod_1: | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 232 | "p \<in> zprime ==> A \<in> bijER (reciR p) ==> [setprod A = #1] (mod p)" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 233 | apply (unfold reciR_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 234 | apply (erule bijER.induct) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 235 | apply (subgoal_tac [2] "a = #1 \<or> a = p - #1") | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 236 | apply (rule_tac [3] zcong_square_zless) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 237 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 238 | apply (subst setprod_insert) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 239 | prefer 3 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 240 | apply (subst setprod_insert) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 241 | apply (auto simp add: fin_bijER) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 242 | apply (subgoal_tac "zcong ((a * b) * setprod A) (#1 * #1) p") | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 243 | apply (simp add: zmult_assoc) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 244 | apply (rule zcong_zmult) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 245 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 246 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 247 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 248 | theorem Wilson_Bij: "p \<in> zprime ==> [zfact (p - #1) = #-1] (mod p)" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 249 | apply (subgoal_tac "zcong ((p - #1) * zfact (p - #2)) (#-1 * #1) p") | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 250 | apply (rule_tac [2] zcong_zmult) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 251 | apply (simp add: zprime_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 252 | apply (subst zfact.simps) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 253 | apply (rule_tac t = "p - #1 - #1" and s = "p - #2" in subst) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 254 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 255 | apply (simp add: zcong_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 256 | apply (subst d22set_prod_zfact [symmetric]) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 257 | apply (rule bijER_zcong_prod_1) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 258 | apply (rule_tac [2] bijER_d22set) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 259 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 260 | done | 
| 9508 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 261 | |
| 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 262 | end |