| author | wenzelm | 
| Mon, 26 Aug 2013 11:41:17 +0200 | |
| changeset 53199 | 7a9fe70c8b0a | 
| parent 53077 | a1b3784f8129 | 
| child 57512 | cc97b347b301 | 
| permissions | -rw-r--r-- | 
| 38159 | 1 | (* Title: HOL/Old_Number_Theory/Primes.thy | 
| 27106 | 2 | Author: Amine Chaieb, Christophe Tabacznyj and Lawrence C Paulson | 
| 11363 | 3 | Copyright 1996 University of Cambridge | 
| 4 | *) | |
| 5 | ||
| 16762 | 6 | header {* Primality on nat *}
 | 
| 11363 | 7 | |
| 15131 | 8 | theory Primes | 
| 31706 | 9 | imports Complex_Main Legacy_GCD | 
| 15131 | 10 | begin | 
| 11363 | 11 | |
| 38159 | 12 | definition coprime :: "nat => nat => bool" | 
| 13 | where "coprime m n \<longleftrightarrow> gcd m n = 1" | |
| 11363 | 14 | |
| 38159 | 15 | definition prime :: "nat \<Rightarrow> bool" | 
| 16 | where "prime p \<longleftrightarrow> (1 < p \<and> (\<forall>m. m dvd p --> m = 1 \<or> m = p))" | |
| 11363 | 17 | |
| 18 | ||
| 16762 | 19 | lemma two_is_prime: "prime 2" | 
| 20 | apply (auto simp add: prime_def) | |
| 21 | apply (case_tac m) | |
| 22 | apply (auto dest!: dvd_imp_le) | |
| 11363 | 23 | done | 
| 24 | ||
| 27556 | 25 | lemma prime_imp_relprime: "prime p ==> \<not> p dvd n ==> gcd p n = 1" | 
| 11363 | 26 | apply (auto simp add: prime_def) | 
| 50037 | 27 | apply (metis gcd_dvd1 gcd_dvd2) | 
| 11363 | 28 | done | 
| 29 | ||
| 30 | text {*
 | |
| 31 | This theorem leads immediately to a proof of the uniqueness of | |
| 32 |   factorization.  If @{term p} divides a product of primes then it is
 | |
| 33 | one of those primes. | |
| 34 | *} | |
| 35 | ||
| 16663 | 36 | lemma prime_dvd_mult: "prime p ==> p dvd m * n ==> p dvd m \<or> p dvd n" | 
| 12739 | 37 | by (blast intro: relprime_dvd_mult prime_imp_relprime) | 
| 11363 | 38 | |
| 16663 | 39 | lemma prime_dvd_square: "prime p ==> p dvd m^Suc (Suc 0) ==> p dvd m" | 
| 12739 | 40 | by (auto dest: prime_dvd_mult) | 
| 41 | ||
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50037diff
changeset | 42 | lemma prime_dvd_power_two: "prime p ==> p dvd m\<^sup>2 ==> p dvd m" | 
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
13187diff
changeset | 43 | by (rule prime_dvd_square) (simp_all add: power2_eq_square) | 
| 11368 | 44 | |
| 26125 | 45 | |
| 30056 | 46 | lemma exp_eq_1:"(x::nat)^n = 1 \<longleftrightarrow> x = 1 \<or> n = 0" | 
| 47 | by (induct n, auto) | |
| 48 | ||
| 26125 | 49 | lemma exp_mono_lt: "(x::nat) ^ (Suc n) < y ^ (Suc n) \<longleftrightarrow> x < y" | 
| 30056 | 50 | by(metis linorder_not_less not_less0 power_le_imp_le_base power_less_imp_less_base) | 
| 51 | ||
| 26125 | 52 | lemma exp_mono_le: "(x::nat) ^ (Suc n) \<le> y ^ (Suc n) \<longleftrightarrow> x \<le> y" | 
| 30056 | 53 | by (simp only: linorder_not_less[symmetric] exp_mono_lt) | 
| 26125 | 54 | |
| 55 | lemma exp_mono_eq: "(x::nat) ^ Suc n = y ^ Suc n \<longleftrightarrow> x = y" | |
| 56 | using power_inject_base[of x n y] by auto | |
| 57 | ||
| 58 | ||
| 53077 | 59 | lemma even_square: assumes e: "even (n::nat)" shows "\<exists>x. n\<^sup>2 = 4*x" | 
| 26125 | 60 | proof- | 
| 61 | from e have "2 dvd n" by presburger | |
| 62 | then obtain k where k: "n = 2*k" using dvd_def by auto | |
| 53077 | 63 | hence "n\<^sup>2 = 4 * k\<^sup>2" by (simp add: power2_eq_square) | 
| 26125 | 64 | thus ?thesis by blast | 
| 65 | qed | |
| 66 | ||
| 53077 | 67 | lemma odd_square: assumes e: "odd (n::nat)" shows "\<exists>x. n\<^sup>2 = 4*x + 1" | 
| 26125 | 68 | proof- | 
| 69 | from e have np: "n > 0" by presburger | |
| 70 | from e have "2 dvd (n - 1)" by presburger | |
| 71 | then obtain k where "n - 1 = 2*k" using dvd_def by auto | |
| 72 | hence k: "n = 2*k + 1" using e by presburger | |
| 53077 | 73 | hence "n\<^sup>2 = 4* (k\<^sup>2 + k) + 1" by algebra | 
| 26125 | 74 | thus ?thesis by blast | 
| 75 | qed | |
| 76 | ||
| 53077 | 77 | lemma diff_square: "(x::nat)\<^sup>2 - y\<^sup>2 = (x+y)*(x - y)" | 
| 26125 | 78 | proof- | 
| 79 | have "x \<le> y \<or> y \<le> x" by (rule nat_le_linear) | |
| 80 | moreover | |
| 81 |   {assume le: "x \<le> y"
 | |
| 53077 | 82 | hence "x\<^sup>2 \<le> y\<^sup>2" by (simp only: numeral_2_eq_2 exp_mono_le Let_def) | 
| 26125 | 83 | with le have ?thesis by simp } | 
| 84 | moreover | |
| 85 |   {assume le: "y \<le> x"
 | |
| 53077 | 86 | hence le2: "y\<^sup>2 \<le> x\<^sup>2" by (simp only: numeral_2_eq_2 exp_mono_le Let_def) | 
| 26125 | 87 | from le have "\<exists>z. y + z = x" by presburger | 
| 88 | then obtain z where z: "x = y + z" by blast | |
| 53077 | 89 | from le2 have "\<exists>z. x\<^sup>2 = y\<^sup>2 + z" by presburger | 
| 90 | then obtain z2 where z2: "x\<^sup>2 = y\<^sup>2 + z2" by blast | |
| 91 | from z z2 have ?thesis by simp algebra } | |
| 26125 | 92 | ultimately show ?thesis by blast | 
| 93 | qed | |
| 94 | ||
| 26144 | 95 | text {* Elementary theory of divisibility *}
 | 
| 26125 | 96 | lemma divides_ge: "(a::nat) dvd b \<Longrightarrow> b = 0 \<or> a \<le> b" unfolding dvd_def by auto | 
| 97 | lemma divides_antisym: "(x::nat) dvd y \<and> y dvd x \<longleftrightarrow> x = y" | |
| 33657 | 98 | using dvd_antisym[of x y] by auto | 
| 26125 | 99 | |
| 100 | lemma divides_add_revr: assumes da: "(d::nat) dvd a" and dab:"d dvd (a + b)" | |
| 101 | shows "d dvd b" | |
| 102 | proof- | |
| 103 | from da obtain k where k:"a = d*k" by (auto simp add: dvd_def) | |
| 104 | from dab obtain k' where k': "a + b = d*k'" by (auto simp add: dvd_def) | |
| 105 | from k k' have "b = d *(k' - k)" by (simp add : diff_mult_distrib2) | |
| 106 | thus ?thesis unfolding dvd_def by blast | |
| 107 | qed | |
| 108 | ||
| 109 | declare nat_mult_dvd_cancel_disj[presburger] | |
| 110 | lemma nat_mult_dvd_cancel_disj'[presburger]: | |
| 111 | "(m\<Colon>nat)*k dvd n*k \<longleftrightarrow> k = 0 \<or> m dvd n" unfolding mult_commute[of m k] mult_commute[of n k] by presburger | |
| 112 | ||
| 113 | lemma divides_mul_l: "(a::nat) dvd b ==> (c * a) dvd (c * b)" | |
| 114 | by presburger | |
| 115 | ||
| 116 | lemma divides_mul_r: "(a::nat) dvd b ==> (a * c) dvd (b * c)" by presburger | |
| 117 | lemma divides_cases: "(n::nat) dvd m ==> m = 0 \<or> m = n \<or> 2 * n <= m" | |
| 118 | by (auto simp add: dvd_def) | |
| 119 | ||
| 120 | lemma divides_div_not: "(x::nat) = (q * n) + r \<Longrightarrow> 0 < r \<Longrightarrow> r < n ==> ~(n dvd x)" | |
| 121 | proof(auto simp add: dvd_def) | |
| 122 | fix k assume H: "0 < r" "r < n" "q * n + r = n * k" | |
| 123 | from H(3) have r: "r = n* (k -q)" by(simp add: diff_mult_distrib2 mult_commute) | |
| 124 |   {assume "k - q = 0" with r H(1) have False by simp}
 | |
| 125 | moreover | |
| 126 |   {assume "k - q \<noteq> 0" with r have "r \<ge> n" by auto
 | |
| 127 | with H(2) have False by simp} | |
| 128 | ultimately show False by blast | |
| 129 | qed | |
| 130 | lemma divides_exp: "(x::nat) dvd y ==> x ^ n dvd y ^ n" | |
| 131 | by (auto simp add: power_mult_distrib dvd_def) | |
| 132 | ||
| 133 | lemma divides_exp2: "n \<noteq> 0 \<Longrightarrow> (x::nat) ^ n dvd y \<Longrightarrow> x dvd y" | |
| 134 | by (induct n ,auto simp add: dvd_def) | |
| 135 | ||
| 136 | fun fact :: "nat \<Rightarrow> nat" where | |
| 137 | "fact 0 = 1" | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 138 | | "fact (Suc n) = Suc n * fact n" | 
| 26125 | 139 | |
| 140 | lemma fact_lt: "0 < fact n" by(induct n, simp_all) | |
| 141 | lemma fact_le: "fact n \<ge> 1" using fact_lt[of n] by simp | |
| 142 | lemma fact_mono: assumes le: "m \<le> n" shows "fact m \<le> fact n" | |
| 143 | proof- | |
| 144 | from le have "\<exists>i. n = m+i" by presburger | |
| 145 | then obtain i where i: "n = m+i" by blast | |
| 146 | have "fact m \<le> fact (m + i)" | |
| 147 | proof(induct m) | |
| 148 | case 0 thus ?case using fact_le[of i] by simp | |
| 149 | next | |
| 150 | case (Suc m) | |
| 151 | have "fact (Suc m) = Suc m * fact m" by simp | |
| 152 | have th1: "Suc m \<le> Suc (m + i)" by simp | |
| 153 | from mult_le_mono[of "Suc m" "Suc (m+i)" "fact m" "fact (m+i)", OF th1 Suc.hyps] | |
| 154 | show ?case by simp | |
| 155 | qed | |
| 156 | thus ?thesis using i by simp | |
| 157 | qed | |
| 158 | ||
| 159 | lemma divides_fact: "1 <= p \<Longrightarrow> p <= n ==> p dvd fact n" | |
| 160 | proof(induct n arbitrary: p) | |
| 161 | case 0 thus ?case by simp | |
| 162 | next | |
| 163 | case (Suc n p) | |
| 164 | from Suc.prems have "p = Suc n \<or> p \<le> n" by presburger | |
| 165 | moreover | |
| 166 |   {assume "p = Suc n" hence ?case  by (simp only: fact.simps dvd_triv_left)}
 | |
| 167 | moreover | |
| 168 |   {assume "p \<le> n"
 | |
| 169 | with Suc.prems(1) Suc.hyps have th: "p dvd fact n" by simp | |
| 170 | from dvd_mult[OF th] have ?case by (simp only: fact.simps) } | |
| 171 | ultimately show ?case by blast | |
| 172 | qed | |
| 173 | ||
| 174 | declare dvd_triv_left[presburger] | |
| 175 | declare dvd_triv_right[presburger] | |
| 176 | lemma divides_rexp: | |
| 177 | "x dvd y \<Longrightarrow> (x::nat) dvd (y^(Suc n))" by (simp add: dvd_mult2[of x y]) | |
| 178 | ||
| 26144 | 179 | text {* Coprimality *}
 | 
| 26125 | 180 | |
| 181 | lemma coprime: "coprime a b \<longleftrightarrow> (\<forall>d. d dvd a \<and> d dvd b \<longleftrightarrow> d = 1)" | |
| 182 | using gcd_unique[of 1 a b, simplified] by (auto simp add: coprime_def) | |
| 183 | lemma coprime_commute: "coprime a b \<longleftrightarrow> coprime b a" by (simp add: coprime_def gcd_commute) | |
| 184 | ||
| 185 | lemma coprime_bezout: "coprime a b \<longleftrightarrow> (\<exists>x y. a * x - b * y = 1 \<or> b * x - a * y = 1)" | |
| 186 | using coprime_def gcd_bezout by auto | |
| 187 | ||
| 188 | lemma coprime_divprod: "d dvd a * b \<Longrightarrow> coprime d a \<Longrightarrow> d dvd b" | |
| 189 | using relprime_dvd_mult_iff[of d a b] by (auto simp add: coprime_def mult_commute) | |
| 190 | ||
| 191 | lemma coprime_1[simp]: "coprime a 1" by (simp add: coprime_def) | |
| 192 | lemma coprime_1'[simp]: "coprime 1 a" by (simp add: coprime_def) | |
| 193 | lemma coprime_Suc0[simp]: "coprime a (Suc 0)" by (simp add: coprime_def) | |
| 194 | lemma coprime_Suc0'[simp]: "coprime (Suc 0) a" by (simp add: coprime_def) | |
| 195 | ||
| 196 | lemma gcd_coprime: | |
| 27556 | 197 | assumes z: "gcd a b \<noteq> 0" and a: "a = a' * gcd a b" and b: "b = b' * gcd a b" | 
| 26125 | 198 | shows "coprime a' b'" | 
| 199 | proof- | |
| 27556 | 200 | let ?g = "gcd a b" | 
| 26125 | 201 |   {assume bz: "a = 0" from b bz z a have ?thesis by (simp add: gcd_zero coprime_def)}
 | 
| 202 | moreover | |
| 203 |   {assume az: "a\<noteq> 0" 
 | |
| 204 | from z have z': "?g > 0" by simp | |
| 205 | from bezout_gcd_strong[OF az, of b] | |
| 206 | obtain x y where xy: "a*x = b*y + ?g" by blast | |
| 29667 | 207 | from xy a b have "?g * a'*x = ?g * (b'*y + 1)" by (simp add: algebra_simps) | 
| 26125 | 208 | hence "?g * (a'*x) = ?g * (b'*y + 1)" by (simp add: mult_assoc) | 
| 209 | hence "a'*x = (b'*y + 1)" | |
| 210 | by (simp only: nat_mult_eq_cancel1[OF z']) | |
| 211 | hence "a'*x - b'*y = 1" by simp | |
| 212 | with coprime_bezout[of a' b'] have ?thesis by auto} | |
| 213 | ultimately show ?thesis by blast | |
| 214 | qed | |
| 215 | lemma coprime_0: "coprime d 0 \<longleftrightarrow> d = 1" by (simp add: coprime_def) | |
| 216 | lemma coprime_mul: assumes da: "coprime d a" and db: "coprime d b" | |
| 217 | shows "coprime d (a * b)" | |
| 218 | proof- | |
| 27556 | 219 | from da have th: "gcd a d = 1" by (simp add: coprime_def gcd_commute) | 
| 27567 | 220 | from gcd_mult_cancel[of a d b, OF th] db[unfolded coprime_def] have "gcd d (a*b) = 1" | 
| 26125 | 221 | by (simp add: gcd_commute) | 
| 222 | thus ?thesis unfolding coprime_def . | |
| 223 | qed | |
| 224 | lemma coprime_lmul2: assumes dab: "coprime d (a * b)" shows "coprime d b" | |
| 41541 | 225 | using dab unfolding coprime_bezout | 
| 26125 | 226 | apply clarsimp | 
| 227 | apply (case_tac "d * x - a * b * y = Suc 0 ", simp_all) | |
| 228 | apply (rule_tac x="x" in exI) | |
| 229 | apply (rule_tac x="a*y" in exI) | |
| 230 | apply (simp add: mult_ac) | |
| 231 | apply (rule_tac x="a*x" in exI) | |
| 232 | apply (rule_tac x="y" in exI) | |
| 233 | apply (simp add: mult_ac) | |
| 234 | done | |
| 235 | ||
| 236 | lemma coprime_rmul2: "coprime d (a * b) \<Longrightarrow> coprime d a" | |
| 237 | unfolding coprime_bezout | |
| 238 | apply clarsimp | |
| 239 | apply (case_tac "d * x - a * b * y = Suc 0 ", simp_all) | |
| 240 | apply (rule_tac x="x" in exI) | |
| 241 | apply (rule_tac x="b*y" in exI) | |
| 242 | apply (simp add: mult_ac) | |
| 243 | apply (rule_tac x="b*x" in exI) | |
| 244 | apply (rule_tac x="y" in exI) | |
| 245 | apply (simp add: mult_ac) | |
| 246 | done | |
| 247 | lemma coprime_mul_eq: "coprime d (a * b) \<longleftrightarrow> coprime d a \<and> coprime d b" | |
| 248 | using coprime_rmul2[of d a b] coprime_lmul2[of d a b] coprime_mul[of d a b] | |
| 249 | by blast | |
| 250 | ||
| 251 | lemma gcd_coprime_exists: | |
| 27556 | 252 | assumes nz: "gcd a b \<noteq> 0" | 
| 253 | shows "\<exists>a' b'. a = a' * gcd a b \<and> b = b' * gcd a b \<and> coprime a' b'" | |
| 26125 | 254 | proof- | 
| 27556 | 255 | let ?g = "gcd a b" | 
| 26125 | 256 | from gcd_dvd1[of a b] gcd_dvd2[of a b] | 
| 257 | obtain a' b' where "a = ?g*a'" "b = ?g*b'" unfolding dvd_def by blast | |
| 258 | hence ab': "a = a'*?g" "b = b'*?g" by algebra+ | |
| 259 | from ab' gcd_coprime[OF nz ab'] show ?thesis by blast | |
| 260 | qed | |
| 261 | ||
| 262 | lemma coprime_exp: "coprime d a ==> coprime d (a^n)" | |
| 263 | by(induct n, simp_all add: coprime_mul) | |
| 264 | ||
| 265 | lemma coprime_exp_imp: "coprime a b ==> coprime (a ^n) (b ^n)" | |
| 266 | by (induct n, simp_all add: coprime_mul_eq coprime_commute coprime_exp) | |
| 267 | lemma coprime_refl[simp]: "coprime n n \<longleftrightarrow> n = 1" by (simp add: coprime_def) | |
| 268 | lemma coprime_plus1[simp]: "coprime (n + 1) n" | |
| 269 | apply (simp add: coprime_bezout) | |
| 270 | apply (rule exI[where x=1]) | |
| 271 | apply (rule exI[where x=1]) | |
| 272 | apply simp | |
| 273 | done | |
| 274 | lemma coprime_minus1: "n \<noteq> 0 ==> coprime (n - 1) n" | |
| 275 | using coprime_plus1[of "n - 1"] coprime_commute[of "n - 1" n] by auto | |
| 276 | ||
| 27556 | 277 | lemma bezout_gcd_pow: "\<exists>x y. a ^n * x - b ^ n * y = gcd a b ^ n \<or> b ^ n * x - a ^ n * y = gcd a b ^ n" | 
| 26125 | 278 | proof- | 
| 27556 | 279 | let ?g = "gcd a b" | 
| 26125 | 280 |   {assume z: "?g = 0" hence ?thesis 
 | 
| 281 | apply (cases n, simp) | |
| 282 | apply arith | |
| 283 | apply (simp only: z power_0_Suc) | |
| 284 | apply (rule exI[where x=0]) | |
| 285 | apply (rule exI[where x=0]) | |
| 41541 | 286 | apply simp | 
| 287 | done } | |
| 26125 | 288 | moreover | 
| 289 |   {assume z: "?g \<noteq> 0"
 | |
| 290 | from gcd_dvd1[of a b] gcd_dvd2[of a b] obtain a' b' where | |
| 291 | ab': "a = a'*?g" "b = b'*?g" unfolding dvd_def by (auto simp add: mult_ac) | |
| 292 | hence ab'': "?g*a' = a" "?g * b' = b" by algebra+ | |
| 293 | from coprime_exp_imp[OF gcd_coprime[OF z ab'], unfolded coprime_bezout, of n] | |
| 294 | obtain x y where "a'^n * x - b'^n * y = 1 \<or> b'^n * x - a'^n * y = 1" by blast | |
| 295 | hence "?g^n * (a'^n * x - b'^n * y) = ?g^n \<or> ?g^n*(b'^n * x - a'^n * y) = ?g^n" | |
| 296 | using z by auto | |
| 297 | then have "a^n * x - b^n * y = ?g^n \<or> b^n * x - a^n * y = ?g^n" | |
| 298 | using z ab'' by (simp only: power_mult_distrib[symmetric] | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 299 | diff_mult_distrib2 mult_assoc[symmetric]) | 
| 26125 | 300 | hence ?thesis by blast } | 
| 301 | ultimately show ?thesis by blast | |
| 302 | qed | |
| 27567 | 303 | |
| 304 | lemma gcd_exp: "gcd (a^n) (b^n) = gcd a b^n" | |
| 26125 | 305 | proof- | 
| 27556 | 306 | let ?g = "gcd (a^n) (b^n)" | 
| 27567 | 307 | let ?gn = "gcd a b^n" | 
| 26125 | 308 |   {fix e assume H: "e dvd a^n" "e dvd b^n"
 | 
| 309 | from bezout_gcd_pow[of a n b] obtain x y | |
| 310 | where xy: "a ^ n * x - b ^ n * y = ?gn \<or> b ^ n * x - a ^ n * y = ?gn" by blast | |
| 31952 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 nipkow parents: 
31706diff
changeset | 311 | from dvd_diff_nat [OF dvd_mult2[OF H(1), of x] dvd_mult2[OF H(2), of y]] | 
| 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 nipkow parents: 
31706diff
changeset | 312 | dvd_diff_nat [OF dvd_mult2[OF H(2), of x] dvd_mult2[OF H(1), of y]] xy | 
| 27556 | 313 | have "e dvd ?gn" by (cases "a ^ n * x - b ^ n * y = gcd a b ^ n", simp_all)} | 
| 26125 | 314 | hence th: "\<forall>e. e dvd a^n \<and> e dvd b^n \<longrightarrow> e dvd ?gn" by blast | 
| 315 | from divides_exp[OF gcd_dvd1[of a b], of n] divides_exp[OF gcd_dvd2[of a b], of n] th | |
| 316 | gcd_unique have "?gn = ?g" by blast thus ?thesis by simp | |
| 317 | qed | |
| 318 | ||
| 319 | lemma coprime_exp2: "coprime (a ^ Suc n) (b^ Suc n) \<longleftrightarrow> coprime a b" | |
| 320 | by (simp only: coprime_def gcd_exp exp_eq_1) simp | |
| 321 | ||
| 322 | lemma division_decomp: assumes dc: "(a::nat) dvd b * c" | |
| 323 | shows "\<exists>b' c'. a = b' * c' \<and> b' dvd b \<and> c' dvd c" | |
| 324 | proof- | |
| 27556 | 325 | let ?g = "gcd a b" | 
| 26125 | 326 |   {assume "?g = 0" with dc have ?thesis apply (simp add: gcd_zero)
 | 
| 327 | apply (rule exI[where x="0"]) | |
| 328 | by (rule exI[where x="c"], simp)} | |
| 329 | moreover | |
| 330 |   {assume z: "?g \<noteq> 0"
 | |
| 331 | from gcd_coprime_exists[OF z] | |
| 332 | obtain a' b' where ab': "a = a' * ?g" "b = b' * ?g" "coprime a' b'" by blast | |
| 333 | from gcd_dvd2[of a b] have thb: "?g dvd b" . | |
| 334 | from ab'(1) have "a' dvd a" unfolding dvd_def by blast | |
| 335 | with dc have th0: "a' dvd b*c" using dvd_trans[of a' a "b*c"] by simp | |
| 336 | from dc ab'(1,2) have "a'*?g dvd (b'*?g) *c" by auto | |
| 337 | hence "?g*a' dvd ?g * (b' * c)" by (simp add: mult_assoc) | |
| 338 | with z have th_1: "a' dvd b'*c" by simp | |
| 339 | from coprime_divprod[OF th_1 ab'(3)] have thc: "a' dvd c" . | |
| 340 | from ab' have "a = ?g*a'" by algebra | |
| 341 | with thb thc have ?thesis by blast } | |
| 342 | ultimately show ?thesis by blast | |
| 343 | qed | |
| 344 | ||
| 345 | lemma nat_power_eq_0_iff: "(m::nat) ^ n = 0 \<longleftrightarrow> n \<noteq> 0 \<and> m = 0" by (induct n, auto) | |
| 346 | ||
| 347 | lemma divides_rev: assumes ab: "(a::nat) ^ n dvd b ^n" and n:"n \<noteq> 0" shows "a dvd b" | |
| 348 | proof- | |
| 27556 | 349 | let ?g = "gcd a b" | 
| 26125 | 350 | from n obtain m where m: "n = Suc m" by (cases n, simp_all) | 
| 351 |   {assume "?g = 0" with ab n have ?thesis by (simp add: gcd_zero)}
 | |
| 352 | moreover | |
| 353 |   {assume z: "?g \<noteq> 0"
 | |
| 41541 | 354 | hence zn: "?g ^ n \<noteq> 0" using n by simp | 
| 26125 | 355 | from gcd_coprime_exists[OF z] | 
| 356 | obtain a' b' where ab': "a = a' * ?g" "b = b' * ?g" "coprime a' b'" by blast | |
| 357 | from ab have "(a' * ?g) ^ n dvd (b' * ?g)^n" by (simp add: ab'(1,2)[symmetric]) | |
| 358 | hence "?g^n*a'^n dvd ?g^n *b'^n" by (simp only: power_mult_distrib mult_commute) | |
| 359 | with zn z n have th0:"a'^n dvd b'^n" by (auto simp add: nat_power_eq_0_iff) | |
| 360 | have "a' dvd a'^n" by (simp add: m) | |
| 361 | with th0 have "a' dvd b'^n" using dvd_trans[of a' "a'^n" "b'^n"] by simp | |
| 362 | hence th1: "a' dvd b'^m * b'" by (simp add: m mult_commute) | |
| 363 | from coprime_divprod[OF th1 coprime_exp[OF ab'(3), of m]] | |
| 364 | have "a' dvd b'" . | |
| 365 | hence "a'*?g dvd b'*?g" by simp | |
| 366 | with ab'(1,2) have ?thesis by simp } | |
| 367 | ultimately show ?thesis by blast | |
| 368 | qed | |
| 369 | ||
| 370 | lemma divides_mul: assumes mr: "m dvd r" and nr: "n dvd r" and mn:"coprime m n" | |
| 371 | shows "m * n dvd r" | |
| 372 | proof- | |
| 373 | from mr nr obtain m' n' where m': "r = m*m'" and n': "r = n*n'" | |
| 374 | unfolding dvd_def by blast | |
| 375 | from mr n' have "m dvd n'*n" by (simp add: mult_commute) | |
| 376 | hence "m dvd n'" using relprime_dvd_mult_iff[OF mn[unfolded coprime_def]] by simp | |
| 377 | then obtain k where k: "n' = m*k" unfolding dvd_def by blast | |
| 378 | from n' k show ?thesis unfolding dvd_def by auto | |
| 379 | qed | |
| 380 | ||
| 26144 | 381 | |
| 382 | text {* A binary form of the Chinese Remainder Theorem. *}
 | |
| 26125 | 383 | |
| 384 | lemma chinese_remainder: assumes ab: "coprime a b" and a:"a \<noteq> 0" and b:"b \<noteq> 0" | |
| 385 | shows "\<exists>x q1 q2. x = u + q1 * a \<and> x = v + q2 * b" | |
| 386 | proof- | |
| 387 | from bezout_add_strong[OF a, of b] bezout_add_strong[OF b, of a] | |
| 388 | obtain d1 x1 y1 d2 x2 y2 where dxy1: "d1 dvd a" "d1 dvd b" "a * x1 = b * y1 + d1" | |
| 389 | and dxy2: "d2 dvd b" "d2 dvd a" "b * x2 = a * y2 + d2" by blast | |
| 390 | from gcd_unique[of 1 a b, simplified ab[unfolded coprime_def], simplified] | |
| 391 | dxy1(1,2) dxy2(1,2) have d12: "d1 = 1" "d2 =1" by auto | |
| 392 | let ?x = "v * a * x1 + u * b * x2" | |
| 393 | let ?q1 = "v * x1 + u * y2" | |
| 394 | let ?q2 = "v * y1 + u * x2" | |
| 395 | from dxy2(3)[simplified d12] dxy1(3)[simplified d12] | |
| 396 | have "?x = u + ?q1 * a" "?x = v + ?q2 * b" by algebra+ | |
| 397 | thus ?thesis by blast | |
| 398 | qed | |
| 399 | ||
| 26144 | 400 | text {* Primality *}
 | 
| 401 | ||
| 402 | text {* A few useful theorems about primes *}
 | |
| 26125 | 403 | |
| 404 | lemma prime_0[simp]: "~prime 0" by (simp add: prime_def) | |
| 405 | lemma prime_1[simp]: "~ prime 1" by (simp add: prime_def) | |
| 406 | lemma prime_Suc0[simp]: "~ prime (Suc 0)" by (simp add: prime_def) | |
| 407 | ||
| 408 | lemma prime_ge_2: "prime p ==> p \<ge> 2" by (simp add: prime_def) | |
| 409 | lemma prime_factor: assumes n: "n \<noteq> 1" shows "\<exists> p. prime p \<and> p dvd n" | |
| 410 | using n | |
| 411 | proof(induct n rule: nat_less_induct) | |
| 412 | fix n | |
| 413 | assume H: "\<forall>m<n. m \<noteq> 1 \<longrightarrow> (\<exists>p. prime p \<and> p dvd m)" "n \<noteq> 1" | |
| 414 | let ?ths = "\<exists>p. prime p \<and> p dvd n" | |
| 415 |   {assume "n=0" hence ?ths using two_is_prime by auto}
 | |
| 416 | moreover | |
| 417 |   {assume nz: "n\<noteq>0" 
 | |
| 418 |     {assume "prime n" hence ?ths by - (rule exI[where x="n"], simp)}
 | |
| 419 | moreover | |
| 420 |     {assume n: "\<not> prime n"
 | |
| 421 | with nz H(2) | |
| 422 | obtain k where k:"k dvd n" "k \<noteq> 1" "k \<noteq> n" by (auto simp add: prime_def) | |
| 423 | from dvd_imp_le[OF k(1)] nz k(3) have kn: "k < n" by simp | |
| 424 | from H(1)[rule_format, OF kn k(2)] obtain p where p: "prime p" "p dvd k" by blast | |
| 425 | from dvd_trans[OF p(2) k(1)] p(1) have ?ths by blast} | |
| 426 | ultimately have ?ths by blast} | |
| 427 | ultimately show ?ths by blast | |
| 428 | qed | |
| 429 | ||
| 430 | lemma prime_factor_lt: assumes p: "prime p" and n: "n \<noteq> 0" and npm:"n = p * m" | |
| 431 | shows "m < n" | |
| 432 | proof- | |
| 433 |   {assume "m=0" with n have ?thesis by simp}
 | |
| 434 | moreover | |
| 435 |   {assume m: "m \<noteq> 0"
 | |
| 436 | from npm have mn: "m dvd n" unfolding dvd_def by auto | |
| 437 | from npm m have "n \<noteq> m" using p by auto | |
| 438 | with dvd_imp_le[OF mn] n have ?thesis by simp} | |
| 439 | ultimately show ?thesis by blast | |
| 440 | qed | |
| 441 | ||
| 442 | lemma euclid_bound: "\<exists>p. prime p \<and> n < p \<and> p <= Suc (fact n)" | |
| 443 | proof- | |
| 444 | have f1: "fact n + 1 \<noteq> 1" using fact_le[of n] by arith | |
| 445 | from prime_factor[OF f1] obtain p where p: "prime p" "p dvd fact n + 1" by blast | |
| 446 | from dvd_imp_le[OF p(2)] have pfn: "p \<le> fact n + 1" by simp | |
| 447 |   {assume np: "p \<le> n"
 | |
| 448 | from p(1) have p1: "p \<ge> 1" by (cases p, simp_all) | |
| 449 | from divides_fact[OF p1 np] have pfn': "p dvd fact n" . | |
| 450 | from divides_add_revr[OF pfn' p(2)] p(1) have False by simp} | |
| 451 | hence "n < p" by arith | |
| 452 | with p(1) pfn show ?thesis by auto | |
| 453 | qed | |
| 454 | ||
| 455 | lemma euclid: "\<exists>p. prime p \<and> p > n" using euclid_bound by auto | |
| 31044 | 456 | |
| 26125 | 457 | lemma primes_infinite: "\<not> (finite {p. prime p})"
 | 
| 31044 | 458 | apply(simp add: finite_nat_set_iff_bounded_le) | 
| 459 | apply (metis euclid linorder_not_le) | |
| 460 | done | |
| 26125 | 461 | |
| 462 | lemma coprime_prime: assumes ab: "coprime a b" | |
| 463 | shows "~(prime p \<and> p dvd a \<and> p dvd b)" | |
| 464 | proof | |
| 465 | assume "prime p \<and> p dvd a \<and> p dvd b" | |
| 466 | thus False using ab gcd_greatest[of p a b] by (simp add: coprime_def) | |
| 467 | qed | |
| 468 | lemma coprime_prime_eq: "coprime a b \<longleftrightarrow> (\<forall>p. ~(prime p \<and> p dvd a \<and> p dvd b))" | |
| 469 | (is "?lhs = ?rhs") | |
| 470 | proof- | |
| 471 |   {assume "?lhs" with coprime_prime  have ?rhs by blast}
 | |
| 472 | moreover | |
| 473 |   {assume r: "?rhs" and c: "\<not> ?lhs"
 | |
| 474 | then obtain g where g: "g\<noteq>1" "g dvd a" "g dvd b" unfolding coprime_def by blast | |
| 475 | from prime_factor[OF g(1)] obtain p where p: "prime p" "p dvd g" by blast | |
| 476 | from dvd_trans [OF p(2) g(2)] dvd_trans [OF p(2) g(3)] | |
| 477 | have "p dvd a" "p dvd b" . with p(1) r have False by blast} | |
| 478 | ultimately show ?thesis by blast | |
| 479 | qed | |
| 480 | ||
| 481 | lemma prime_coprime: assumes p: "prime p" | |
| 482 | shows "n = 1 \<or> p dvd n \<or> coprime p n" | |
| 483 | using p prime_imp_relprime[of p n] by (auto simp add: coprime_def) | |
| 484 | ||
| 485 | lemma prime_coprime_strong: "prime p \<Longrightarrow> p dvd n \<or> coprime p n" | |
| 486 | using prime_coprime[of p n] by auto | |
| 487 | ||
| 488 | declare coprime_0[simp] | |
| 489 | ||
| 490 | lemma coprime_0'[simp]: "coprime 0 d \<longleftrightarrow> d = 1" by (simp add: coprime_commute[of 0 d]) | |
| 491 | lemma coprime_bezout_strong: assumes ab: "coprime a b" and b: "b \<noteq> 1" | |
| 492 | shows "\<exists>x y. a * x = b * y + 1" | |
| 493 | proof- | |
| 494 | from ab b have az: "a \<noteq> 0" by - (rule ccontr, auto) | |
| 495 | from bezout_gcd_strong[OF az, of b] ab[unfolded coprime_def] | |
| 496 | show ?thesis by auto | |
| 497 | qed | |
| 498 | ||
| 499 | lemma bezout_prime: assumes p: "prime p" and pa: "\<not> p dvd a" | |
| 500 | shows "\<exists>x y. a*x = p*y + 1" | |
| 501 | proof- | |
| 502 | from p have p1: "p \<noteq> 1" using prime_1 by blast | |
| 503 | from prime_coprime[OF p, of a] p1 pa have ap: "coprime a p" | |
| 504 | by (auto simp add: coprime_commute) | |
| 505 | from coprime_bezout_strong[OF ap p1] show ?thesis . | |
| 506 | qed | |
| 507 | lemma prime_divprod: assumes p: "prime p" and pab: "p dvd a*b" | |
| 508 | shows "p dvd a \<or> p dvd b" | |
| 509 | proof- | |
| 510 |   {assume "a=1" hence ?thesis using pab by simp }
 | |
| 511 | moreover | |
| 512 |   {assume "p dvd a" hence ?thesis by blast}
 | |
| 513 | moreover | |
| 514 |   {assume pa: "coprime p a" from coprime_divprod[OF pab pa]  have ?thesis .. }
 | |
| 515 | ultimately show ?thesis using prime_coprime[OF p, of a] by blast | |
| 516 | qed | |
| 517 | ||
| 518 | lemma prime_divprod_eq: assumes p: "prime p" | |
| 519 | shows "p dvd a*b \<longleftrightarrow> p dvd a \<or> p dvd b" | |
| 520 | using p prime_divprod dvd_mult dvd_mult2 by auto | |
| 521 | ||
| 522 | lemma prime_divexp: assumes p:"prime p" and px: "p dvd x^n" | |
| 523 | shows "p dvd x" | |
| 524 | using px | |
| 525 | proof(induct n) | |
| 526 | case 0 thus ?case by simp | |
| 527 | next | |
| 528 | case (Suc n) | |
| 529 | hence th: "p dvd x*x^n" by simp | |
| 530 |   {assume H: "p dvd x^n"
 | |
| 531 | from Suc.hyps[OF H] have ?case .} | |
| 532 | with prime_divprod[OF p th] show ?case by blast | |
| 533 | qed | |
| 534 | ||
| 535 | lemma prime_divexp_n: "prime p \<Longrightarrow> p dvd x^n \<Longrightarrow> p^n dvd x^n" | |
| 536 | using prime_divexp[of p x n] divides_exp[of p x n] by blast | |
| 537 | ||
| 538 | lemma coprime_prime_dvd_ex: assumes xy: "\<not>coprime x y" | |
| 539 | shows "\<exists>p. prime p \<and> p dvd x \<and> p dvd y" | |
| 540 | proof- | |
| 541 | from xy[unfolded coprime_def] obtain g where g: "g \<noteq> 1" "g dvd x" "g dvd y" | |
| 542 | by blast | |
| 543 | from prime_factor[OF g(1)] obtain p where p: "prime p" "p dvd g" by blast | |
| 544 | from g(2,3) dvd_trans[OF p(2)] p(1) show ?thesis by auto | |
| 545 | qed | |
| 546 | lemma coprime_sos: assumes xy: "coprime x y" | |
| 53077 | 547 | shows "coprime (x * y) (x\<^sup>2 + y\<^sup>2)" | 
| 26125 | 548 | proof- | 
| 53077 | 549 |   {assume c: "\<not> coprime (x * y) (x\<^sup>2 + y\<^sup>2)"
 | 
| 26125 | 550 | from coprime_prime_dvd_ex[OF c] obtain p | 
| 53077 | 551 | where p: "prime p" "p dvd x*y" "p dvd x\<^sup>2 + y\<^sup>2" by blast | 
| 26125 | 552 |     {assume px: "p dvd x"
 | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27567diff
changeset | 553 | from dvd_mult[OF px, of x] p(3) | 
| 53077 | 554 | obtain r s where "x * x = p * r" and "x\<^sup>2 + y\<^sup>2 = p * s" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27567diff
changeset | 555 | by (auto elim!: dvdE) | 
| 53077 | 556 | then have "y\<^sup>2 = p * (s - r)" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27567diff
changeset | 557 | by (auto simp add: power2_eq_square diff_mult_distrib2) | 
| 53077 | 558 | then have "p dvd y\<^sup>2" .. | 
| 26125 | 559 | with prime_divexp[OF p(1), of y 2] have py: "p dvd y" . | 
| 560 | from p(1) px py xy[unfolded coprime, rule_format, of p] prime_1 | |
| 561 | have False by simp } | |
| 562 | moreover | |
| 563 |     {assume py: "p dvd y"
 | |
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27567diff
changeset | 564 | from dvd_mult[OF py, of y] p(3) | 
| 53077 | 565 | obtain r s where "y * y = p * r" and "x\<^sup>2 + y\<^sup>2 = p * s" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27567diff
changeset | 566 | by (auto elim!: dvdE) | 
| 53077 | 567 | then have "x\<^sup>2 = p * (s - r)" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27567diff
changeset | 568 | by (auto simp add: power2_eq_square diff_mult_distrib2) | 
| 53077 | 569 | then have "p dvd x\<^sup>2" .. | 
| 26125 | 570 | with prime_divexp[OF p(1), of x 2] have px: "p dvd x" . | 
| 571 | from p(1) px py xy[unfolded coprime, rule_format, of p] prime_1 | |
| 572 | have False by simp } | |
| 573 | ultimately have False using prime_divprod[OF p(1,2)] by blast} | |
| 574 | thus ?thesis by blast | |
| 575 | qed | |
| 576 | ||
| 577 | lemma distinct_prime_coprime: "prime p \<Longrightarrow> prime q \<Longrightarrow> p \<noteq> q \<Longrightarrow> coprime p q" | |
| 578 | unfolding prime_def coprime_prime_eq by blast | |
| 579 | ||
| 580 | lemma prime_coprime_lt: assumes p: "prime p" and x: "0 < x" and xp: "x < p" | |
| 581 | shows "coprime x p" | |
| 582 | proof- | |
| 583 |   {assume c: "\<not> coprime x p"
 | |
| 584 | then obtain g where g: "g \<noteq> 1" "g dvd x" "g dvd p" unfolding coprime_def by blast | |
| 585 | from dvd_imp_le[OF g(2)] x xp have gp: "g < p" by arith | |
| 586 | from g(2) x have "g \<noteq> 0" by - (rule ccontr, simp) | |
| 587 | with g gp p[unfolded prime_def] have False by blast} | |
| 588 | thus ?thesis by blast | |
| 589 | qed | |
| 590 | ||
| 591 | lemma even_dvd[simp]: "even (n::nat) \<longleftrightarrow> 2 dvd n" by presburger | |
| 592 | lemma prime_odd: "prime p \<Longrightarrow> p = 2 \<or> odd p" unfolding prime_def by auto | |
| 593 | ||
| 26144 | 594 | |
| 595 | text {* One property of coprimality is easier to prove via prime factors. *}
 | |
| 26125 | 596 | |
| 597 | lemma prime_divprod_pow: | |
| 598 | assumes p: "prime p" and ab: "coprime a b" and pab: "p^n dvd a * b" | |
| 599 | shows "p^n dvd a \<or> p^n dvd b" | |
| 600 | proof- | |
| 601 |   {assume "n = 0 \<or> a = 1 \<or> b = 1" with pab have ?thesis 
 | |
| 602 | apply (cases "n=0", simp_all) | |
| 603 | apply (cases "a=1", simp_all) done} | |
| 604 | moreover | |
| 605 |   {assume n: "n \<noteq> 0" and a: "a\<noteq>1" and b: "b\<noteq>1" 
 | |
| 606 | then obtain m where m: "n = Suc m" by (cases n, auto) | |
| 607 | from divides_exp2[OF n pab] have pab': "p dvd a*b" . | |
| 608 | from prime_divprod[OF p pab'] | |
| 609 | have "p dvd a \<or> p dvd b" . | |
| 610 | moreover | |
| 611 |     {assume pa: "p dvd a"
 | |
| 612 | have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute) | |
| 613 | from coprime_prime[OF ab, of p] p pa have "\<not> p dvd b" by blast | |
| 614 | with prime_coprime[OF p, of b] b | |
| 615 | have cpb: "coprime b p" using coprime_commute by blast | |
| 616 | from coprime_exp[OF cpb] have pnb: "coprime (p^n) b" | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 617 | by (simp add: coprime_commute) | 
| 26125 | 618 | from coprime_divprod[OF pnba pnb] have ?thesis by blast } | 
| 619 | moreover | |
| 620 |     {assume pb: "p dvd b"
 | |
| 621 | have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute) | |
| 622 | from coprime_prime[OF ab, of p] p pb have "\<not> p dvd a" by blast | |
| 623 | with prime_coprime[OF p, of a] a | |
| 624 | have cpb: "coprime a p" using coprime_commute by blast | |
| 625 | from coprime_exp[OF cpb] have pnb: "coprime (p^n) a" | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 626 | by (simp add: coprime_commute) | 
| 26125 | 627 | from coprime_divprod[OF pab pnb] have ?thesis by blast } | 
| 628 | ultimately have ?thesis by blast} | |
| 629 | ultimately show ?thesis by blast | |
| 630 | qed | |
| 631 | ||
| 632 | lemma nat_mult_eq_one: "(n::nat) * m = 1 \<longleftrightarrow> n = 1 \<and> m = 1" (is "?lhs \<longleftrightarrow> ?rhs") | |
| 633 | proof | |
| 634 | assume H: "?lhs" | |
| 635 | hence "n dvd 1" "m dvd 1" unfolding dvd_def by (auto simp add: mult_commute) | |
| 636 | thus ?rhs by auto | |
| 637 | next | |
| 638 | assume ?rhs then show ?lhs by auto | |
| 639 | qed | |
| 640 | ||
| 41541 | 641 | lemma power_Suc0: "Suc 0 ^ n = Suc 0" | 
| 26125 | 642 | unfolding One_nat_def[symmetric] power_one .. | 
| 41541 | 643 | |
| 26125 | 644 | lemma coprime_pow: assumes ab: "coprime a b" and abcn: "a * b = c ^n" | 
| 645 | shows "\<exists>r s. a = r^n \<and> b = s ^n" | |
| 646 | using ab abcn | |
| 647 | proof(induct c arbitrary: a b rule: nat_less_induct) | |
| 648 | fix c a b | |
| 649 | assume H: "\<forall>m<c. \<forall>a b. coprime a b \<longrightarrow> a * b = m ^ n \<longrightarrow> (\<exists>r s. a = r ^ n \<and> b = s ^ n)" "coprime a b" "a * b = c ^ n" | |
| 650 | let ?ths = "\<exists>r s. a = r^n \<and> b = s ^n" | |
| 651 |   {assume n: "n = 0"
 | |
| 652 | with H(3) power_one have "a*b = 1" by simp | |
| 653 | hence "a = 1 \<and> b = 1" by simp | |
| 654 | hence ?ths | |
| 655 | apply - | |
| 656 | apply (rule exI[where x=1]) | |
| 657 | apply (rule exI[where x=1]) | |
| 658 | using power_one[of n] | |
| 659 | by simp} | |
| 660 | moreover | |
| 661 |   {assume n: "n \<noteq> 0" then obtain m where m: "n = Suc m" by (cases n, auto)
 | |
| 662 |     {assume c: "c = 0"
 | |
| 663 | with H(3) m H(2) have ?ths apply simp | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 664 | apply (cases "a=0", simp_all) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 665 | apply (rule exI[where x="0"], simp) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 666 | apply (rule exI[where x="0"], simp) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 667 | done} | 
| 26125 | 668 | moreover | 
| 669 |     {assume "c=1" with H(3) power_one have "a*b = 1" by simp 
 | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 670 | hence "a = 1 \<and> b = 1" by simp | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 671 | hence ?ths | 
| 26125 | 672 | apply - | 
| 673 | apply (rule exI[where x=1]) | |
| 674 | apply (rule exI[where x=1]) | |
| 675 | using power_one[of n] | |
| 676 | by simp} | |
| 677 | moreover | |
| 678 |   {assume c: "c\<noteq>1" "c \<noteq> 0"
 | |
| 679 | from prime_factor[OF c(1)] obtain p where p: "prime p" "p dvd c" by blast | |
| 680 | from prime_divprod_pow[OF p(1) H(2), unfolded H(3), OF divides_exp[OF p(2), of n]] | |
| 681 | have pnab: "p ^ n dvd a \<or> p^n dvd b" . | |
| 682 | from p(2) obtain l where l: "c = p*l" unfolding dvd_def by blast | |
| 41541 | 683 | have pn0: "p^n \<noteq> 0" using n prime_ge_2 [OF p(1)] by simp | 
| 26125 | 684 |     {assume pa: "p^n dvd a"
 | 
| 685 | then obtain k where k: "a = p^n * k" unfolding dvd_def by blast | |
| 686 | from l have "l dvd c" by auto | |
| 687 | with dvd_imp_le[of l c] c have "l \<le> c" by auto | |
| 688 |       moreover {assume "l = c" with l c  have "p = 1" by simp with p have False by simp}
 | |
| 689 | ultimately have lc: "l < c" by arith | |
| 690 | from coprime_lmul2 [OF H(2)[unfolded k coprime_commute[of "p^n*k" b]]] | |
| 691 | have kb: "coprime k b" by (simp add: coprime_commute) | |
| 692 | from H(3) l k pn0 have kbln: "k * b = l ^ n" | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 693 | by (auto simp add: power_mult_distrib) | 
| 26125 | 694 | from H(1)[rule_format, OF lc kb kbln] | 
| 695 | obtain r s where rs: "k = r ^n" "b = s^n" by blast | |
| 696 | from k rs(1) have "a = (p*r)^n" by (simp add: power_mult_distrib) | |
| 697 | with rs(2) have ?ths by blast } | |
| 698 | moreover | |
| 699 |     {assume pb: "p^n dvd b"
 | |
| 700 | then obtain k where k: "b = p^n * k" unfolding dvd_def by blast | |
| 701 | from l have "l dvd c" by auto | |
| 702 | with dvd_imp_le[of l c] c have "l \<le> c" by auto | |
| 703 |       moreover {assume "l = c" with l c  have "p = 1" by simp with p have False by simp}
 | |
| 704 | ultimately have lc: "l < c" by arith | |
| 705 | from coprime_lmul2 [OF H(2)[unfolded k coprime_commute[of "p^n*k" a]]] | |
| 706 | have kb: "coprime k a" by (simp add: coprime_commute) | |
| 707 | from H(3) l k pn0 n have kbln: "k * a = l ^ n" | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 708 | by (simp add: power_mult_distrib mult_commute) | 
| 26125 | 709 | from H(1)[rule_format, OF lc kb kbln] | 
| 710 | obtain r s where rs: "k = r ^n" "a = s^n" by blast | |
| 711 | from k rs(1) have "b = (p*r)^n" by (simp add: power_mult_distrib) | |
| 712 | with rs(2) have ?ths by blast } | |
| 713 | ultimately have ?ths using pnab by blast} | |
| 714 | ultimately have ?ths by blast} | |
| 715 | ultimately show ?ths by blast | |
| 716 | qed | |
| 717 | ||
| 26144 | 718 | text {* More useful lemmas. *}
 | 
| 26125 | 719 | lemma prime_product: | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27567diff
changeset | 720 | assumes "prime (p * q)" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27567diff
changeset | 721 | shows "p = 1 \<or> q = 1" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27567diff
changeset | 722 | proof - | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27567diff
changeset | 723 | from assms have | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27567diff
changeset | 724 | "1 < p * q" and P: "\<And>m. m dvd p * q \<Longrightarrow> m = 1 \<or> m = p * q" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27567diff
changeset | 725 | unfolding prime_def by auto | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27567diff
changeset | 726 | from `1 < p * q` have "p \<noteq> 0" by (cases p) auto | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27567diff
changeset | 727 | then have Q: "p = p * q \<longleftrightarrow> q = 1" by auto | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27567diff
changeset | 728 | have "p dvd p * q" by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27567diff
changeset | 729 | then have "p = 1 \<or> p = p * q" by (rule P) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27567diff
changeset | 730 | then show ?thesis by (simp add: Q) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27567diff
changeset | 731 | qed | 
| 26125 | 732 | |
| 733 | lemma prime_exp: "prime (p^n) \<longleftrightarrow> prime p \<and> n = 1" | |
| 734 | proof(induct n) | |
| 735 | case 0 thus ?case by simp | |
| 736 | next | |
| 737 | case (Suc n) | |
| 738 |   {assume "p = 0" hence ?case by simp}
 | |
| 739 | moreover | |
| 740 |   {assume "p=1" hence ?case by simp}
 | |
| 741 | moreover | |
| 742 |   {assume p: "p \<noteq> 0" "p\<noteq>1"
 | |
| 743 |     {assume pp: "prime (p^Suc n)"
 | |
| 744 | hence "p = 1 \<or> p^n = 1" using prime_product[of p "p^n"] by simp | |
| 745 | with p have n: "n = 0" | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 746 | by (simp only: exp_eq_1 ) simp | 
| 26125 | 747 | with pp have "prime p \<and> Suc n = 1" by simp} | 
| 748 | moreover | |
| 749 |     {assume n: "prime p \<and> Suc n = 1" hence "prime (p^Suc n)" by simp}
 | |
| 750 | ultimately have ?case by blast} | |
| 751 | ultimately show ?case by blast | |
| 752 | qed | |
| 753 | ||
| 754 | lemma prime_power_mult: | |
| 755 | assumes p: "prime p" and xy: "x * y = p ^ k" | |
| 756 | shows "\<exists>i j. x = p ^i \<and> y = p^ j" | |
| 757 | using xy | |
| 758 | proof(induct k arbitrary: x y) | |
| 759 | case 0 thus ?case apply simp by (rule exI[where x="0"], simp) | |
| 760 | next | |
| 761 | case (Suc k x y) | |
| 762 | from Suc.prems have pxy: "p dvd x*y" by auto | |
| 763 | from prime_divprod[OF p pxy] have pxyc: "p dvd x \<or> p dvd y" . | |
| 764 | from p have p0: "p \<noteq> 0" by - (rule ccontr, simp) | |
| 765 |   {assume px: "p dvd x"
 | |
| 766 | then obtain d where d: "x = p*d" unfolding dvd_def by blast | |
| 767 | from Suc.prems d have "p*d*y = p^Suc k" by simp | |
| 768 | hence th: "d*y = p^k" using p0 by simp | |
| 769 | from Suc.hyps[OF th] obtain i j where ij: "d = p^i" "y = p^j" by blast | |
| 770 | with d have "x = p^Suc i" by simp | |
| 771 | with ij(2) have ?case by blast} | |
| 772 | moreover | |
| 773 |   {assume px: "p dvd y"
 | |
| 774 | then obtain d where d: "y = p*d" unfolding dvd_def by blast | |
| 775 | from Suc.prems d have "p*d*x = p^Suc k" by (simp add: mult_commute) | |
| 776 | hence th: "d*x = p^k" using p0 by simp | |
| 777 | from Suc.hyps[OF th] obtain i j where ij: "d = p^i" "x = p^j" by blast | |
| 778 | with d have "y = p^Suc i" by simp | |
| 779 | with ij(2) have ?case by blast} | |
| 780 | ultimately show ?case using pxyc by blast | |
| 781 | qed | |
| 782 | ||
| 783 | lemma prime_power_exp: assumes p: "prime p" and n:"n \<noteq> 0" | |
| 784 | and xn: "x^n = p^k" shows "\<exists>i. x = p^i" | |
| 785 | using n xn | |
| 786 | proof(induct n arbitrary: k) | |
| 787 | case 0 thus ?case by simp | |
| 788 | next | |
| 789 | case (Suc n k) hence th: "x*x^n = p^k" by simp | |
| 41541 | 790 |   {assume "n = 0" with Suc have ?case by simp (rule exI[where x="k"], simp)}
 | 
| 26125 | 791 | moreover | 
| 792 |   {assume n: "n \<noteq> 0"
 | |
| 793 | from prime_power_mult[OF p th] | |
| 794 | obtain i j where ij: "x = p^i" "x^n = p^j"by blast | |
| 795 | from Suc.hyps[OF n ij(2)] have ?case .} | |
| 796 | ultimately show ?case by blast | |
| 797 | qed | |
| 798 | ||
| 799 | lemma divides_primepow: assumes p: "prime p" | |
| 800 | shows "d dvd p^k \<longleftrightarrow> (\<exists> i. i \<le> k \<and> d = p ^i)" | |
| 801 | proof | |
| 802 | assume H: "d dvd p^k" then obtain e where e: "d*e = p^k" | |
| 803 | unfolding dvd_def apply (auto simp add: mult_commute) by blast | |
| 804 | from prime_power_mult[OF p e] obtain i j where ij: "d = p^i" "e=p^j" by blast | |
| 805 | from prime_ge_2[OF p] have p1: "p > 1" by arith | |
| 806 | from e ij have "p^(i + j) = p^k" by (simp add: power_add) | |
| 807 | hence "i + j = k" using power_inject_exp[of p "i+j" k, OF p1] by simp | |
| 808 | hence "i \<le> k" by arith | |
| 809 | with ij(1) show "\<exists>i\<le>k. d = p ^ i" by blast | |
| 810 | next | |
| 811 |   {fix i assume H: "i \<le> k" "d = p^i"
 | |
| 812 | hence "\<exists>j. k = i + j" by arith | |
| 813 | then obtain j where j: "k = i + j" by blast | |
| 814 | hence "p^k = p^j*d" using H(2) by (simp add: power_add) | |
| 815 | hence "d dvd p^k" unfolding dvd_def by auto} | |
| 816 | thus "\<exists>i\<le>k. d = p ^ i \<Longrightarrow> d dvd p ^ k" by blast | |
| 817 | qed | |
| 818 | ||
| 819 | lemma coprime_divisors: "d dvd a \<Longrightarrow> e dvd b \<Longrightarrow> coprime a b \<Longrightarrow> coprime d e" | |
| 820 | by (auto simp add: dvd_def coprime) | |
| 821 | ||
| 34223 | 822 | lemma mult_inj_if_coprime_nat: | 
| 823 | "inj_on f A \<Longrightarrow> inj_on g B \<Longrightarrow> ALL a:A. ALL b:B. coprime (f a) (g b) | |
| 824 | \<Longrightarrow> inj_on (%(a,b). f a * g b::nat) (A \<times> B)" | |
| 825 | apply(auto simp add:inj_on_def) | |
| 826 | apply(metis coprime_def dvd_triv_left gcd_proj2_if_dvd_nat gcd_semilattice_nat.inf_commute relprime_dvd_mult) | |
| 827 | apply(metis coprime_commute coprime_divprod dvd.neq_le_trans dvd_triv_right) | |
| 828 | done | |
| 829 | ||
| 26159 | 830 | declare power_Suc0[simp del] | 
| 831 | declare even_dvd[simp del] | |
| 26757 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 haftmann parents: 
26159diff
changeset | 832 | |
| 11363 | 833 | end |