| author | wenzelm | 
| Tue, 29 Sep 2020 20:08:08 +0200 | |
| changeset 72341 | 0973a594be72 | 
| parent 71252 | c5914bdd896b | 
| child 75963 | 884dbbc8e1b3 | 
| 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.Multiplicative_Group" | 
| 65465 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 14 | Totient | 
| 31719 | 15 | begin | 
| 16 | ||
| 66305 | 17 | definition QuadRes :: "int \<Rightarrow> int \<Rightarrow> bool" | 
| 18 | 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 | 19 | |
| 66305 | 20 | definition Legendre :: "int \<Rightarrow> int \<Rightarrow> int" | 
| 21 | where "Legendre a p = | |
| 22 | (if ([a = 0] (mod p)) then 0 | |
| 23 | else if QuadRes p a then 1 | |
| 24 | else -1)" | |
| 25 | ||
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: 
64272diff
changeset | 26 | |
| 60527 | 27 | subsection \<open>A locale for residue rings\<close> | 
| 31719 | 28 | |
| 60527 | 29 | definition residue_ring :: "int \<Rightarrow> int ring" | 
| 66305 | 30 | where | 
| 31 | "residue_ring m = | |
| 32 |       \<lparr>carrier = {0..m - 1},
 | |
| 33 | monoid.mult = \<lambda>x y. (x * y) mod m, | |
| 34 | one = 1, | |
| 35 | zero = 0, | |
| 36 | add = \<lambda>x y. (x + y) mod m\<rparr>" | |
| 31719 | 37 | |
| 38 | locale residues = | |
| 39 | fixes m :: int and R (structure) | |
| 40 | assumes m_gt_one: "m > 1" | |
| 60527 | 41 | defines "R \<equiv> residue_ring m" | 
| 44872 | 42 | begin | 
| 31719 | 43 | |
| 44 | lemma abelian_group: "abelian_group R" | |
| 65066 | 45 | proof - | 
| 46 |   have "\<exists>y\<in>{0..m - 1}. (x + y) mod m = 0" if "0 \<le> x" "x < m" for x
 | |
| 47 | proof (cases "x = 0") | |
| 48 | case True | |
| 49 | with m_gt_one show ?thesis by simp | |
| 50 | next | |
| 51 | case False | |
| 52 | then have "(x + (m - x)) mod m = 0" | |
| 53 | by simp | |
| 54 | with m_gt_one that show ?thesis | |
| 55 | by (metis False atLeastAtMost_iff diff_ge_0_iff_ge diff_left_mono int_one_le_iff_zero_less less_le) | |
| 56 | qed | |
| 57 | with m_gt_one show ?thesis | |
| 58 | by (fastforce simp add: R_def residue_ring_def mod_add_right_eq ac_simps intro!: abelian_groupI) | |
| 65899 | 59 | qed | 
| 31719 | 60 | |
| 61 | lemma comm_monoid: "comm_monoid R" | |
| 65066 | 62 | unfolding R_def residue_ring_def | 
| 31719 | 63 | apply (rule comm_monoidI) | 
| 65066 | 64 | using m_gt_one apply auto | 
| 65 | apply (metis mod_mult_right_eq mult.assoc mult.commute) | |
| 66 | apply (metis mult.commute) | |
| 41541 | 67 | done | 
| 31719 | 68 | |
| 69 | lemma cring: "cring R" | |
| 65066 | 70 | apply (intro cringI abelian_group comm_monoid) | 
| 71 | unfolding R_def residue_ring_def | |
| 72 | apply (auto simp add: comm_semiring_class.distrib mod_add_eq mod_mult_left_eq) | |
| 41541 | 73 | done | 
| 31719 | 74 | |
| 75 | end | |
| 76 | ||
| 77 | sublocale residues < cring | |
| 78 | by (rule cring) | |
| 79 | ||
| 80 | ||
| 41541 | 81 | context residues | 
| 82 | begin | |
| 31719 | 83 | |
| 60527 | 84 | text \<open> | 
| 85 | These lemmas translate back and forth between internal and | |
| 86 | external concepts. | |
| 87 | \<close> | |
| 31719 | 88 | |
| 89 | lemma res_carrier_eq: "carrier R = {0..m - 1}"
 | |
| 66305 | 90 | by (auto simp: R_def residue_ring_def) | 
| 31719 | 91 | |
| 92 | lemma res_add_eq: "x \<oplus> y = (x + y) mod m" | |
| 66305 | 93 | by (auto simp: R_def residue_ring_def) | 
| 31719 | 94 | |
| 95 | lemma res_mult_eq: "x \<otimes> y = (x * y) mod m" | |
| 66305 | 96 | by (auto simp: R_def residue_ring_def) | 
| 31719 | 97 | |
| 98 | lemma res_zero_eq: "\<zero> = 0" | |
| 66305 | 99 | by (auto simp: R_def residue_ring_def) | 
| 31719 | 100 | |
| 101 | lemma res_one_eq: "\<one> = 1" | |
| 66305 | 102 | by (auto simp: R_def residue_ring_def units_of_def) | 
| 31719 | 103 | |
| 60527 | 104 | lemma res_units_eq: "Units R = {x. 0 < x \<and> x < m \<and> coprime x m}"
 | 
| 65066 | 105 | using m_gt_one | 
| 67051 | 106 | apply (auto simp add: Units_def R_def residue_ring_def ac_simps invertible_coprime intro: ccontr) | 
| 31952 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 nipkow parents: 
31798diff
changeset | 107 | apply (subst (asm) coprime_iff_invertible'_int) | 
| 67051 | 108 | apply (auto simp add: cong_def) | 
| 41541 | 109 | done | 
| 31719 | 110 | |
| 111 | lemma res_neg_eq: "\<ominus> x = (- x) mod m" | |
| 65066 | 112 | using m_gt_one unfolding R_def a_inv_def m_inv_def residue_ring_def | 
| 113 | apply simp | |
| 31719 | 114 | apply (rule the_equality) | 
| 66305 | 115 | apply (simp add: mod_add_right_eq) | 
| 116 | apply (simp add: add.commute mod_add_right_eq) | |
| 65066 | 117 | apply (metis add.right_neutral minus_add_cancel mod_add_right_eq mod_pos_pos_trivial) | 
| 41541 | 118 | done | 
| 31719 | 119 | |
| 44872 | 120 | lemma finite [iff]: "finite (carrier R)" | 
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 121 | by (simp add: res_carrier_eq) | 
| 31719 | 122 | |
| 44872 | 123 | lemma finite_Units [iff]: "finite (Units R)" | 
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 124 | by (simp add: finite_ring_finite_units) | 
| 31719 | 125 | |
| 60527 | 126 | text \<open> | 
| 63167 | 127 | The function \<open>a \<mapsto> a mod m\<close> maps the integers to the | 
| 60527 | 128 | residue classes. The following lemmas show that this mapping | 
| 129 | respects addition and multiplication on the integers. | |
| 130 | \<close> | |
| 31719 | 131 | |
| 60527 | 132 | lemma mod_in_carrier [iff]: "a mod m \<in> carrier R" | 
| 133 | unfolding res_carrier_eq | |
| 134 | using insert m_gt_one by auto | |
| 31719 | 135 | |
| 136 | lemma add_cong: "(x mod m) \<oplus> (y mod m) = (x + y) mod m" | |
| 66305 | 137 | by (auto simp: R_def residue_ring_def mod_simps) | 
| 31719 | 138 | |
| 139 | lemma mult_cong: "(x mod m) \<otimes> (y mod m) = (x * y) mod m" | |
| 66305 | 140 | by (auto simp: R_def residue_ring_def mod_simps) | 
| 31719 | 141 | |
| 142 | lemma zero_cong: "\<zero> = 0" | |
| 66305 | 143 | by (auto simp: R_def residue_ring_def) | 
| 31719 | 144 | |
| 145 | lemma one_cong: "\<one> = 1 mod m" | |
| 66305 | 146 | using m_gt_one by (auto simp: R_def residue_ring_def) | 
| 31719 | 147 | |
| 60527 | 148 | (* FIXME revise algebra library to use 1? *) | 
| 67341 
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
 nipkow parents: 
67091diff
changeset | 149 | lemma pow_cong: "(x mod m) [^] n = x^n mod m" | 
| 65066 | 150 | using m_gt_one | 
| 31719 | 151 | apply (induct n) | 
| 41541 | 152 | 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 | 153 | apply (metis mult.commute mult_cong) | 
| 41541 | 154 | done | 
| 31719 | 155 | |
| 156 | lemma neg_cong: "\<ominus> (x mod m) = (- x) mod m" | |
| 55352 | 157 | by (metis mod_minus_eq res_neg_eq) | 
| 31719 | 158 | |
| 60528 | 159 | 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 | 160 | by (induct set: finite) (auto simp: one_cong mult_cong) | 
| 31719 | 161 | |
| 60528 | 162 | 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 | 163 | by (induct set: finite) (auto simp: zero_cong add_cong) | 
| 31719 | 164 | |
| 60688 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60528diff
changeset | 165 | 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 | 166 | 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 | 167 | 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 | 168 | proof (cases "a mod m = 0") | 
| 66305 | 169 | case True | 
| 170 | 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 | 171 | 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 | 172 | next | 
| 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60528diff
changeset | 173 | case False | 
| 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60528diff
changeset | 174 | from assms have "0 < m" by simp | 
| 66305 | 175 | 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 | 176 | 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 | 177 | 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 | 178 | 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 | 179 | qed | 
| 31719 | 180 | |
| 60528 | 181 | lemma res_eq_to_cong: "(a mod m) = (b mod m) \<longleftrightarrow> [a = b] (mod m)" | 
| 66888 | 182 | by (auto simp: cong_def) | 
| 31719 | 183 | |
| 184 | ||
| 60528 | 185 | text \<open>Simplifying with these will translate a ring equation in R to a congruence.\<close> | 
| 66305 | 186 | lemmas res_to_cong_simps = | 
| 187 | add_cong mult_cong pow_cong one_cong | |
| 188 | prod_cong sum_cong neg_cong res_eq_to_cong | |
| 31719 | 189 | |
| 60527 | 190 | text \<open>Other useful facts about the residue ring.\<close> | 
| 31719 | 191 | lemma one_eq_neg_one: "\<one> = \<ominus> \<one> \<Longrightarrow> m = 2" | 
| 192 | 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 | 193 | apply (metis add.commute add_diff_cancel mod_mod_trivial one_add_one uminus_add_conv_diff | 
| 60528 | 194 | zero_neq_one zmod_zminus1_eq_if) | 
| 41541 | 195 | done | 
| 31719 | 196 | |
| 197 | end | |
| 198 | ||
| 199 | ||
| 60527 | 200 | subsection \<open>Prime residues\<close> | 
| 31719 | 201 | |
| 202 | locale residues_prime = | |
| 63534 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 eberlm <eberlm@in.tum.de> parents: 
63417diff
changeset | 203 | fixes p :: nat and R (structure) | 
| 31719 | 204 | assumes p_prime [intro]: "prime p" | 
| 63534 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 eberlm <eberlm@in.tum.de> parents: 
63417diff
changeset | 205 | defines "R \<equiv> residue_ring (int p)" | 
| 31719 | 206 | |
| 207 | sublocale residues_prime < residues p | |
| 65066 | 208 | unfolding R_def residues_def | 
| 31719 | 209 | using p_prime apply auto | 
| 62348 | 210 | apply (metis (full_types) of_nat_1 of_nat_less_iff prime_gt_1_nat) | 
| 41541 | 211 | done | 
| 31719 | 212 | |
| 44872 | 213 | context residues_prime | 
| 214 | begin | |
| 31719 | 215 | |
| 67051 | 216 | lemma p_coprime_left: | 
| 217 | "coprime p a \<longleftrightarrow> \<not> p dvd a" | |
| 218 | using p_prime by (auto intro: prime_imp_coprime dest: coprime_common_divisor) | |
| 219 | ||
| 220 | lemma p_coprime_right: | |
| 221 | "coprime a p \<longleftrightarrow> \<not> p dvd a" | |
| 222 | using p_coprime_left [of a] by (simp add: ac_simps) | |
| 223 | ||
| 224 | lemma p_coprime_left_int: | |
| 225 | "coprime (int p) a \<longleftrightarrow> \<not> int p dvd a" | |
| 226 | using p_prime by (auto intro: prime_imp_coprime dest: coprime_common_divisor) | |
| 227 | ||
| 228 | lemma p_coprime_right_int: | |
| 229 | "coprime a (int p) \<longleftrightarrow> \<not> int p dvd a" | |
| 230 | using p_coprime_left_int [of a] by (simp add: ac_simps) | |
| 231 | ||
| 31719 | 232 | lemma is_field: "field R" | 
| 65066 | 233 | proof - | 
| 66837 | 234 | have "0 < x \<Longrightarrow> x < int p \<Longrightarrow> coprime (int p) x" for x | 
| 235 | by (rule prime_imp_coprime) (auto simp add: zdvd_not_zless) | |
| 65066 | 236 | then show ?thesis | 
| 66837 | 237 | by (intro cring.field_intro2 cring) | 
| 238 | (auto simp add: res_carrier_eq res_one_eq res_zero_eq res_units_eq ac_simps) | |
| 65066 | 239 | qed | 
| 31719 | 240 | |
| 241 | lemma res_prime_units_eq: "Units R = {1..p - 1}"
 | |
| 242 | apply (subst res_units_eq) | |
| 67051 | 243 | apply (auto simp add: p_coprime_right_int zdvd_not_zless) | 
| 41541 | 244 | done | 
| 31719 | 245 | |
| 246 | end | |
| 247 | ||
| 248 | sublocale residues_prime < field | |
| 249 | by (rule is_field) | |
| 250 | ||
| 251 | ||
| 60527 | 252 | section \<open>Test cases: Euler's theorem and Wilson's theorem\<close> | 
| 31719 | 253 | |
| 60527 | 254 | subsection \<open>Euler's theorem\<close> | 
| 31719 | 255 | |
| 67051 | 256 | lemma (in residues) totatives_eq: | 
| 257 | "totatives (nat m) = nat ` Units R" | |
| 55261 
ad3604df6bc6
new lemmas involving phi from Lehmer AFP entry
 paulson <lp15@cam.ac.uk> parents: 
55242diff
changeset | 258 | proof - | 
| 67051 | 259 | from m_gt_one have "\<bar>m\<bar> > 1" | 
| 260 | by simp | |
| 261 | then have "totatives (nat \<bar>m\<bar>) = nat ` abs ` Units R" | |
| 262 | by (auto simp add: totatives_def res_units_eq image_iff le_less) | |
| 263 | (use m_gt_one zless_nat_eq_int_zless in force) | |
| 264 | moreover have "\<bar>m\<bar> = m" "abs ` Units R = Units R" | |
| 265 | using m_gt_one by (auto simp add: res_units_eq image_iff) | |
| 266 | ultimately show ?thesis | |
| 267 | by simp | |
| 268 | qed | |
| 269 | ||
| 270 | lemma (in residues) totient_eq: | |
| 271 | "totient (nat m) = card (Units R)" | |
| 272 | proof - | |
| 65465 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 273 | have *: "inj_on nat (Units R)" | 
| 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 274 | by (rule inj_onI) (auto simp add: res_units_eq) | 
| 67051 | 275 | then show ?thesis | 
| 276 | by (simp add: totient_def totatives_eq card_image) | |
| 55261 
ad3604df6bc6
new lemmas involving phi from Lehmer AFP entry
 paulson <lp15@cam.ac.uk> parents: 
55242diff
changeset | 277 | qed | 
| 
ad3604df6bc6
new lemmas involving phi from Lehmer AFP entry
 paulson <lp15@cam.ac.uk> parents: 
55242diff
changeset | 278 | |
| 65465 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 279 | 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 | 280 | using totient_eq by (simp add: res_prime_units_eq) | 
| 31719 | 281 | |
| 65465 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 282 | lemma (in residues) euler_theorem: | 
| 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 283 | assumes "coprime a m" | 
| 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 284 | shows "[a ^ totient (nat m) = 1] (mod m)" | 
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 285 | proof - | 
| 65465 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 286 | 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 | 287 | by (metis assms finite_Units m_gt_one mod_in_res_units one_cong totient_eq pow_cong units_power_order_eq_one) | 
| 65066 | 288 | then show ?thesis | 
| 289 | using res_eq_to_cong by blast | |
| 31719 | 290 | qed | 
| 291 | ||
| 292 | lemma euler_theorem: | |
| 65465 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 293 | fixes a m :: nat | 
| 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 294 | assumes "coprime a m" | 
| 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 295 | shows "[a ^ totient m = 1] (mod m)" | 
| 67091 | 296 | proof (cases "m = 0 \<or> m = 1") | 
| 60527 | 297 | case True | 
| 44872 | 298 | then show ?thesis by auto | 
| 31719 | 299 | next | 
| 60527 | 300 | case False | 
| 41541 | 301 | with assms show ?thesis | 
| 66954 | 302 | using residues.euler_theorem [of "int m" "int a"] cong_int_iff | 
| 303 | by (auto simp add: residues_def gcd_int_def) fastforce | |
| 31719 | 304 | qed | 
| 305 | ||
| 306 | lemma fermat_theorem: | |
| 65465 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 307 | fixes p a :: nat | 
| 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 308 | assumes "prime p" and "\<not> p dvd a" | 
| 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 309 | shows "[a ^ (p - 1) = 1] (mod p)" | 
| 31719 | 310 | proof - | 
| 65465 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 311 | 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 | 312 | by (auto simp add: ac_simps) | 
| 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 313 | then have "[a ^ totient p = 1] (mod p)" | 
| 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 314 | by (rule euler_theorem) | 
| 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 315 | also have "totient p = p - 1" | 
| 65726 
f5d64d094efe
More material on totient function
 eberlm <eberlm@in.tum.de> parents: 
65465diff
changeset | 316 | by (rule totient_prime) (rule assms) | 
| 65465 
067210a08a22
more fundamental euler's totient function on nat rather than int;
 haftmann parents: 
65416diff
changeset | 317 | finally show ?thesis . | 
| 31719 | 318 | qed | 
| 319 | ||
| 320 | ||
| 60526 | 321 | subsection \<open>Wilson's theorem\<close> | 
| 31719 | 322 | |
| 60527 | 323 | lemma (in field) inv_pair_lemma: "x \<in> Units R \<Longrightarrow> y \<in> Units R \<Longrightarrow> | 
| 324 |     {x, inv x} \<noteq> {y, inv y} \<Longrightarrow> {x, inv x} \<inter> {y, inv y} = {}"
 | |
| 31719 | 325 | apply auto | 
| 55352 | 326 | apply (metis Units_inv_inv)+ | 
| 41541 | 327 | done | 
| 31719 | 328 | |
| 329 | lemma (in residues_prime) wilson_theorem1: | |
| 330 | assumes a: "p > 2" | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59667diff
changeset | 331 | shows "[fact (p - 1) = (-1::int)] (mod p)" | 
| 31719 | 332 | proof - | 
| 60527 | 333 |   let ?Inverse_Pairs = "{{x, inv x}| x. x \<in> Units R - {\<one>, \<ominus> \<one>}}"
 | 
| 334 |   have UR: "Units R = {\<one>, \<ominus> \<one>} \<union> \<Union>?Inverse_Pairs"
 | |
| 31719 | 335 | by auto | 
| 60527 | 336 |   have "(\<Otimes>i\<in>Units R. i) = (\<Otimes>i\<in>{\<one>, \<ominus> \<one>}. i) \<otimes> (\<Otimes>i\<in>\<Union>?Inverse_Pairs. i)"
 | 
| 31732 | 337 | apply (subst UR) | 
| 31719 | 338 | apply (subst finprod_Un_disjoint) | 
| 66305 | 339 | apply (auto intro: funcsetI) | 
| 60527 | 340 | using inv_one apply auto[1] | 
| 341 | using inv_eq_neg_one_eq apply auto | |
| 31719 | 342 | done | 
| 60527 | 343 |   also have "(\<Otimes>i\<in>{\<one>, \<ominus> \<one>}. i) = \<ominus> \<one>"
 | 
| 31719 | 344 | apply (subst finprod_insert) | 
| 66305 | 345 | apply auto | 
| 31719 | 346 | apply (frule one_eq_neg_one) | 
| 60527 | 347 | using a apply force | 
| 31719 | 348 | done | 
| 60527 | 349 | also have "(\<Otimes>i\<in>(\<Union>?Inverse_Pairs). i) = (\<Otimes>A\<in>?Inverse_Pairs. (\<Otimes>y\<in>A. y))" | 
| 350 | apply (subst finprod_Union_disjoint) | |
| 68458 | 351 | apply (auto simp: pairwise_def disjnt_def) | 
| 66305 | 352 | apply (metis Units_inv_inv)+ | 
| 31719 | 353 | done | 
| 354 | also have "\<dots> = \<one>" | |
| 68447 
0beb927eed89
Adjusting Number_Theory for new Algebra
 paulson <lp15@cam.ac.uk> parents: 
67341diff
changeset | 355 | apply (rule finprod_one_eqI) | 
| 66305 | 356 | apply auto | 
| 60527 | 357 | apply (subst finprod_insert) | 
| 66305 | 358 | apply auto | 
| 55352 | 359 | apply (metis inv_eq_self) | 
| 31719 | 360 | done | 
| 60527 | 361 | finally have "(\<Otimes>i\<in>Units R. i) = \<ominus> \<one>" | 
| 31719 | 362 | by simp | 
| 60527 | 363 | also have "(\<Otimes>i\<in>Units R. i) = (\<Otimes>i\<in>Units R. i mod p)" | 
| 65066 | 364 | by (rule finprod_cong') (auto simp: res_units_eq) | 
| 60527 | 365 | also have "\<dots> = (\<Prod>i\<in>Units R. i) mod p" | 
| 65066 | 366 | by (rule prod_cong) auto | 
| 31719 | 367 | also have "\<dots> = fact (p - 1) mod p" | 
| 64272 | 368 | apply (simp add: fact_prod) | 
| 65066 | 369 | 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 | 370 | apply (subst res_prime_units_eq) | 
| 64272 | 371 | apply (simp add: int_prod zmod_int prod_int_eq) | 
| 31719 | 372 | done | 
| 60527 | 373 | finally have "fact (p - 1) mod p = \<ominus> \<one>" . | 
| 374 | then show ?thesis | |
| 66888 | 375 | by (simp add: cong_def res_neg_eq res_one_eq zmod_int) | 
| 31719 | 376 | qed | 
| 377 | ||
| 55352 | 378 | lemma wilson_theorem: | 
| 60527 | 379 | assumes "prime p" | 
| 380 | shows "[fact (p - 1) = - 1] (mod p)" | |
| 55352 | 381 | 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 | 382 | case True | 
| 55352 | 383 | then show ?thesis | 
| 66888 | 384 | by (simp add: cong_def fact_prod) | 
| 55352 | 385 | next | 
| 386 | case False | |
| 387 | then show ?thesis | |
| 388 | using assms prime_ge_2_nat | |
| 389 | by (metis residues_prime.wilson_theorem1 residues_prime.intro le_eq_less_or_eq) | |
| 390 | qed | |
| 31719 | 391 | |
| 66304 | 392 | text \<open> | 
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 393 | This result can be transferred to the multiplicative group of | 
| 66305 | 394 | \<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 | 395 | |
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 396 | lemma mod_nat_int_pow_eq: | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 397 | fixes n :: nat and p a :: int | 
| 66305 | 398 | 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 | 399 | 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 | 400 | |
| 69785 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 401 | theorem residue_prime_mult_group_has_gen: | 
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 402 | fixes p :: nat | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 403 | assumes prime_p : "prime p" | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 404 |  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 | 405 | proof - | 
| 66305 | 406 | have "p \<ge> 2" | 
| 407 | using prime_gt_1_nat[OF prime_p] by simp | |
| 408 | interpret R: residues_prime p "residue_ring p" | |
| 409 | by (simp add: residues_prime_def prime_p) | |
| 410 |   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 | 411 | by (auto simp add: R.zero_cong R.res_carrier_eq) | 
| 66305 | 412 | |
| 67341 
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
 nipkow parents: 
67091diff
changeset | 413 | have "x [^]\<^bsub>residue_ring (int p)\<^esub> i = x ^ i mod (int p)" | 
| 66305 | 414 |     if "x \<in> {1 .. int p - 1}" for x and i :: nat
 | 
| 415 | using that R.pow_cong[of x i] by auto | |
| 416 | moreover | |
| 417 |   obtain a where a: "a \<in> {1 .. int p - 1}"
 | |
| 67341 
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
 nipkow parents: 
67091diff
changeset | 418 |     and a_gen: "{1 .. int p - 1} = {a[^]\<^bsub>residue_ring (int p)\<^esub>i|i::nat . i \<in> UNIV}"
 | 
| 66305 | 419 | 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 | 420 | by (auto simp add: car[symmetric] carrier_mult_of) | 
| 66305 | 421 | moreover | 
| 422 |   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 | 423 | proof | 
| 66305 | 424 | have "n \<in> ?R" if "n \<in> ?L" for n | 
| 425 | using that \<open>p\<ge>2\<close> by force | |
| 426 | then show "?L \<subseteq> ?R" by blast | |
| 427 | have "n \<in> ?L" if "n \<in> ?R" for n | |
| 66837 | 428 | using that \<open>p\<ge>2\<close> by (auto intro: rev_image_eqI [of "int n"]) | 
| 66305 | 429 | then show "?R \<subseteq> ?L" by blast | 
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 430 | qed | 
| 66305 | 431 | moreover | 
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 432 |   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 | 433 | proof | 
| 66305 | 434 | have "x \<in> ?R" if "x \<in> ?L" for x | 
| 435 | proof - | |
| 436 | from that obtain i where i: "x = nat (a^i mod (int p))" | |
| 437 | by blast | |
| 438 | then have "x = nat a ^ i mod p" | |
| 439 | using mod_nat_int_pow_eq[of a "int p" i] a \<open>p\<ge>2\<close> by auto | |
| 440 | with i show ?thesis by blast | |
| 441 | qed | |
| 442 | then show "?L \<subseteq> ?R" by blast | |
| 443 | have "x \<in> ?L" if "x \<in> ?R" for x | |
| 444 | proof - | |
| 445 | from that obtain i where i: "x = nat a^i mod p" | |
| 446 | by blast | |
| 447 | with mod_nat_int_pow_eq[of a "int p" i] a \<open>p\<ge>2\<close> show ?thesis | |
| 448 | by auto | |
| 449 | qed | |
| 450 | then show "?R \<subseteq> ?L" by blast | |
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 451 | qed | 
| 66305 | 452 |   ultimately have "{1 .. p - 1} = {nat a^i mod p | i. i \<in> UNIV}"
 | 
| 453 | by presburger | |
| 454 |   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 | 455 | ultimately show ?thesis .. | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 456 | qed | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65066diff
changeset | 457 | |
| 69785 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 458 | |
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 459 | subsection \<open>Upper bound for the number of $n$-th roots\<close> | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 460 | |
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 461 | lemma roots_mod_prime_bound: | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 462 | fixes n c p :: nat | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 463 | assumes "prime p" "n > 0" | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 464 |   defines "A \<equiv> {x\<in>{..<p}. [x ^ n = c] (mod p)}"
 | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 465 | shows "card A \<le> n" | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 466 | proof - | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 467 | define R where "R = residue_ring (int p)" | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 468 | from assms(1) interpret residues_prime p R | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 469 | by unfold_locales (simp_all add: R_def) | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 470 | interpret R: UP_domain R "UP R" by (unfold_locales) | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 471 | |
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 472 | let ?f = "UnivPoly.monom (UP R) \<one>\<^bsub>R\<^esub> n \<ominus>\<^bsub>(UP R)\<^esub> UnivPoly.monom (UP R) (int (c mod p)) 0" | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 473 | have in_carrier: "int (c mod p) \<in> carrier R" | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 474 | using prime_gt_1_nat[OF assms(1)] by (simp add: R_def residue_ring_def) | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 475 | |
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 476 | have "deg R ?f = n" | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 477 | using assms in_carrier by (simp add: R.deg_minus_eq) | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 478 | hence f_not_zero: "?f \<noteq> \<zero>\<^bsub>UP R\<^esub>" using assms by (auto simp add : R.deg_nzero_nzero) | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 479 |   have roots_bound: "finite {a \<in> carrier R. UnivPoly.eval R R id a ?f = \<zero>\<^bsub>R\<^esub>} \<and>
 | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 480 |                      card {a \<in> carrier R. UnivPoly.eval R R id a ?f = \<zero>\<^bsub>R\<^esub>} \<le> deg R ?f"
 | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 481 | using finite in_carrier by (intro R.roots_bound[OF _ f_not_zero]) simp | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 482 |   have subs: "{x \<in> carrier R. x [^]\<^bsub>R\<^esub> n = int (c mod p)} \<subseteq>
 | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 483 |                 {a \<in> carrier R. UnivPoly.eval R R id a ?f = \<zero>\<^bsub>R\<^esub>}"
 | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 484 | using in_carrier by (auto simp: R.evalRR_simps) | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 485 |   then have "card {x \<in> carrier R. x [^]\<^bsub>R\<^esub> n = int (c mod p)} \<le>
 | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 486 |                card {a \<in> carrier R. UnivPoly.eval R R id a ?f = \<zero>\<^bsub>R\<^esub>}"
 | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 487 | using finite by (intro card_mono) auto | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 488 | also have "\<dots> \<le> n" | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 489 | using \<open>deg R ?f = n\<close> roots_bound by linarith | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 490 |   also {
 | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 491 | fix x assume "x \<in> carrier R" | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 492 | hence "x [^]\<^bsub>R\<^esub> n = (x ^ n) mod (int p)" | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 493 | by (subst pow_cong [symmetric]) (auto simp: R_def residue_ring_def) | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 494 | } | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 495 |   hence "{x \<in> carrier R. x [^]\<^bsub>R\<^esub> n = int (c mod p)} = {x \<in> carrier R. [x ^ n = int c] (mod p)}"
 | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 496 | by (fastforce simp: cong_def zmod_int) | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 497 |   also have "bij_betw int A {x \<in> carrier R. [x ^ n = int c] (mod p)}"
 | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 498 | by (rule bij_betwI[of int _ _ nat]) | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 499 | (use cong_int_iff in \<open>force simp: R_def residue_ring_def A_def\<close>)+ | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 500 |   from bij_betw_same_card[OF this] have "card {x \<in> carrier R. [x ^ n = int c] (mod p)} = card A" ..
 | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 501 | finally show ?thesis . | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 502 | qed | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 503 | |
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
68458diff
changeset | 504 | |
| 31719 | 505 | end |