| author | haftmann | 
| Sat, 23 Mar 2013 20:50:39 +0100 | |
| changeset 51489 | f738e6dbd844 | 
| parent 50027 | 7747a9f4c358 | 
| child 54489 | 03ff4d1e6784 | 
| 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 | ||
| 8 | header {* Residue rings *}
 | |
| 9 | ||
| 10 | theory Residues | |
| 11 | imports | |
| 41541 | 12 | UniqueFactorization | 
| 13 | Binomial | |
| 14 | MiscAlgebra | |
| 31719 | 15 | begin | 
| 16 | ||
| 17 | ||
| 18 | (* | |
| 44872 | 19 | |
| 31719 | 20 | A locale for residue rings | 
| 21 | ||
| 22 | *) | |
| 23 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32479diff
changeset | 24 | definition residue_ring :: "int => int ring" where | 
| 44872 | 25 | "residue_ring m == (| | 
| 26 |     carrier =       {0..m - 1},
 | |
| 31719 | 27 | mult = (%x y. (x * y) mod m), | 
| 28 | one = 1, | |
| 29 | zero = 0, | |
| 30 | add = (%x y. (x + y) mod m) |)" | |
| 31 | ||
| 32 | locale residues = | |
| 33 | fixes m :: int and R (structure) | |
| 34 | assumes m_gt_one: "m > 1" | |
| 35 | defines "R == residue_ring m" | |
| 36 | ||
| 44872 | 37 | context residues | 
| 38 | begin | |
| 31719 | 39 | |
| 40 | lemma abelian_group: "abelian_group R" | |
| 41 | apply (insert m_gt_one) | |
| 42 | apply (rule abelian_groupI) | |
| 43 | apply (unfold R_def residue_ring_def) | |
| 41541 | 44 | apply (auto simp add: mod_add_right_eq [symmetric] add_ac) | 
| 31719 | 45 | apply (case_tac "x = 0") | 
| 46 | apply force | |
| 47 | apply (subgoal_tac "(x + (m - x)) mod m = 0") | |
| 48 | apply (erule bexI) | |
| 49 | apply auto | |
| 41541 | 50 | done | 
| 31719 | 51 | |
| 52 | lemma comm_monoid: "comm_monoid R" | |
| 53 | apply (insert m_gt_one) | |
| 54 | apply (unfold R_def residue_ring_def) | |
| 55 | apply (rule comm_monoidI) | |
| 56 | apply auto | |
| 57 | apply (subgoal_tac "x * y mod m * z mod m = z * (x * y mod m) mod m") | |
| 58 | apply (erule ssubst) | |
| 47163 | 59 | apply (subst mod_mult_right_eq [symmetric])+ | 
| 31719 | 60 | apply (simp_all only: mult_ac) | 
| 41541 | 61 | done | 
| 31719 | 62 | |
| 63 | lemma cring: "cring R" | |
| 64 | apply (rule cringI) | |
| 65 | apply (rule abelian_group) | |
| 66 | apply (rule comm_monoid) | |
| 67 | apply (unfold R_def residue_ring_def, auto) | |
| 68 | apply (subst mod_add_eq [symmetric]) | |
| 69 | apply (subst mult_commute) | |
| 47163 | 70 | apply (subst mod_mult_right_eq [symmetric]) | 
| 36350 | 71 | apply (simp add: field_simps) | 
| 41541 | 72 | done | 
| 31719 | 73 | |
| 74 | end | |
| 75 | ||
| 76 | sublocale residues < cring | |
| 77 | by (rule cring) | |
| 78 | ||
| 79 | ||
| 41541 | 80 | context residues | 
| 81 | begin | |
| 31719 | 82 | |
| 44872 | 83 | (* These lemmas translate back and forth between internal and | 
| 31719 | 84 | external concepts *) | 
| 85 | ||
| 86 | lemma res_carrier_eq: "carrier R = {0..m - 1}"
 | |
| 44872 | 87 | unfolding R_def residue_ring_def by auto | 
| 31719 | 88 | |
| 89 | lemma res_add_eq: "x \<oplus> y = (x + y) mod m" | |
| 44872 | 90 | unfolding R_def residue_ring_def by auto | 
| 31719 | 91 | |
| 92 | lemma res_mult_eq: "x \<otimes> y = (x * y) mod m" | |
| 44872 | 93 | unfolding R_def residue_ring_def by auto | 
| 31719 | 94 | |
| 95 | lemma res_zero_eq: "\<zero> = 0" | |
| 44872 | 96 | unfolding R_def residue_ring_def by auto | 
| 31719 | 97 | |
| 98 | lemma res_one_eq: "\<one> = 1" | |
| 44872 | 99 | unfolding R_def residue_ring_def units_of_def by auto | 
| 31719 | 100 | |
| 101 | lemma res_units_eq: "Units R = { x. 0 < x & x < m & coprime x m}"
 | |
| 102 | apply (insert m_gt_one) | |
| 103 | apply (unfold Units_def R_def residue_ring_def) | |
| 104 | apply auto | |
| 105 | apply (subgoal_tac "x ~= 0") | |
| 106 | apply auto | |
| 31952 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 nipkow parents: 
31798diff
changeset | 107 | apply (rule invertible_coprime_int) | 
| 31719 | 108 | apply (subgoal_tac "x ~= 0") | 
| 109 | apply auto | |
| 31952 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 nipkow parents: 
31798diff
changeset | 110 | apply (subst (asm) coprime_iff_invertible'_int) | 
| 31719 | 111 | apply (rule m_gt_one) | 
| 112 | apply (auto simp add: cong_int_def mult_commute) | |
| 41541 | 113 | done | 
| 31719 | 114 | |
| 115 | lemma res_neg_eq: "\<ominus> x = (- x) mod m" | |
| 116 | apply (insert m_gt_one) | |
| 117 | apply (unfold R_def a_inv_def m_inv_def residue_ring_def) | |
| 118 | apply auto | |
| 119 | apply (rule the_equality) | |
| 120 | apply auto | |
| 121 | apply (subst mod_add_right_eq [symmetric]) | |
| 122 | apply auto | |
| 123 | apply (subst mod_add_left_eq [symmetric]) | |
| 124 | apply auto | |
| 125 | apply (subgoal_tac "y mod m = - x mod m") | |
| 126 | apply simp | |
| 127 | apply (subst zmod_eq_dvd_iff) | |
| 128 | apply auto | |
| 41541 | 129 | done | 
| 31719 | 130 | |
| 44872 | 131 | lemma finite [iff]: "finite (carrier R)" | 
| 31719 | 132 | by (subst res_carrier_eq, auto) | 
| 133 | ||
| 50027 
7747a9f4c358
adjusting proofs as the set_comprehension_pointfree simproc breaks some existing proofs
 bulwahn parents: 
47163diff
changeset | 134 | declare [[simproc del: finite_Collect]] | 
| 44872 | 135 | lemma finite_Units [iff]: "finite (Units R)" | 
| 50027 
7747a9f4c358
adjusting proofs as the set_comprehension_pointfree simproc breaks some existing proofs
 bulwahn parents: 
47163diff
changeset | 136 | by (subst res_units_eq) auto | 
| 
7747a9f4c358
adjusting proofs as the set_comprehension_pointfree simproc breaks some existing proofs
 bulwahn parents: 
47163diff
changeset | 137 | declare [[simproc add: finite_Collect]] | 
| 31719 | 138 | |
| 44872 | 139 | (* The function a -> a mod m maps the integers to the | 
| 140 | residue classes. The following lemmas show that this mapping | |
| 31719 | 141 | respects addition and multiplication on the integers. *) | 
| 142 | ||
| 143 | lemma mod_in_carrier [iff]: "a mod m : carrier R" | |
| 144 | apply (unfold res_carrier_eq) | |
| 145 | apply (insert m_gt_one, auto) | |
| 41541 | 146 | done | 
| 31719 | 147 | |
| 148 | lemma add_cong: "(x mod m) \<oplus> (y mod m) = (x + y) mod m" | |
| 44872 | 149 | unfolding R_def residue_ring_def | 
| 150 | apply auto | |
| 151 | apply presburger | |
| 152 | done | |
| 31719 | 153 | |
| 154 | lemma mult_cong: "(x mod m) \<otimes> (y mod m) = (x * y) mod m" | |
| 155 | apply (unfold R_def residue_ring_def, auto) | |
| 47163 | 156 | apply (subst mod_mult_right_eq [symmetric]) | 
| 31719 | 157 | apply (subst mult_commute) | 
| 47163 | 158 | apply (subst mod_mult_right_eq [symmetric]) | 
| 31719 | 159 | apply (subst mult_commute) | 
| 160 | apply auto | |
| 41541 | 161 | done | 
| 31719 | 162 | |
| 163 | lemma zero_cong: "\<zero> = 0" | |
| 44872 | 164 | unfolding R_def residue_ring_def by auto | 
| 31719 | 165 | |
| 166 | lemma one_cong: "\<one> = 1 mod m" | |
| 44872 | 167 | using m_gt_one unfolding R_def residue_ring_def by auto | 
| 31719 | 168 | |
| 169 | (* revise algebra library to use 1? *) | |
| 170 | lemma pow_cong: "(x mod m) (^) n = x^n mod m" | |
| 171 | apply (insert m_gt_one) | |
| 172 | apply (induct n) | |
| 41541 | 173 | apply (auto simp add: nat_pow_def one_cong) | 
| 31719 | 174 | apply (subst mult_commute) | 
| 175 | apply (rule mult_cong) | |
| 41541 | 176 | done | 
| 31719 | 177 | |
| 178 | lemma neg_cong: "\<ominus> (x mod m) = (- x) mod m" | |
| 179 | apply (rule sym) | |
| 180 | apply (rule sum_zero_eq_neg) | |
| 181 | apply auto | |
| 182 | apply (subst add_cong) | |
| 183 | apply (subst zero_cong) | |
| 184 | apply auto | |
| 41541 | 185 | done | 
| 31719 | 186 | |
| 44872 | 187 | lemma (in residues) prod_cong: | 
| 188 | "finite A \<Longrightarrow> (\<Otimes> i:A. (f i) mod m) = (PROD i:A. f i) mod m" | |
| 31719 | 189 | apply (induct set: finite) | 
| 31727 | 190 | apply (auto simp: one_cong mult_cong) | 
| 41541 | 191 | done | 
| 31719 | 192 | |
| 193 | lemma (in residues) sum_cong: | |
| 44872 | 194 | "finite A \<Longrightarrow> (\<Oplus> i:A. (f i) mod m) = (SUM i: A. f i) mod m" | 
| 31719 | 195 | apply (induct set: finite) | 
| 31727 | 196 | apply (auto simp: zero_cong add_cong) | 
| 41541 | 197 | done | 
| 31719 | 198 | |
| 44872 | 199 | lemma mod_in_res_units [simp]: "1 < m \<Longrightarrow> coprime a m \<Longrightarrow> | 
| 31719 | 200 | a mod m : Units R" | 
| 201 | apply (subst res_units_eq, auto) | |
| 202 | apply (insert pos_mod_sign [of m a]) | |
| 203 | apply (subgoal_tac "a mod m ~= 0") | |
| 204 | apply arith | |
| 205 | apply auto | |
| 31952 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 nipkow parents: 
31798diff
changeset | 206 | apply (subst (asm) gcd_red_int) | 
| 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 nipkow parents: 
31798diff
changeset | 207 | apply (subst gcd_commute_int, assumption) | 
| 41541 | 208 | done | 
| 31719 | 209 | |
| 44872 | 210 | lemma res_eq_to_cong: "((a mod m) = (b mod m)) = [a = b] (mod (m::int))" | 
| 31719 | 211 | unfolding cong_int_def by auto | 
| 212 | ||
| 44872 | 213 | (* Simplifying with these will translate a ring equation in R to a | 
| 31719 | 214 | congruence. *) | 
| 215 | ||
| 216 | lemmas res_to_cong_simps = add_cong mult_cong pow_cong one_cong | |
| 217 | prod_cong sum_cong neg_cong res_eq_to_cong | |
| 218 | ||
| 219 | (* Other useful facts about the residue ring *) | |
| 220 | ||
| 221 | lemma one_eq_neg_one: "\<one> = \<ominus> \<one> \<Longrightarrow> m = 2" | |
| 222 | apply (simp add: res_one_eq res_neg_eq) | |
| 223 | apply (insert m_gt_one) | |
| 224 | apply (subgoal_tac "~(m > 2)") | |
| 225 | apply arith | |
| 226 | apply (rule notI) | |
| 227 | apply (subgoal_tac "-1 mod m = m - 1") | |
| 228 | apply force | |
| 229 | apply (subst mod_add_self2 [symmetric]) | |
| 230 | apply (subst mod_pos_pos_trivial) | |
| 231 | apply auto | |
| 41541 | 232 | done | 
| 31719 | 233 | |
| 234 | end | |
| 235 | ||
| 236 | ||
| 237 | (* prime residues *) | |
| 238 | ||
| 239 | locale residues_prime = | |
| 240 | fixes p :: int and R (structure) | |
| 241 | assumes p_prime [intro]: "prime p" | |
| 242 | defines "R == residue_ring p" | |
| 243 | ||
| 244 | sublocale residues_prime < residues p | |
| 245 | apply (unfold R_def residues_def) | |
| 246 | using p_prime apply auto | |
| 41541 | 247 | done | 
| 31719 | 248 | |
| 44872 | 249 | context residues_prime | 
| 250 | begin | |
| 31719 | 251 | |
| 252 | lemma is_field: "field R" | |
| 253 | apply (rule cring.field_intro2) | |
| 254 | apply (rule cring) | |
| 44872 | 255 | apply (auto simp add: res_carrier_eq res_one_eq res_zero_eq res_units_eq) | 
| 31719 | 256 | apply (rule classical) | 
| 257 | apply (erule notE) | |
| 31952 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 nipkow parents: 
31798diff
changeset | 258 | apply (subst gcd_commute_int) | 
| 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 nipkow parents: 
31798diff
changeset | 259 | apply (rule prime_imp_coprime_int) | 
| 31719 | 260 | apply (rule p_prime) | 
| 261 | apply (rule notI) | |
| 262 | apply (frule zdvd_imp_le) | |
| 263 | apply auto | |
| 41541 | 264 | done | 
| 31719 | 265 | |
| 266 | lemma res_prime_units_eq: "Units R = {1..p - 1}"
 | |
| 267 | apply (subst res_units_eq) | |
| 268 | apply auto | |
| 31952 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 nipkow parents: 
31798diff
changeset | 269 | apply (subst gcd_commute_int) | 
| 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 nipkow parents: 
31798diff
changeset | 270 | apply (rule prime_imp_coprime_int) | 
| 31719 | 271 | apply (rule p_prime) | 
| 272 | apply (rule zdvd_not_zless) | |
| 273 | apply auto | |
| 41541 | 274 | done | 
| 31719 | 275 | |
| 276 | end | |
| 277 | ||
| 278 | sublocale residues_prime < field | |
| 279 | by (rule is_field) | |
| 280 | ||
| 281 | ||
| 282 | (* | |
| 283 | Test cases: Euler's theorem and Wilson's theorem. | |
| 284 | *) | |
| 285 | ||
| 286 | ||
| 287 | subsection{* Euler's theorem *}
 | |
| 288 | ||
| 289 | (* the definition of the phi function *) | |
| 290 | ||
| 44872 | 291 | definition phi :: "int => nat" | 
| 292 |   where "phi m = card({ x. 0 < x & x < m & gcd x m = 1})"
 | |
| 31719 | 293 | |
| 294 | lemma phi_zero [simp]: "phi 0 = 0" | |
| 295 | apply (subst phi_def) | |
| 44872 | 296 | (* Auto hangs here. Once again, where is the simplification rule | 
| 31719 | 297 | 1 == Suc 0 coming from? *) | 
| 298 | apply (auto simp add: card_eq_0_iff) | |
| 299 | (* Add card_eq_0_iff as a simp rule? delete card_empty_imp? *) | |
| 41541 | 300 | done | 
| 31719 | 301 | |
| 302 | lemma phi_one [simp]: "phi 1 = 0" | |
| 44872 | 303 | by (auto simp add: phi_def card_eq_0_iff) | 
| 31719 | 304 | |
| 305 | lemma (in residues) phi_eq: "phi m = card(Units R)" | |
| 306 | by (simp add: phi_def res_units_eq) | |
| 307 | ||
| 44872 | 308 | lemma (in residues) euler_theorem1: | 
| 31719 | 309 | assumes a: "gcd a m = 1" | 
| 310 | shows "[a^phi m = 1] (mod m)" | |
| 311 | proof - | |
| 312 | from a m_gt_one have [simp]: "a mod m : Units R" | |
| 313 | by (intro mod_in_res_units) | |
| 314 | from phi_eq have "(a mod m) (^) (phi m) = (a mod m) (^) (card (Units R))" | |
| 315 | by simp | |
| 44872 | 316 | also have "\<dots> = \<one>" | 
| 31719 | 317 | by (intro units_power_order_eq_one, auto) | 
| 318 | finally show ?thesis | |
| 319 | by (simp add: res_to_cong_simps) | |
| 320 | qed | |
| 321 | ||
| 322 | (* In fact, there is a two line proof! | |
| 323 | ||
| 44872 | 324 | lemma (in residues) euler_theorem1: | 
| 31719 | 325 | assumes a: "gcd a m = 1" | 
| 326 | shows "[a^phi m = 1] (mod m)" | |
| 327 | proof - | |
| 328 | have "(a mod m) (^) (phi m) = \<one>" | |
| 329 | by (simp add: phi_eq units_power_order_eq_one a m_gt_one) | |
| 44872 | 330 | then show ?thesis | 
| 31719 | 331 | by (simp add: res_to_cong_simps) | 
| 332 | qed | |
| 333 | ||
| 334 | *) | |
| 335 | ||
| 336 | (* outside the locale, we can relax the restriction m > 1 *) | |
| 337 | ||
| 338 | lemma euler_theorem: | |
| 339 | assumes "m >= 0" and "gcd a m = 1" | |
| 340 | shows "[a^phi m = 1] (mod m)" | |
| 341 | proof (cases) | |
| 342 | assume "m = 0 | m = 1" | |
| 44872 | 343 | then show ?thesis by auto | 
| 31719 | 344 | next | 
| 345 | assume "~(m = 0 | m = 1)" | |
| 41541 | 346 | with assms show ?thesis | 
| 31719 | 347 | by (intro residues.euler_theorem1, unfold residues_def, auto) | 
| 348 | qed | |
| 349 | ||
| 350 | lemma (in residues_prime) phi_prime: "phi p = (nat p - 1)" | |
| 351 | apply (subst phi_eq) | |
| 352 | apply (subst res_prime_units_eq) | |
| 353 | apply auto | |
| 41541 | 354 | done | 
| 31719 | 355 | |
| 356 | lemma phi_prime: "prime p \<Longrightarrow> phi p = (nat p - 1)" | |
| 357 | apply (rule residues_prime.phi_prime) | |
| 358 | apply (erule residues_prime.intro) | |
| 41541 | 359 | done | 
| 31719 | 360 | |
| 361 | lemma fermat_theorem: | |
| 362 | assumes "prime p" and "~ (p dvd a)" | |
| 363 | shows "[a^(nat p - 1) = 1] (mod p)" | |
| 364 | proof - | |
| 41541 | 365 | from assms have "[a^phi p = 1] (mod p)" | 
| 31719 | 366 | apply (intro euler_theorem) | 
| 367 | (* auto should get this next part. matching across | |
| 368 | substitutions is needed. *) | |
| 31952 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 nipkow parents: 
31798diff
changeset | 369 | apply (frule prime_gt_1_int, arith) | 
| 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 nipkow parents: 
31798diff
changeset | 370 | apply (subst gcd_commute_int, erule prime_imp_coprime_int, assumption) | 
| 31719 | 371 | done | 
| 372 | also have "phi p = nat p - 1" | |
| 41541 | 373 | by (rule phi_prime, rule assms) | 
| 31719 | 374 | finally show ?thesis . | 
| 375 | qed | |
| 376 | ||
| 377 | ||
| 378 | subsection {* Wilson's theorem *}
 | |
| 379 | ||
| 44872 | 380 | lemma (in field) inv_pair_lemma: "x : Units R \<Longrightarrow> y : Units R \<Longrightarrow> | 
| 381 |     {x, inv x} ~= {y, inv y} \<Longrightarrow> {x, inv x} Int {y, inv y} = {}"
 | |
| 31719 | 382 | apply auto | 
| 383 | apply (erule notE) | |
| 384 | apply (erule inv_eq_imp_eq) | |
| 385 | apply auto | |
| 386 | apply (erule notE) | |
| 387 | apply (erule inv_eq_imp_eq) | |
| 388 | apply auto | |
| 41541 | 389 | done | 
| 31719 | 390 | |
| 391 | lemma (in residues_prime) wilson_theorem1: | |
| 392 | assumes a: "p > 2" | |
| 393 | shows "[fact (p - 1) = - 1] (mod p)" | |
| 394 | proof - | |
| 44872 | 395 |   let ?InversePairs = "{ {x, inv x} | x. x : Units R - {\<one>, \<ominus> \<one>}}"
 | 
| 31732 | 396 |   have UR: "Units R = {\<one>, \<ominus> \<one>} Un (Union ?InversePairs)"
 | 
| 31719 | 397 | by auto | 
| 44872 | 398 | have "(\<Otimes>i: Units R. i) = | 
| 31719 | 399 |     (\<Otimes>i: {\<one>, \<ominus> \<one>}. i) \<otimes> (\<Otimes>i: Union ?InversePairs. i)"
 | 
| 31732 | 400 | apply (subst UR) | 
| 31719 | 401 | apply (subst finprod_Un_disjoint) | 
| 31732 | 402 | apply (auto intro:funcsetI) | 
| 31719 | 403 | apply (drule sym, subst (asm) inv_eq_one_eq) | 
| 404 | apply auto | |
| 405 | apply (drule sym, subst (asm) inv_eq_neg_one_eq) | |
| 406 | apply auto | |
| 407 | done | |
| 408 |   also have "(\<Otimes>i: {\<one>, \<ominus> \<one>}. i) = \<ominus> \<one>"
 | |
| 409 | apply (subst finprod_insert) | |
| 410 | apply auto | |
| 411 | apply (frule one_eq_neg_one) | |
| 412 | apply (insert a, force) | |
| 413 | done | |
| 44872 | 414 | also have "(\<Otimes>i:(Union ?InversePairs). i) = | 
| 41541 | 415 | (\<Otimes>A: ?InversePairs. (\<Otimes>y:A. y))" | 
| 31719 | 416 | apply (subst finprod_Union_disjoint) | 
| 417 | apply force | |
| 418 | apply force | |
| 419 | apply clarify | |
| 420 | apply (rule inv_pair_lemma) | |
| 421 | apply auto | |
| 422 | done | |
| 423 | also have "\<dots> = \<one>" | |
| 424 | apply (rule finprod_one) | |
| 425 | apply auto | |
| 426 | apply (subst finprod_insert) | |
| 427 | apply auto | |
| 428 | apply (frule inv_eq_self) | |
| 31732 | 429 | apply (auto) | 
| 31719 | 430 | done | 
| 431 | finally have "(\<Otimes>i: Units R. i) = \<ominus> \<one>" | |
| 432 | by simp | |
| 433 | also have "(\<Otimes>i: Units R. i) = (\<Otimes>i: Units R. i mod p)" | |
| 434 | apply (rule finprod_cong') | |
| 31732 | 435 | apply (auto) | 
| 31719 | 436 | apply (subst (asm) res_prime_units_eq) | 
| 437 | apply auto | |
| 438 | done | |
| 439 | also have "\<dots> = (PROD i: Units R. i) mod p" | |
| 440 | apply (rule prod_cong) | |
| 441 | apply auto | |
| 442 | done | |
| 443 | also have "\<dots> = fact (p - 1) mod p" | |
| 31952 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 nipkow parents: 
31798diff
changeset | 444 | apply (subst fact_altdef_int) | 
| 41541 | 445 | apply (insert assms, force) | 
| 31719 | 446 | apply (subst res_prime_units_eq, rule refl) | 
| 447 | done | |
| 448 | finally have "fact (p - 1) mod p = \<ominus> \<one>". | |
| 44872 | 449 | then show ?thesis by (simp add: res_to_cong_simps) | 
| 31719 | 450 | qed | 
| 451 | ||
| 452 | lemma wilson_theorem: "prime (p::int) \<Longrightarrow> [fact (p - 1) = - 1] (mod p)" | |
| 31952 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 nipkow parents: 
31798diff
changeset | 453 | apply (frule prime_gt_1_int) | 
| 31719 | 454 | apply (case_tac "p = 2") | 
| 31952 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 nipkow parents: 
31798diff
changeset | 455 | apply (subst fact_altdef_int, simp) | 
| 31719 | 456 | apply (subst cong_int_def) | 
| 457 | apply simp | |
| 458 | apply (rule residues_prime.wilson_theorem1) | |
| 459 | apply (rule residues_prime.intro) | |
| 460 | apply auto | |
| 44872 | 461 | done | 
| 31719 | 462 | |
| 463 | end |