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