| author | wenzelm | 
| Fri, 15 Aug 2008 15:51:06 +0200 | |
| changeset 27890 | 62ac62e30ab1 | 
| parent 24893 | b8ef7afe3a6b | 
| child 32960 | 69916a850301 | 
| permissions | -rw-r--r-- | 
| 1792 
75c54074cd8c
The "divides" relation, the greatest common divisor and Euclid's algorithm
 paulson parents: diff
changeset | 1 | (* Title: ZF/ex/Primes.thy | 
| 
75c54074cd8c
The "divides" relation, the greatest common divisor and Euclid's algorithm
 paulson parents: diff
changeset | 2 | ID: $Id$ | 
| 
75c54074cd8c
The "divides" relation, the greatest common divisor and Euclid's algorithm
 paulson parents: diff
changeset | 3 | Author: Christophe Tabacznyj and Lawrence C Paulson | 
| 
75c54074cd8c
The "divides" relation, the greatest common divisor and Euclid's algorithm
 paulson parents: diff
changeset | 4 | Copyright 1996 University of Cambridge | 
| 15863 | 5 | *) | 
| 1792 
75c54074cd8c
The "divides" relation, the greatest common divisor and Euclid's algorithm
 paulson parents: diff
changeset | 6 | |
| 15863 | 7 | header{*The Divides Relation and Euclid's algorithm for the GCD*}
 | 
| 1792 
75c54074cd8c
The "divides" relation, the greatest common divisor and Euclid's algorithm
 paulson parents: diff
changeset | 8 | |
| 16417 | 9 | theory Primes imports Main begin | 
| 24893 | 10 | |
| 11 | definition | |
| 12 | divides :: "[i,i]=>o" (infixl "dvd" 50) where | |
| 12867 | 13 | "m dvd n == m \<in> nat & n \<in> nat & (\<exists>k \<in> nat. n = m#*k)" | 
| 1792 
75c54074cd8c
The "divides" relation, the greatest common divisor and Euclid's algorithm
 paulson parents: diff
changeset | 14 | |
| 24893 | 15 | definition | 
| 16 |   is_gcd  :: "[i,i,i]=>o"     --{*definition of great common divisor*}  where
 | |
| 12867 | 17 | "is_gcd(p,m,n) == ((p dvd m) & (p dvd n)) & | 
| 18 | (\<forall>d\<in>nat. (d dvd m) & (d dvd n) --> d dvd p)" | |
| 1792 
75c54074cd8c
The "divides" relation, the greatest common divisor and Euclid's algorithm
 paulson parents: diff
changeset | 19 | |
| 24893 | 20 | definition | 
| 21 |   gcd     :: "[i,i]=>i"       --{*Euclid's algorithm for the gcd*}  where
 | |
| 12867 | 22 | "gcd(m,n) == transrec(natify(n), | 
| 23 | %n f. \<lambda>m \<in> nat. | |
| 11382 
a816fead971a
now more like the HOL versions, and with the Square Root example added
 paulson parents: 
11316diff
changeset | 24 | if n=0 then m else f`(m mod n)`n) ` natify(m)" | 
| 1792 
75c54074cd8c
The "divides" relation, the greatest common divisor and Euclid's algorithm
 paulson parents: diff
changeset | 25 | |
| 24893 | 26 | definition | 
| 27 |   coprime :: "[i,i]=>o"       --{*the coprime relation*}  where
 | |
| 11382 
a816fead971a
now more like the HOL versions, and with the Square Root example added
 paulson parents: 
11316diff
changeset | 28 | "coprime(m,n) == gcd(m,n) = 1" | 
| 
a816fead971a
now more like the HOL versions, and with the Square Root example added
 paulson parents: 
11316diff
changeset | 29 | |
| 24893 | 30 | definition | 
| 31 |   prime   :: i                --{*the set of prime numbers*}  where
 | |
| 12867 | 32 |    "prime == {p \<in> nat. 1<p & (\<forall>m \<in> nat. m dvd p --> m=1 | m=p)}"
 | 
| 33 | ||
| 15863 | 34 | |
| 35 | subsection{*The Divides Relation*}
 | |
| 12867 | 36 | |
| 37 | lemma dvdD: "m dvd n ==> m \<in> nat & n \<in> nat & (\<exists>k \<in> nat. n = m#*k)" | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13259diff
changeset | 38 | by (unfold divides_def, assumption) | 
| 12867 | 39 | |
| 40 | lemma dvdE: | |
| 41 | "[|m dvd n; !!k. [|m \<in> nat; n \<in> nat; k \<in> nat; n = m#*k|] ==> P|] ==> P" | |
| 42 | by (blast dest!: dvdD) | |
| 43 | ||
| 44 | lemmas dvd_imp_nat1 = dvdD [THEN conjunct1, standard] | |
| 45 | lemmas dvd_imp_nat2 = dvdD [THEN conjunct2, THEN conjunct1, standard] | |
| 46 | ||
| 47 | ||
| 48 | lemma dvd_0_right [simp]: "m \<in> nat ==> m dvd 0" | |
| 15863 | 49 | apply (simp add: divides_def) | 
| 12867 | 50 | apply (fast intro: nat_0I mult_0_right [symmetric]) | 
| 51 | done | |
| 52 | ||
| 53 | lemma dvd_0_left: "0 dvd m ==> m = 0" | |
| 15863 | 54 | by (simp add: divides_def) | 
| 12867 | 55 | |
| 56 | lemma dvd_refl [simp]: "m \<in> nat ==> m dvd m" | |
| 15863 | 57 | apply (simp add: divides_def) | 
| 12867 | 58 | apply (fast intro: nat_1I mult_1_right [symmetric]) | 
| 59 | done | |
| 60 | ||
| 61 | lemma dvd_trans: "[| m dvd n; n dvd p |] ==> m dvd p" | |
| 15863 | 62 | by (auto simp add: divides_def intro: mult_assoc mult_type) | 
| 12867 | 63 | |
| 64 | lemma dvd_anti_sym: "[| m dvd n; n dvd m |] ==> m=n" | |
| 15863 | 65 | apply (simp add: divides_def) | 
| 12867 | 66 | apply (force dest: mult_eq_self_implies_10 | 
| 67 | simp add: mult_assoc mult_eq_1_iff) | |
| 68 | done | |
| 69 | ||
| 70 | lemma dvd_mult_left: "[|(i#*j) dvd k; i \<in> nat|] ==> i dvd k" | |
| 15863 | 71 | by (auto simp add: divides_def mult_assoc) | 
| 12867 | 72 | |
| 73 | lemma dvd_mult_right: "[|(i#*j) dvd k; j \<in> nat|] ==> j dvd k" | |
| 15863 | 74 | apply (simp add: divides_def, clarify) | 
| 12867 | 75 | apply (rule_tac x = "i#*k" in bexI) | 
| 76 | apply (simp add: mult_ac) | |
| 77 | apply (rule mult_type) | |
| 78 | done | |
| 79 | ||
| 80 | ||
| 15863 | 81 | subsection{*Euclid's Algorithm for the GCD*}
 | 
| 12867 | 82 | |
| 83 | lemma gcd_0 [simp]: "gcd(m,0) = natify(m)" | |
| 15863 | 84 | apply (simp add: gcd_def) | 
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13259diff
changeset | 85 | apply (subst transrec, simp) | 
| 12867 | 86 | done | 
| 87 | ||
| 88 | lemma gcd_natify1 [simp]: "gcd(natify(m),n) = gcd(m,n)" | |
| 89 | by (simp add: gcd_def) | |
| 90 | ||
| 91 | lemma gcd_natify2 [simp]: "gcd(m, natify(n)) = gcd(m,n)" | |
| 92 | by (simp add: gcd_def) | |
| 93 | ||
| 94 | lemma gcd_non_0_raw: | |
| 95 | "[| 0<n; n \<in> nat |] ==> gcd(m,n) = gcd(n, m mod n)" | |
| 15863 | 96 | apply (simp add: gcd_def) | 
| 12867 | 97 | apply (rule_tac P = "%z. ?left (z) = ?right" in transrec [THEN ssubst]) | 
| 98 | apply (simp add: ltD [THEN mem_imp_not_eq, THEN not_sym] | |
| 99 | mod_less_divisor [THEN ltD]) | |
| 100 | done | |
| 101 | ||
| 102 | lemma gcd_non_0: "0 < natify(n) ==> gcd(m,n) = gcd(n, m mod n)" | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13259diff
changeset | 103 | apply (cut_tac m = m and n = "natify (n) " in gcd_non_0_raw) | 
| 12867 | 104 | apply auto | 
| 105 | done | |
| 106 | ||
| 107 | lemma gcd_1 [simp]: "gcd(m,1) = 1" | |
| 108 | by (simp (no_asm_simp) add: gcd_non_0) | |
| 109 | ||
| 110 | lemma dvd_add: "[| k dvd a; k dvd b |] ==> k dvd (a #+ b)" | |
| 15863 | 111 | apply (simp add: divides_def) | 
| 12867 | 112 | apply (fast intro: add_mult_distrib_left [symmetric] add_type) | 
| 113 | done | |
| 114 | ||
| 115 | lemma dvd_mult: "k dvd n ==> k dvd (m #* n)" | |
| 15863 | 116 | apply (simp add: divides_def) | 
| 12867 | 117 | apply (fast intro: mult_left_commute mult_type) | 
| 118 | done | |
| 119 | ||
| 120 | lemma dvd_mult2: "k dvd m ==> k dvd (m #* n)" | |
| 121 | apply (subst mult_commute) | |
| 122 | apply (blast intro: dvd_mult) | |
| 123 | done | |
| 124 | ||
| 125 | (* k dvd (m*k) *) | |
| 126 | lemmas dvdI1 [simp] = dvd_refl [THEN dvd_mult, standard] | |
| 127 | lemmas dvdI2 [simp] = dvd_refl [THEN dvd_mult2, standard] | |
| 128 | ||
| 129 | lemma dvd_mod_imp_dvd_raw: | |
| 130 | "[| a \<in> nat; b \<in> nat; k dvd b; k dvd (a mod b) |] ==> k dvd a" | |
| 15863 | 131 | apply (case_tac "b=0") | 
| 13259 | 132 | apply (simp add: DIVISION_BY_ZERO_MOD) | 
| 12867 | 133 | apply (blast intro: mod_div_equality [THEN subst] | 
| 134 | elim: dvdE | |
| 135 | intro!: dvd_add dvd_mult mult_type mod_type div_type) | |
| 136 | done | |
| 137 | ||
| 138 | lemma dvd_mod_imp_dvd: "[| k dvd (a mod b); k dvd b; a \<in> nat |] ==> k dvd a" | |
| 13259 | 139 | apply (cut_tac b = "natify (b)" in dvd_mod_imp_dvd_raw) | 
| 12867 | 140 | apply auto | 
| 141 | apply (simp add: divides_def) | |
| 142 | done | |
| 143 | ||
| 144 | (*Imitating TFL*) | |
| 145 | lemma gcd_induct_lemma [rule_format (no_asm)]: "[| n \<in> nat; | |
| 146 | \<forall>m \<in> nat. P(m,0); | |
| 147 | \<forall>m \<in> nat. \<forall>n \<in> nat. 0<n --> P(n, m mod n) --> P(m,n) |] | |
| 148 | ==> \<forall>m \<in> nat. P (m,n)" | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13259diff
changeset | 149 | apply (erule_tac i = n in complete_induct) | 
| 12867 | 150 | apply (case_tac "x=0") | 
| 151 | apply (simp (no_asm_simp)) | |
| 152 | apply clarify | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13259diff
changeset | 153 | apply (drule_tac x1 = m and x = x in bspec [THEN bspec]) | 
| 12867 | 154 | apply (simp_all add: Ord_0_lt_iff) | 
| 155 | apply (blast intro: mod_less_divisor [THEN ltD]) | |
| 156 | done | |
| 157 | ||
| 158 | lemma gcd_induct: "!!P. [| m \<in> nat; n \<in> nat; | |
| 159 | !!m. m \<in> nat ==> P(m,0); | |
| 160 | !!m n. [|m \<in> nat; n \<in> nat; 0<n; P(n, m mod n)|] ==> P(m,n) |] | |
| 161 | ==> P (m,n)" | |
| 162 | by (blast intro: gcd_induct_lemma) | |
| 163 | ||
| 164 | ||
| 15863 | 165 | subsection{*Basic Properties of @{term gcd}*}
 | 
| 12867 | 166 | |
| 15863 | 167 | text{*type of gcd*}
 | 
| 12867 | 168 | lemma gcd_type [simp,TC]: "gcd(m, n) \<in> nat" | 
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13259diff
changeset | 169 | apply (subgoal_tac "gcd (natify (m), natify (n)) \<in> nat") | 
| 12867 | 170 | apply simp | 
| 13259 | 171 | apply (rule_tac m = "natify (m)" and n = "natify (n)" in gcd_induct) | 
| 12867 | 172 | apply auto | 
| 173 | apply (simp add: gcd_non_0) | |
| 174 | done | |
| 175 | ||
| 176 | ||
| 15863 | 177 | text{* Property 1: gcd(a,b) divides a and b *}
 | 
| 12867 | 178 | |
| 179 | lemma gcd_dvd_both: | |
| 180 | "[| m \<in> nat; n \<in> nat |] ==> gcd (m, n) dvd m & gcd (m, n) dvd n" | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13259diff
changeset | 181 | apply (rule_tac m = m and n = n in gcd_induct) | 
| 12867 | 182 | apply (simp_all add: gcd_non_0) | 
| 183 | apply (blast intro: dvd_mod_imp_dvd_raw nat_into_Ord [THEN Ord_0_lt]) | |
| 184 | done | |
| 185 | ||
| 186 | lemma gcd_dvd1 [simp]: "m \<in> nat ==> gcd(m,n) dvd m" | |
| 13259 | 187 | apply (cut_tac m = "natify (m)" and n = "natify (n)" in gcd_dvd_both) | 
| 12867 | 188 | apply auto | 
| 189 | done | |
| 190 | ||
| 191 | lemma gcd_dvd2 [simp]: "n \<in> nat ==> gcd(m,n) dvd n" | |
| 13259 | 192 | apply (cut_tac m = "natify (m)" and n = "natify (n)" in gcd_dvd_both) | 
| 12867 | 193 | apply auto | 
| 194 | done | |
| 195 | ||
| 15863 | 196 | text{* if f divides a and b then f divides gcd(a,b) *}
 | 
| 12867 | 197 | |
| 198 | lemma dvd_mod: "[| f dvd a; f dvd b |] ==> f dvd (a mod b)" | |
| 15863 | 199 | apply (simp add: divides_def) | 
| 13259 | 200 | apply (case_tac "b=0") | 
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13259diff
changeset | 201 | apply (simp add: DIVISION_BY_ZERO_MOD, auto) | 
| 12867 | 202 | apply (blast intro: mod_mult_distrib2 [symmetric]) | 
| 203 | done | |
| 204 | ||
| 15863 | 205 | text{* Property 2: for all a,b,f naturals, 
 | 
| 206 | if f divides a and f divides b then f divides gcd(a,b)*} | |
| 12867 | 207 | |
| 208 | lemma gcd_greatest_raw [rule_format]: | |
| 209 | "[| m \<in> nat; n \<in> nat; f \<in> nat |] | |
| 210 | ==> (f dvd m) --> (f dvd n) --> f dvd gcd(m,n)" | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13259diff
changeset | 211 | apply (rule_tac m = m and n = n in gcd_induct) | 
| 12867 | 212 | apply (simp_all add: gcd_non_0 dvd_mod) | 
| 213 | done | |
| 214 | ||
| 215 | lemma gcd_greatest: "[| f dvd m; f dvd n; f \<in> nat |] ==> f dvd gcd(m,n)" | |
| 216 | apply (rule gcd_greatest_raw) | |
| 217 | apply (auto simp add: divides_def) | |
| 218 | done | |
| 219 | ||
| 220 | lemma gcd_greatest_iff [simp]: "[| k \<in> nat; m \<in> nat; n \<in> nat |] | |
| 221 | ==> (k dvd gcd (m, n)) <-> (k dvd m & k dvd n)" | |
| 222 | by (blast intro!: gcd_greatest gcd_dvd1 gcd_dvd2 intro: dvd_trans) | |
| 223 | ||
| 224 | ||
| 15863 | 225 | subsection{*The Greatest Common Divisor*}
 | 
| 226 | ||
| 227 | text{*The GCD exists and function gcd computes it.*}
 | |
| 12867 | 228 | |
| 229 | lemma is_gcd: "[| m \<in> nat; n \<in> nat |] ==> is_gcd(gcd(m,n), m, n)" | |
| 230 | by (simp add: is_gcd_def) | |
| 231 | ||
| 15863 | 232 | text{*The GCD is unique*}
 | 
| 12867 | 233 | |
| 234 | lemma is_gcd_unique: "[|is_gcd(m,a,b); is_gcd(n,a,b); m\<in>nat; n\<in>nat|] ==> m=n" | |
| 15863 | 235 | apply (simp add: is_gcd_def) | 
| 12867 | 236 | apply (blast intro: dvd_anti_sym) | 
| 237 | done | |
| 238 | ||
| 239 | lemma is_gcd_commute: "is_gcd(k,m,n) <-> is_gcd(k,n,m)" | |
| 15863 | 240 | by (simp add: is_gcd_def, blast) | 
| 12867 | 241 | |
| 242 | lemma gcd_commute_raw: "[| m \<in> nat; n \<in> nat |] ==> gcd(m,n) = gcd(n,m)" | |
| 243 | apply (rule is_gcd_unique) | |
| 244 | apply (rule is_gcd) | |
| 245 | apply (rule_tac [3] is_gcd_commute [THEN iffD1]) | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13259diff
changeset | 246 | apply (rule_tac [3] is_gcd, auto) | 
| 12867 | 247 | done | 
| 248 | ||
| 249 | lemma gcd_commute: "gcd(m,n) = gcd(n,m)" | |
| 13259 | 250 | apply (cut_tac m = "natify (m)" and n = "natify (n)" in gcd_commute_raw) | 
| 12867 | 251 | apply auto | 
| 252 | done | |
| 253 | ||
| 254 | lemma gcd_assoc_raw: "[| k \<in> nat; m \<in> nat; n \<in> nat |] | |
| 255 | ==> gcd (gcd (k, m), n) = gcd (k, gcd (m, n))" | |
| 256 | apply (rule is_gcd_unique) | |
| 257 | apply (rule is_gcd) | |
| 258 | apply (simp_all add: is_gcd_def) | |
| 259 | apply (blast intro: gcd_dvd1 gcd_dvd2 gcd_type intro: dvd_trans) | |
| 260 | done | |
| 261 | ||
| 262 | lemma gcd_assoc: "gcd (gcd (k, m), n) = gcd (k, gcd (m, n))" | |
| 13259 | 263 | apply (cut_tac k = "natify (k)" and m = "natify (m)" and n = "natify (n) " | 
| 12867 | 264 | in gcd_assoc_raw) | 
| 265 | apply auto | |
| 266 | done | |
| 267 | ||
| 268 | lemma gcd_0_left [simp]: "gcd (0, m) = natify(m)" | |
| 269 | by (simp add: gcd_commute [of 0]) | |
| 270 | ||
| 271 | lemma gcd_1_left [simp]: "gcd (1, m) = 1" | |
| 272 | by (simp add: gcd_commute [of 1]) | |
| 273 | ||
| 274 | ||
| 15863 | 275 | subsection{*Addition laws*}
 | 
| 276 | ||
| 277 | lemma gcd_add1 [simp]: "gcd (m #+ n, n) = gcd (m, n)" | |
| 278 | apply (subgoal_tac "gcd (m #+ natify (n), natify (n)) = gcd (m, natify (n))") | |
| 279 | apply simp | |
| 280 | apply (case_tac "natify (n) = 0") | |
| 281 | apply (auto simp add: Ord_0_lt_iff gcd_non_0) | |
| 282 | done | |
| 283 | ||
| 284 | lemma gcd_add2 [simp]: "gcd (m, m #+ n) = gcd (m, n)" | |
| 285 | apply (rule gcd_commute [THEN trans]) | |
| 286 | apply (subst add_commute, simp) | |
| 287 | apply (rule gcd_commute) | |
| 288 | done | |
| 289 | ||
| 290 | lemma gcd_add2' [simp]: "gcd (m, n #+ m) = gcd (m, n)" | |
| 291 | by (subst add_commute, rule gcd_add2) | |
| 292 | ||
| 293 | lemma gcd_add_mult_raw: "k \<in> nat ==> gcd (m, k #* m #+ n) = gcd (m, n)" | |
| 294 | apply (erule nat_induct) | |
| 295 | apply (auto simp add: gcd_add2 add_assoc) | |
| 296 | done | |
| 297 | ||
| 298 | lemma gcd_add_mult: "gcd (m, k #* m #+ n) = gcd (m, n)" | |
| 299 | apply (cut_tac k = "natify (k)" in gcd_add_mult_raw) | |
| 300 | apply auto | |
| 301 | done | |
| 302 | ||
| 303 | ||
| 304 | subsection{* Multiplication Laws*}
 | |
| 12867 | 305 | |
| 306 | lemma gcd_mult_distrib2_raw: | |
| 307 | "[| k \<in> nat; m \<in> nat; n \<in> nat |] | |
| 308 | ==> k #* gcd (m, n) = gcd (k #* m, k #* n)" | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13259diff
changeset | 309 | apply (erule_tac m = m and n = n in gcd_induct, assumption) | 
| 12867 | 310 | apply simp | 
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13259diff
changeset | 311 | apply (case_tac "k = 0", simp) | 
| 12867 | 312 | apply (simp add: mod_geq gcd_non_0 mod_mult_distrib2 Ord_0_lt_iff) | 
| 313 | done | |
| 314 | ||
| 315 | lemma gcd_mult_distrib2: "k #* gcd (m, n) = gcd (k #* m, k #* n)" | |
| 13259 | 316 | apply (cut_tac k = "natify (k)" and m = "natify (m)" and n = "natify (n) " | 
| 12867 | 317 | in gcd_mult_distrib2_raw) | 
| 318 | apply auto | |
| 319 | done | |
| 320 | ||
| 321 | lemma gcd_mult [simp]: "gcd (k, k #* n) = natify(k)" | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13259diff
changeset | 322 | by (cut_tac k = k and m = 1 and n = n in gcd_mult_distrib2, auto) | 
| 12867 | 323 | |
| 324 | lemma gcd_self [simp]: "gcd (k, k) = natify(k)" | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13259diff
changeset | 325 | by (cut_tac k = k and n = 1 in gcd_mult, auto) | 
| 12867 | 326 | |
| 327 | lemma relprime_dvd_mult: | |
| 328 | "[| gcd (k,n) = 1; k dvd (m #* n); m \<in> nat |] ==> k dvd m" | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13259diff
changeset | 329 | apply (cut_tac k = m and m = k and n = n in gcd_mult_distrib2, auto) | 
| 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13259diff
changeset | 330 | apply (erule_tac b = m in ssubst) | 
| 12867 | 331 | apply (simp add: dvd_imp_nat1) | 
| 332 | done | |
| 333 | ||
| 334 | lemma relprime_dvd_mult_iff: | |
| 335 | "[| gcd (k,n) = 1; m \<in> nat |] ==> k dvd (m #* n) <-> k dvd m" | |
| 336 | by (blast intro: dvdI2 relprime_dvd_mult dvd_trans) | |
| 337 | ||
| 338 | lemma prime_imp_relprime: | |
| 339 | "[| p \<in> prime; ~ (p dvd n); n \<in> nat |] ==> gcd (p, n) = 1" | |
| 15863 | 340 | apply (simp add: prime_def, clarify) | 
| 13259 | 341 | apply (drule_tac x = "gcd (p,n)" in bspec) | 
| 12867 | 342 | apply auto | 
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13259diff
changeset | 343 | apply (cut_tac m = p and n = n in gcd_dvd2, auto) | 
| 12867 | 344 | done | 
| 345 | ||
| 346 | lemma prime_into_nat: "p \<in> prime ==> p \<in> nat" | |
| 15863 | 347 | by (simp add: prime_def) | 
| 12867 | 348 | |
| 349 | lemma prime_nonzero: "p \<in> prime \<Longrightarrow> p\<noteq>0" | |
| 15863 | 350 | by (auto simp add: prime_def) | 
| 12867 | 351 | |
| 352 | ||
| 15863 | 353 | text{*This theorem leads immediately to a proof of the uniqueness of
 | 
| 354 |   factorization.  If @{term p} divides a product of primes then it is
 | |
| 355 | one of those primes.*} | |
| 12867 | 356 | |
| 357 | lemma prime_dvd_mult: | |
| 358 | "[|p dvd m #* n; p \<in> prime; m \<in> nat; n \<in> nat |] ==> p dvd m \<or> p dvd n" | |
| 359 | by (blast intro: relprime_dvd_mult prime_imp_relprime prime_into_nat) | |
| 360 | ||
| 361 | ||
| 362 | lemma gcd_mult_cancel_raw: | |
| 363 | "[|gcd (k,n) = 1; m \<in> nat; n \<in> nat|] ==> gcd (k #* m, n) = gcd (m, n)" | |
| 364 | apply (rule dvd_anti_sym) | |
| 365 | apply (rule gcd_greatest) | |
| 366 | apply (rule relprime_dvd_mult [of _ k]) | |
| 367 | apply (simp add: gcd_assoc) | |
| 368 | apply (simp add: gcd_commute) | |
| 369 | apply (simp_all add: mult_commute) | |
| 370 | apply (blast intro: dvdI1 gcd_dvd1 dvd_trans) | |
| 371 | done | |
| 372 | ||
| 373 | lemma gcd_mult_cancel: "gcd (k,n) = 1 ==> gcd (k #* m, n) = gcd (m, n)" | |
| 13259 | 374 | apply (cut_tac m = "natify (m)" and n = "natify (n)" in gcd_mult_cancel_raw) | 
| 12867 | 375 | apply auto | 
| 376 | done | |
| 377 | ||
| 378 | ||
| 15863 | 379 | subsection{*The Square Root of a Prime is Irrational: Key Lemma*}
 | 
| 12867 | 380 | |
| 381 | lemma prime_dvd_other_side: | |
| 382 | "\<lbrakk>n#*n = p#*(k#*k); p \<in> prime; n \<in> nat\<rbrakk> \<Longrightarrow> p dvd n" | |
| 383 | apply (subgoal_tac "p dvd n#*n") | |
| 384 | apply (blast dest: prime_dvd_mult) | |
| 385 | apply (rule_tac j = "k#*k" in dvd_mult_left) | |
| 386 | apply (auto simp add: prime_def) | |
| 387 | done | |
| 388 | ||
| 389 | lemma reduction: | |
| 390 | "\<lbrakk>k#*k = p#*(j#*j); p \<in> prime; 0 < k; j \<in> nat; k \<in> nat\<rbrakk> | |
| 391 | \<Longrightarrow> k < p#*j & 0 < j" | |
| 392 | apply (rule ccontr) | |
| 393 | apply (simp add: not_lt_iff_le prime_into_nat) | |
| 394 | apply (erule disjE) | |
| 395 | apply (frule mult_le_mono, assumption+) | |
| 396 | apply (simp add: mult_ac) | |
| 397 | apply (auto dest!: natify_eqE | |
| 398 | simp add: not_lt_iff_le prime_into_nat mult_le_cancel_le1) | |
| 399 | apply (simp add: prime_def) | |
| 400 | apply (blast dest: lt_trans1) | |
| 401 | done | |
| 402 | ||
| 403 | lemma rearrange: "j #* (p#*j) = k#*k \<Longrightarrow> k#*k = p#*(j#*j)" | |
| 404 | by (simp add: mult_ac) | |
| 405 | ||
| 406 | lemma prime_not_square: | |
| 407 | "\<lbrakk>m \<in> nat; p \<in> prime\<rbrakk> \<Longrightarrow> \<forall>k \<in> nat. 0<k \<longrightarrow> m#*m \<noteq> p#*(k#*k)" | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13259diff
changeset | 408 | apply (erule complete_induct, clarify) | 
| 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13259diff
changeset | 409 | apply (frule prime_dvd_other_side, assumption) | 
| 12867 | 410 | apply assumption | 
| 411 | apply (erule dvdE) | |
| 412 | apply (simp add: mult_assoc mult_cancel1 prime_nonzero prime_into_nat) | |
| 413 | apply (blast dest: rearrange reduction ltD) | |
| 414 | done | |
| 1792 
75c54074cd8c
The "divides" relation, the greatest common divisor and Euclid's algorithm
 paulson parents: diff
changeset | 415 | |
| 
75c54074cd8c
The "divides" relation, the greatest common divisor and Euclid's algorithm
 paulson parents: diff
changeset | 416 | end |