| author | Manuel Eberl <eberlm@in.tum.de> | 
| Sun, 20 Aug 2017 03:35:20 +0200 | |
| changeset 66456 | 621897f47fab | 
| parent 66453 | cc19f7ca2ed6 | 
| child 66817 | 0b12755ccbb2 | 
| permissions | -rw-r--r-- | 
| 41959 | 1 | (* Title: HOL/Number_Theory/Residues.thy | 
| 31719 | 2 | Author: Jeremy Avigad | 
| 3 | ||
| 41541 | 4 | An algebraic treatment of residue rings, and resulting proofs of | 
| 41959 | 5 | Euler's theorem and Wilson's theorem. | 
| 31719 | 6 | *) | 
| 7 | ||
| 60526 | 8 | section \<open>Residue rings\<close> | 
| 31719 | 9 | |
| 10 | theory Residues | |
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 11 | imports | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 12 | Cong | 
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
66305diff
changeset | 13 | "HOL-Algebra.More_Group" | 
| 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
66305diff
changeset | 14 | "HOL-Algebra.More_Ring" | 
| 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
66305diff
changeset | 15 | "HOL-Algebra.More_Finite_Product" | 
| 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
66305diff
changeset | 16 | "HOL-Algebra.Multiplicative_Group" | 
| 65465 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 17 | Totient | 
| 31719 | 18 | begin | 
| 19 | ||
| 66305 | 20 | definition QuadRes :: "int \<Rightarrow> int \<Rightarrow> bool" | 
| 21 | where "QuadRes p a = (\<exists>y. ([y^2 = a] (mod p)))" | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: 
64272diff
changeset | 22 | |
| 66305 | 23 | definition Legendre :: "int \<Rightarrow> int \<Rightarrow> int" | 
| 24 | where "Legendre a p = | |
| 25 | (if ([a = 0] (mod p)) then 0 | |
| 26 | else if QuadRes p a then 1 | |
| 27 | else -1)" | |
| 28 | ||
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: 
64272diff
changeset | 29 | |
| 60527 | 30 | subsection \<open>A locale for residue rings\<close> | 
| 31719 | 31 | |
| 60527 | 32 | definition residue_ring :: "int \<Rightarrow> int ring" | 
| 66305 | 33 | where | 
| 34 | "residue_ring m = | |
| 35 |       \<lparr>carrier = {0..m - 1},
 | |
| 36 | monoid.mult = \<lambda>x y. (x * y) mod m, | |
| 37 | one = 1, | |
| 38 | zero = 0, | |
| 39 | add = \<lambda>x y. (x + y) mod m\<rparr>" | |
| 31719 | 40 | |
| 41 | locale residues = | |
| 42 | fixes m :: int and R (structure) | |
| 43 | assumes m_gt_one: "m > 1" | |
| 60527 | 44 | defines "R \<equiv> residue_ring m" | 
| 44872 | 45 | begin | 
| 31719 | 46 | |
| 47 | lemma abelian_group: "abelian_group R" | |
| 65066 | 48 | proof - | 
| 49 |   have "\<exists>y\<in>{0..m - 1}. (x + y) mod m = 0" if "0 \<le> x" "x < m" for x
 | |
| 50 | proof (cases "x = 0") | |
| 51 | case True | |
| 52 | with m_gt_one show ?thesis by simp | |
| 53 | next | |
| 54 | case False | |
| 55 | then have "(x + (m - x)) mod m = 0" | |
| 56 | by simp | |
| 57 | with m_gt_one that show ?thesis | |
| 58 | by (metis False atLeastAtMost_iff diff_ge_0_iff_ge diff_left_mono int_one_le_iff_zero_less less_le) | |
| 59 | qed | |
| 60 | with m_gt_one show ?thesis | |
| 61 | by (fastforce simp add: R_def residue_ring_def mod_add_right_eq ac_simps intro!: abelian_groupI) | |
| 65899 | 62 | qed | 
| 31719 | 63 | |
| 64 | lemma comm_monoid: "comm_monoid R" | |
| 65066 | 65 | unfolding R_def residue_ring_def | 
| 31719 | 66 | apply (rule comm_monoidI) | 
| 65066 | 67 | using m_gt_one apply auto | 
| 68 | apply (metis mod_mult_right_eq mult.assoc mult.commute) | |
| 69 | apply (metis mult.commute) | |
| 41541 | 70 | done | 
| 31719 | 71 | |
| 72 | lemma cring: "cring R" | |
| 65066 | 73 | apply (intro cringI abelian_group comm_monoid) | 
| 74 | unfolding R_def residue_ring_def | |
| 75 | apply (auto simp add: comm_semiring_class.distrib mod_add_eq mod_mult_left_eq) | |
| 41541 | 76 | done | 
| 31719 | 77 | |
| 78 | end | |
| 79 | ||
| 80 | sublocale residues < cring | |
| 81 | by (rule cring) | |
| 82 | ||
| 83 | ||
| 41541 | 84 | context residues | 
| 85 | begin | |
| 31719 | 86 | |
| 60527 | 87 | text \<open> | 
| 88 | These lemmas translate back and forth between internal and | |
| 89 | external concepts. | |
| 90 | \<close> | |
| 31719 | 91 | |
| 92 | lemma res_carrier_eq: "carrier R = {0..m - 1}"
 | |
| 66305 | 93 | by (auto simp: R_def residue_ring_def) | 
| 31719 | 94 | |
| 95 | lemma res_add_eq: "x \<oplus> y = (x + y) mod m" | |
| 66305 | 96 | by (auto simp: R_def residue_ring_def) | 
| 31719 | 97 | |
| 98 | lemma res_mult_eq: "x \<otimes> y = (x * y) mod m" | |
| 66305 | 99 | by (auto simp: R_def residue_ring_def) | 
| 31719 | 100 | |
| 101 | lemma res_zero_eq: "\<zero> = 0" | |
| 66305 | 102 | by (auto simp: R_def residue_ring_def) | 
| 31719 | 103 | |
| 104 | lemma res_one_eq: "\<one> = 1" | |
| 66305 | 105 | by (auto simp: R_def residue_ring_def units_of_def) | 
| 31719 | 106 | |
| 60527 | 107 | lemma res_units_eq: "Units R = {x. 0 < x \<and> x < m \<and> coprime x m}"
 | 
| 65066 | 108 | using m_gt_one | 
| 109 | unfolding Units_def R_def residue_ring_def | |
| 31719 | 110 | apply auto | 
| 66305 | 111 | apply (subgoal_tac "x \<noteq> 0") | 
| 112 | apply auto | |
| 113 | apply (metis invertible_coprime_int) | |
| 31952 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 nipkow parents: 
31798diff
changeset | 114 | apply (subst (asm) coprime_iff_invertible'_int) | 
| 66305 | 115 | apply (auto simp add: cong_int_def mult.commute) | 
| 41541 | 116 | done | 
| 31719 | 117 | |
| 118 | lemma res_neg_eq: "\<ominus> x = (- x) mod m" | |
| 65066 | 119 | using m_gt_one unfolding R_def a_inv_def m_inv_def residue_ring_def | 
| 120 | apply simp | |
| 31719 | 121 | apply (rule the_equality) | 
| 66305 | 122 | apply (simp add: mod_add_right_eq) | 
| 123 | apply (simp add: add.commute mod_add_right_eq) | |
| 65066 | 124 | apply (metis add.right_neutral minus_add_cancel mod_add_right_eq mod_pos_pos_trivial) | 
| 41541 | 125 | done | 
| 31719 | 126 | |
| 44872 | 127 | lemma finite [iff]: "finite (carrier R)" | 
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 128 | by (simp add: res_carrier_eq) | 
| 31719 | 129 | |
| 44872 | 130 | lemma finite_Units [iff]: "finite (Units R)" | 
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 131 | by (simp add: finite_ring_finite_units) | 
| 31719 | 132 | |
| 60527 | 133 | text \<open> | 
| 63167 | 134 | The function \<open>a \<mapsto> a mod m\<close> maps the integers to the | 
| 60527 | 135 | residue classes. The following lemmas show that this mapping | 
| 136 | respects addition and multiplication on the integers. | |
| 137 | \<close> | |
| 31719 | 138 | |
| 60527 | 139 | lemma mod_in_carrier [iff]: "a mod m \<in> carrier R" | 
| 140 | unfolding res_carrier_eq | |
| 141 | using insert m_gt_one by auto | |
| 31719 | 142 | |
| 143 | lemma add_cong: "(x mod m) \<oplus> (y mod m) = (x + y) mod m" | |
| 66305 | 144 | by (auto simp: R_def residue_ring_def mod_simps) | 
| 31719 | 145 | |
| 146 | lemma mult_cong: "(x mod m) \<otimes> (y mod m) = (x * y) mod m" | |
| 66305 | 147 | by (auto simp: R_def residue_ring_def mod_simps) | 
| 31719 | 148 | |
| 149 | lemma zero_cong: "\<zero> = 0" | |
| 66305 | 150 | by (auto simp: R_def residue_ring_def) | 
| 31719 | 151 | |
| 152 | lemma one_cong: "\<one> = 1 mod m" | |
| 66305 | 153 | using m_gt_one by (auto simp: R_def residue_ring_def) | 
| 31719 | 154 | |
| 60527 | 155 | (* FIXME revise algebra library to use 1? *) | 
| 31719 | 156 | lemma pow_cong: "(x mod m) (^) n = x^n mod m" | 
| 65066 | 157 | using m_gt_one | 
| 31719 | 158 | apply (induct n) | 
| 41541 | 159 | apply (auto simp add: nat_pow_def one_cong) | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
55352diff
changeset | 160 | apply (metis mult.commute mult_cong) | 
| 41541 | 161 | done | 
| 31719 | 162 | |
| 163 | lemma neg_cong: "\<ominus> (x mod m) = (- x) mod m" | |
| 55352 | 164 | by (metis mod_minus_eq res_neg_eq) | 
| 31719 | 165 | |
| 60528 | 166 | lemma (in residues) prod_cong: "finite A \<Longrightarrow> (\<Otimes>i\<in>A. (f i) mod m) = (\<Prod>i\<in>A. f i) mod m" | 
| 55352 | 167 | by (induct set: finite) (auto simp: one_cong mult_cong) | 
| 31719 | 168 | |
| 60528 | 169 | lemma (in residues) sum_cong: "finite A \<Longrightarrow> (\<Oplus>i\<in>A. (f i) mod m) = (\<Sum>i\<in>A. f i) mod m" | 
| 55352 | 170 | by (induct set: finite) (auto simp: zero_cong add_cong) | 
| 31719 | 171 | |
| 60688 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60528diff
changeset | 172 | lemma mod_in_res_units [simp]: | 
| 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60528diff
changeset | 173 | assumes "1 < m" and "coprime a m" | 
| 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60528diff
changeset | 174 | shows "a mod m \<in> Units R" | 
| 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60528diff
changeset | 175 | proof (cases "a mod m = 0") | 
| 66305 | 176 | case True | 
| 177 | with assms show ?thesis | |
| 60688 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60528diff
changeset | 178 | by (auto simp add: res_units_eq gcd_red_int [symmetric]) | 
| 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60528diff
changeset | 179 | next | 
| 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60528diff
changeset | 180 | case False | 
| 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60528diff
changeset | 181 | from assms have "0 < m" by simp | 
| 66305 | 182 | then have "0 \<le> a mod m" by (rule pos_mod_sign [of m a]) | 
| 60688 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60528diff
changeset | 183 | with False have "0 < a mod m" by simp | 
| 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60528diff
changeset | 184 | with assms show ?thesis | 
| 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60528diff
changeset | 185 | by (auto simp add: res_units_eq gcd_red_int [symmetric] ac_simps) | 
| 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60528diff
changeset | 186 | qed | 
| 31719 | 187 | |
| 60528 | 188 | lemma res_eq_to_cong: "(a mod m) = (b mod m) \<longleftrightarrow> [a = b] (mod m)" | 
| 66305 | 189 | by (auto simp: cong_int_def) | 
| 31719 | 190 | |
| 191 | ||
| 60528 | 192 | text \<open>Simplifying with these will translate a ring equation in R to a congruence.\<close> | 
| 66305 | 193 | lemmas res_to_cong_simps = | 
| 194 | add_cong mult_cong pow_cong one_cong | |
| 195 | prod_cong sum_cong neg_cong res_eq_to_cong | |
| 31719 | 196 | |
| 60527 | 197 | text \<open>Other useful facts about the residue ring.\<close> | 
| 31719 | 198 | lemma one_eq_neg_one: "\<one> = \<ominus> \<one> \<Longrightarrow> m = 2" | 
| 199 | apply (simp add: res_one_eq res_neg_eq) | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
55352diff
changeset | 200 | apply (metis add.commute add_diff_cancel mod_mod_trivial one_add_one uminus_add_conv_diff | 
| 60528 | 201 | zero_neq_one zmod_zminus1_eq_if) | 
| 41541 | 202 | done | 
| 31719 | 203 | |
| 204 | end | |
| 205 | ||
| 206 | ||
| 60527 | 207 | subsection \<open>Prime residues\<close> | 
| 31719 | 208 | |
| 209 | locale residues_prime = | |
| 63534 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 eberlm <eberlm@in.tum.de> parents: 
63417diff
changeset | 210 | fixes p :: nat and R (structure) | 
| 31719 | 211 | assumes p_prime [intro]: "prime p" | 
| 63534 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 eberlm <eberlm@in.tum.de> parents: 
63417diff
changeset | 212 | defines "R \<equiv> residue_ring (int p)" | 
| 31719 | 213 | |
| 214 | sublocale residues_prime < residues p | |
| 65066 | 215 | unfolding R_def residues_def | 
| 31719 | 216 | using p_prime apply auto | 
| 62348 | 217 | apply (metis (full_types) of_nat_1 of_nat_less_iff prime_gt_1_nat) | 
| 41541 | 218 | done | 
| 31719 | 219 | |
| 44872 | 220 | context residues_prime | 
| 221 | begin | |
| 31719 | 222 | |
| 223 | lemma is_field: "field R" | |
| 65066 | 224 | proof - | 
| 66305 | 225 | have "gcd x (int p) \<noteq> 1 \<Longrightarrow> 0 \<le> x \<Longrightarrow> x < int p \<Longrightarrow> x = 0" for x | 
| 65066 | 226 | by (metis dual_order.order_iff_strict gcd.commute less_le_not_le p_prime prime_imp_coprime prime_nat_int_transfer zdvd_imp_le) | 
| 227 | then show ?thesis | |
| 66305 | 228 | apply (intro cring.field_intro2 cring) | 
| 229 | apply (auto simp add: res_carrier_eq res_one_eq res_zero_eq res_units_eq) | |
| 65066 | 230 | done | 
| 231 | qed | |
| 31719 | 232 | |
| 233 | lemma res_prime_units_eq: "Units R = {1..p - 1}"
 | |
| 234 | apply (subst res_units_eq) | |
| 235 | apply auto | |
| 62348 | 236 | apply (subst gcd.commute) | 
| 55352 | 237 | apply (auto simp add: p_prime prime_imp_coprime_int zdvd_not_zless) | 
| 41541 | 238 | done | 
| 31719 | 239 | |
| 240 | end | |
| 241 | ||
| 242 | sublocale residues_prime < field | |
| 243 | by (rule is_field) | |
| 244 | ||
| 245 | ||
| 60527 | 246 | section \<open>Test cases: Euler's theorem and Wilson's theorem\<close> | 
| 31719 | 247 | |
| 60527 | 248 | subsection \<open>Euler's theorem\<close> | 
| 31719 | 249 | |
| 66305 | 250 | lemma (in residues) totient_eq: "totient (nat m) = card (Units R)" | 
| 55261 
ad3604df6bc6
new lemmas involving phi from Lehmer AFP entry
 paulson <lp15@cam.ac.uk> parents: 
55242diff
changeset | 251 | proof - | 
| 65465 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 252 | have *: "inj_on nat (Units R)" | 
| 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 253 | by (rule inj_onI) (auto simp add: res_units_eq) | 
| 65726 
f5d64d094efe
More material on totient function
 eberlm <eberlm@in.tum.de> parents: 
65465diff
changeset | 254 | define m' where "m' = nat m" | 
| 66305 | 255 | from m_gt_one have "m = int m'" "m' > 1" | 
| 256 | by (simp_all add: m'_def) | |
| 257 | then have "x \<in> Units R \<longleftrightarrow> x \<in> int ` totatives m'" for x | |
| 65726 
f5d64d094efe
More material on totient function
 eberlm <eberlm@in.tum.de> parents: 
65465diff
changeset | 258 | unfolding res_units_eq | 
| 
f5d64d094efe
More material on totient function
 eberlm <eberlm@in.tum.de> parents: 
65465diff
changeset | 259 | by (cases x; cases "x = m") (auto simp: totatives_def transfer_int_nat_gcd) | 
| 66305 | 260 | then have "Units R = int ` totatives m'" | 
| 261 | by blast | |
| 262 | then have "totatives m' = nat ` Units R" | |
| 263 | by (simp add: image_image) | |
| 65465 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 264 | then have "card (totatives (nat m)) = card (nat ` Units R)" | 
| 65726 
f5d64d094efe
More material on totient function
 eberlm <eberlm@in.tum.de> parents: 
65465diff
changeset | 265 | by (simp add: m'_def) | 
| 65465 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 266 | also have "\<dots> = card (Units R)" | 
| 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 267 | using * card_image [of nat "Units R"] by auto | 
| 66305 | 268 | finally show ?thesis | 
| 269 | by (simp add: totient_def) | |
| 55261 
ad3604df6bc6
new lemmas involving phi from Lehmer AFP entry
 paulson <lp15@cam.ac.uk> parents: 
55242diff
changeset | 270 | qed | 
| 
ad3604df6bc6
new lemmas involving phi from Lehmer AFP entry
 paulson <lp15@cam.ac.uk> parents: 
55242diff
changeset | 271 | |
| 65465 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 272 | lemma (in residues_prime) totient_eq: "totient p = p - 1" | 
| 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 273 | using totient_eq by (simp add: res_prime_units_eq) | 
| 31719 | 274 | |
| 65465 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 275 | lemma (in residues) euler_theorem: | 
| 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 276 | assumes "coprime a m" | 
| 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 277 | shows "[a ^ totient (nat m) = 1] (mod m)" | 
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 278 | proof - | 
| 65465 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 279 | have "a ^ totient (nat m) mod m = 1 mod m" | 
| 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 280 | by (metis assms finite_Units m_gt_one mod_in_res_units one_cong totient_eq pow_cong units_power_order_eq_one) | 
| 65066 | 281 | then show ?thesis | 
| 282 | using res_eq_to_cong by blast | |
| 31719 | 283 | qed | 
| 284 | ||
| 285 | lemma euler_theorem: | |
| 65465 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 286 | fixes a m :: nat | 
| 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 287 | assumes "coprime a m" | 
| 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 288 | shows "[a ^ totient m = 1] (mod m)" | 
| 60527 | 289 | proof (cases "m = 0 | m = 1") | 
| 290 | case True | |
| 44872 | 291 | then show ?thesis by auto | 
| 31719 | 292 | next | 
| 60527 | 293 | case False | 
| 41541 | 294 | with assms show ?thesis | 
| 65465 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 295 | using residues.euler_theorem [of "int m" "int a"] transfer_int_nat_cong | 
| 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 296 | by (auto simp add: residues_def transfer_int_nat_gcd(1)) force | 
| 31719 | 297 | qed | 
| 298 | ||
| 299 | lemma fermat_theorem: | |
| 65465 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 300 | fixes p a :: nat | 
| 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 301 | assumes "prime p" and "\<not> p dvd a" | 
| 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 302 | shows "[a ^ (p - 1) = 1] (mod p)" | 
| 31719 | 303 | proof - | 
| 65465 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 304 | from assms prime_imp_coprime [of p a] have "coprime a p" | 
| 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 305 | by (auto simp add: ac_simps) | 
| 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 306 | then have "[a ^ totient p = 1] (mod p)" | 
| 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 307 | by (rule euler_theorem) | 
| 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 308 | also have "totient p = p - 1" | 
| 65726 
f5d64d094efe
More material on totient function
 eberlm <eberlm@in.tum.de> parents: 
65465diff
changeset | 309 | by (rule totient_prime) (rule assms) | 
| 65465 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 310 | finally show ?thesis . | 
| 31719 | 311 | qed | 
| 312 | ||
| 313 | ||
| 60526 | 314 | subsection \<open>Wilson's theorem\<close> | 
| 31719 | 315 | |
| 60527 | 316 | lemma (in field) inv_pair_lemma: "x \<in> Units R \<Longrightarrow> y \<in> Units R \<Longrightarrow> | 
| 317 |     {x, inv x} \<noteq> {y, inv y} \<Longrightarrow> {x, inv x} \<inter> {y, inv y} = {}"
 | |
| 31719 | 318 | apply auto | 
| 55352 | 319 | apply (metis Units_inv_inv)+ | 
| 41541 | 320 | done | 
| 31719 | 321 | |
| 322 | lemma (in residues_prime) wilson_theorem1: | |
| 323 | assumes a: "p > 2" | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59667diff
changeset | 324 | shows "[fact (p - 1) = (-1::int)] (mod p)" | 
| 31719 | 325 | proof - | 
| 60527 | 326 |   let ?Inverse_Pairs = "{{x, inv x}| x. x \<in> Units R - {\<one>, \<ominus> \<one>}}"
 | 
| 327 |   have UR: "Units R = {\<one>, \<ominus> \<one>} \<union> \<Union>?Inverse_Pairs"
 | |
| 31719 | 328 | by auto | 
| 60527 | 329 |   have "(\<Otimes>i\<in>Units R. i) = (\<Otimes>i\<in>{\<one>, \<ominus> \<one>}. i) \<otimes> (\<Otimes>i\<in>\<Union>?Inverse_Pairs. i)"
 | 
| 31732 | 330 | apply (subst UR) | 
| 31719 | 331 | apply (subst finprod_Un_disjoint) | 
| 66305 | 332 | apply (auto intro: funcsetI) | 
| 60527 | 333 | using inv_one apply auto[1] | 
| 334 | using inv_eq_neg_one_eq apply auto | |
| 31719 | 335 | done | 
| 60527 | 336 |   also have "(\<Otimes>i\<in>{\<one>, \<ominus> \<one>}. i) = \<ominus> \<one>"
 | 
| 31719 | 337 | apply (subst finprod_insert) | 
| 66305 | 338 | apply auto | 
| 31719 | 339 | apply (frule one_eq_neg_one) | 
| 60527 | 340 | using a apply force | 
| 31719 | 341 | done | 
| 60527 | 342 | also have "(\<Otimes>i\<in>(\<Union>?Inverse_Pairs). i) = (\<Otimes>A\<in>?Inverse_Pairs. (\<Otimes>y\<in>A. y))" | 
| 343 | apply (subst finprod_Union_disjoint) | |
| 66305 | 344 | apply auto | 
| 345 | apply (metis Units_inv_inv)+ | |
| 31719 | 346 | done | 
| 347 | also have "\<dots> = \<one>" | |
| 60527 | 348 | apply (rule finprod_one) | 
| 66305 | 349 | apply auto | 
| 60527 | 350 | apply (subst finprod_insert) | 
| 66305 | 351 | apply auto | 
| 55352 | 352 | apply (metis inv_eq_self) | 
| 31719 | 353 | done | 
| 60527 | 354 | finally have "(\<Otimes>i\<in>Units R. i) = \<ominus> \<one>" | 
| 31719 | 355 | by simp | 
| 60527 | 356 | also have "(\<Otimes>i\<in>Units R. i) = (\<Otimes>i\<in>Units R. i mod p)" | 
| 65066 | 357 | by (rule finprod_cong') (auto simp: res_units_eq) | 
| 60527 | 358 | also have "\<dots> = (\<Prod>i\<in>Units R. i) mod p" | 
| 65066 | 359 | by (rule prod_cong) auto | 
| 31719 | 360 | also have "\<dots> = fact (p - 1) mod p" | 
| 64272 | 361 | apply (simp add: fact_prod) | 
| 65066 | 362 | using assms | 
| 55242 
413ec965f95d
Number_Theory: prime is no longer overloaded, but only for nat. Automatic coercion to int enabled.
 paulson <lp15@cam.ac.uk> parents: 
55227diff
changeset | 363 | apply (subst res_prime_units_eq) | 
| 64272 | 364 | apply (simp add: int_prod zmod_int prod_int_eq) | 
| 31719 | 365 | done | 
| 60527 | 366 | finally have "fact (p - 1) mod p = \<ominus> \<one>" . | 
| 367 | then show ?thesis | |
| 60528 | 368 | by (metis of_nat_fact Divides.transfer_int_nat_functions(2) | 
| 66305 | 369 | cong_int_def res_neg_eq res_one_eq) | 
| 31719 | 370 | qed | 
| 371 | ||
| 55352 | 372 | lemma wilson_theorem: | 
| 60527 | 373 | assumes "prime p" | 
| 374 | shows "[fact (p - 1) = - 1] (mod p)" | |
| 55352 | 375 | proof (cases "p = 2") | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 376 | case True | 
| 55352 | 377 | then show ?thesis | 
| 64272 | 378 | by (simp add: cong_int_def fact_prod) | 
| 55352 | 379 | next | 
| 380 | case False | |
| 381 | then show ?thesis | |
| 382 | using assms prime_ge_2_nat | |
| 383 | by (metis residues_prime.wilson_theorem1 residues_prime.intro le_eq_less_or_eq) | |
| 384 | qed | |
| 31719 | 385 | |
| 66304 | 386 | text \<open> | 
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 387 | This result can be transferred to the multiplicative group of | 
| 66305 | 388 | \<open>\<int>/p\<int>\<close> for \<open>p\<close> prime.\<close> | 
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 389 | |
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 390 | lemma mod_nat_int_pow_eq: | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 391 | fixes n :: nat and p a :: int | 
| 66305 | 392 | shows "a \<ge> 0 \<Longrightarrow> p \<ge> 0 \<Longrightarrow> (nat a ^ n) mod (nat p) = nat ((a ^ n) mod p)" | 
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 393 | by (simp add: int_one_le_iff_zero_less nat_mod_distrib order_less_imp_le nat_power_eq[symmetric]) | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 394 | |
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 395 | theorem residue_prime_mult_group_has_gen : | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 396 | fixes p :: nat | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 397 | assumes prime_p : "prime p" | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 398 |  shows "\<exists>a \<in> {1 .. p - 1}. {1 .. p - 1} = {a^i mod p|i . i \<in> UNIV}"
 | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 399 | proof - | 
| 66305 | 400 | have "p \<ge> 2" | 
| 401 | using prime_gt_1_nat[OF prime_p] by simp | |
| 402 | interpret R: residues_prime p "residue_ring p" | |
| 403 | by (simp add: residues_prime_def prime_p) | |
| 404 |   have car: "carrier (residue_ring (int p)) - {\<zero>\<^bsub>residue_ring (int p)\<^esub>} = {1 .. int p - 1}"
 | |
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 405 | by (auto simp add: R.zero_cong R.res_carrier_eq) | 
| 66305 | 406 | |
| 407 | have "x (^)\<^bsub>residue_ring (int p)\<^esub> i = x ^ i mod (int p)" | |
| 408 |     if "x \<in> {1 .. int p - 1}" for x and i :: nat
 | |
| 409 | using that R.pow_cong[of x i] by auto | |
| 410 | moreover | |
| 411 |   obtain a where a: "a \<in> {1 .. int p - 1}"
 | |
| 412 |     and a_gen: "{1 .. int p - 1} = {a(^)\<^bsub>residue_ring (int p)\<^esub>i|i::nat . i \<in> UNIV}"
 | |
| 413 | using field.finite_field_mult_group_has_gen[OF R.is_field] | |
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 414 | by (auto simp add: car[symmetric] carrier_mult_of) | 
| 66305 | 415 | moreover | 
| 416 |   have "nat ` {1 .. int p - 1} = {1 .. p - 1}" (is "?L = ?R")
 | |
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 417 | proof | 
| 66305 | 418 | have "n \<in> ?R" if "n \<in> ?L" for n | 
| 419 | using that \<open>p\<ge>2\<close> by force | |
| 420 | then show "?L \<subseteq> ?R" by blast | |
| 421 | have "n \<in> ?L" if "n \<in> ?R" for n | |
| 422 | using that \<open>p\<ge>2\<close> Set_Interval.transfer_nat_int_set_functions(2) by fastforce | |
| 423 | then show "?R \<subseteq> ?L" by blast | |
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 424 | qed | 
| 66305 | 425 | moreover | 
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 426 |   have "nat ` {a^i mod (int p) | i::nat. i \<in> UNIV} = {nat a^i mod p | i . i \<in> UNIV}" (is "?L = ?R")
 | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 427 | proof | 
| 66305 | 428 | have "x \<in> ?R" if "x \<in> ?L" for x | 
| 429 | proof - | |
| 430 | from that obtain i where i: "x = nat (a^i mod (int p))" | |
| 431 | by blast | |
| 432 | then have "x = nat a ^ i mod p" | |
| 433 | using mod_nat_int_pow_eq[of a "int p" i] a \<open>p\<ge>2\<close> by auto | |
| 434 | with i show ?thesis by blast | |
| 435 | qed | |
| 436 | then show "?L \<subseteq> ?R" by blast | |
| 437 | have "x \<in> ?L" if "x \<in> ?R" for x | |
| 438 | proof - | |
| 439 | from that obtain i where i: "x = nat a^i mod p" | |
| 440 | by blast | |
| 441 | with mod_nat_int_pow_eq[of a "int p" i] a \<open>p\<ge>2\<close> show ?thesis | |
| 442 | by auto | |
| 443 | qed | |
| 444 | then show "?R \<subseteq> ?L" by blast | |
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 445 | qed | 
| 66305 | 446 |   ultimately have "{1 .. p - 1} = {nat a^i mod p | i. i \<in> UNIV}"
 | 
| 447 | by presburger | |
| 448 |   moreover from a have "nat a \<in> {1 .. p - 1}" by force
 | |
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 449 | ultimately show ?thesis .. | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 450 | qed | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 451 | |
| 31719 | 452 | end |