| author | wenzelm | 
| Wed, 31 Dec 2008 00:01:51 +0100 | |
| changeset 29263 | bf99ccf71b7c | 
| parent 28952 | 15a4b2cf8c34 | 
| child 29655 | ac31940cfb69 | 
| permissions | -rw-r--r-- | 
| 21256 | 1 | (* Title: HOL/GCD.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Christophe Tabacznyj and Lawrence C Paulson | |
| 4 | Copyright 1996 University of Cambridge | |
| 5 | *) | |
| 6 | ||
| 7 | header {* The Greatest Common Divisor *}
 | |
| 8 | ||
| 9 | theory GCD | |
| 28952 
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
 haftmann parents: 
28562diff
changeset | 10 | imports Plain Presburger | 
| 21256 | 11 | begin | 
| 12 | ||
| 13 | text {*
 | |
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 14 |   See \cite{davenport92}. \bigskip
 | 
| 21256 | 15 | *} | 
| 16 | ||
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 17 | subsection {* Specification of GCD on nats *}
 | 
| 21256 | 18 | |
| 21263 | 19 | definition | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 20 |   is_gcd :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where -- {* @{term gcd} as a relation *}
 | 
| 28562 | 21 | [code del]: "is_gcd m n p \<longleftrightarrow> p dvd m \<and> p dvd n \<and> | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 22 | (\<forall>d. d dvd m \<longrightarrow> d dvd n \<longrightarrow> d dvd p)" | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 23 | |
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 24 | text {* Uniqueness *}
 | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 25 | |
| 27556 | 26 | lemma is_gcd_unique: "is_gcd a b m \<Longrightarrow> is_gcd a b n \<Longrightarrow> m = n" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 27 | by (simp add: is_gcd_def) (blast intro: dvd_anti_sym) | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 28 | |
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 29 | text {* Connection to divides relation *}
 | 
| 21256 | 30 | |
| 27556 | 31 | lemma is_gcd_dvd: "is_gcd a b m \<Longrightarrow> k dvd a \<Longrightarrow> k dvd b \<Longrightarrow> k dvd m" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 32 | by (auto simp add: is_gcd_def) | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 33 | |
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 34 | text {* Commutativity *}
 | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 35 | |
| 27556 | 36 | lemma is_gcd_commute: "is_gcd m n k = is_gcd n m k" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 37 | by (auto simp add: is_gcd_def) | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 38 | |
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 39 | |
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 40 | subsection {* GCD on nat by Euclid's algorithm *}
 | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 41 | |
| 27568 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 42 | fun | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 43 | gcd :: "nat => nat => nat" | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 44 | where | 
| 27556 | 45 | "gcd m n = (if n = 0 then m else gcd n (m mod n))" | 
| 46 | lemma gcd_induct [case_names "0" rec]: | |
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 47 | fixes m n :: nat | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 48 | assumes "\<And>m. P m 0" | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 49 | and "\<And>m n. 0 < n \<Longrightarrow> P n (m mod n) \<Longrightarrow> P m n" | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 50 | shows "P m n" | 
| 27556 | 51 | proof (induct m n rule: gcd.induct) | 
| 52 | case (1 m n) with assms show ?case by (cases "n = 0") simp_all | |
| 53 | qed | |
| 21256 | 54 | |
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 55 | lemma gcd_0 [simp, algebra]: "gcd m 0 = m" | 
| 21263 | 56 | by simp | 
| 21256 | 57 | |
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 58 | lemma gcd_0_left [simp,algebra]: "gcd 0 m = m" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 59 | by simp | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 60 | |
| 27556 | 61 | lemma gcd_non_0: "n > 0 \<Longrightarrow> gcd m n = gcd n (m mod n)" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 62 | by simp | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 63 | |
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 64 | lemma gcd_1 [simp, algebra]: "gcd m (Suc 0) = 1" | 
| 21263 | 65 | by simp | 
| 21256 | 66 | |
| 67 | declare gcd.simps [simp del] | |
| 68 | ||
| 69 | text {*
 | |
| 27556 | 70 |   \medskip @{term "gcd m n"} divides @{text m} and @{text n}.  The
 | 
| 21256 | 71 | conjunctions don't seem provable separately. | 
| 72 | *} | |
| 73 | ||
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 74 | lemma gcd_dvd1 [iff, algebra]: "gcd m n dvd m" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 75 | and gcd_dvd2 [iff, algebra]: "gcd m n dvd n" | 
| 21256 | 76 | apply (induct m n rule: gcd_induct) | 
| 21263 | 77 | apply (simp_all add: gcd_non_0) | 
| 21256 | 78 | apply (blast dest: dvd_mod_imp_dvd) | 
| 79 | done | |
| 80 | ||
| 81 | text {*
 | |
| 82 |   \medskip Maximality: for all @{term m}, @{term n}, @{term k}
 | |
| 83 |   naturals, if @{term k} divides @{term m} and @{term k} divides
 | |
| 27556 | 84 |   @{term n} then @{term k} divides @{term "gcd m n"}.
 | 
| 21256 | 85 | *} | 
| 86 | ||
| 27556 | 87 | lemma gcd_greatest: "k dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd gcd m n" | 
| 21263 | 88 | by (induct m n rule: gcd_induct) (simp_all add: gcd_non_0 dvd_mod) | 
| 21256 | 89 | |
| 90 | text {*
 | |
| 91 | \medskip Function gcd yields the Greatest Common Divisor. | |
| 92 | *} | |
| 93 | ||
| 27556 | 94 | lemma is_gcd: "is_gcd m n (gcd m n) " | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 95 | by (simp add: is_gcd_def gcd_greatest) | 
| 21256 | 96 | |
| 97 | ||
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 98 | subsection {* Derived laws for GCD *}
 | 
| 21256 | 99 | |
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 100 | lemma gcd_greatest_iff [iff, algebra]: "k dvd gcd m n \<longleftrightarrow> k dvd m \<and> k dvd n" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 101 | by (blast intro!: gcd_greatest intro: dvd_trans) | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 102 | |
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 103 | lemma gcd_zero[algebra]: "gcd m n = 0 \<longleftrightarrow> m = 0 \<and> n = 0" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 104 | by (simp only: dvd_0_left_iff [symmetric] gcd_greatest_iff) | 
| 21256 | 105 | |
| 27556 | 106 | lemma gcd_commute: "gcd m n = gcd n m" | 
| 21256 | 107 | apply (rule is_gcd_unique) | 
| 108 | apply (rule is_gcd) | |
| 109 | apply (subst is_gcd_commute) | |
| 110 | apply (simp add: is_gcd) | |
| 111 | done | |
| 112 | ||
| 27556 | 113 | lemma gcd_assoc: "gcd (gcd k m) n = gcd k (gcd m n)" | 
| 21256 | 114 | apply (rule is_gcd_unique) | 
| 115 | apply (rule is_gcd) | |
| 116 | apply (simp add: is_gcd_def) | |
| 117 | apply (blast intro: dvd_trans) | |
| 118 | done | |
| 119 | ||
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 120 | lemma gcd_1_left [simp, algebra]: "gcd (Suc 0) m = 1" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 121 | by (simp add: gcd_commute) | 
| 21256 | 122 | |
| 123 | text {*
 | |
| 124 | \medskip Multiplication laws | |
| 125 | *} | |
| 126 | ||
| 27556 | 127 | lemma gcd_mult_distrib2: "k * gcd m n = gcd (k * m) (k * n)" | 
| 21256 | 128 |     -- {* \cite[page 27]{davenport92} *}
 | 
| 129 | apply (induct m n rule: gcd_induct) | |
| 130 | apply simp | |
| 131 | apply (case_tac "k = 0") | |
| 132 | apply (simp_all add: mod_geq gcd_non_0 mod_mult_distrib2) | |
| 133 | done | |
| 134 | ||
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 135 | lemma gcd_mult [simp, algebra]: "gcd k (k * n) = k" | 
| 21256 | 136 | apply (rule gcd_mult_distrib2 [of k 1 n, simplified, symmetric]) | 
| 137 | done | |
| 138 | ||
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 139 | lemma gcd_self [simp, algebra]: "gcd k k = k" | 
| 21256 | 140 | apply (rule gcd_mult [of k 1, simplified]) | 
| 141 | done | |
| 142 | ||
| 27556 | 143 | lemma relprime_dvd_mult: "gcd k n = 1 ==> k dvd m * n ==> k dvd m" | 
| 21256 | 144 | apply (insert gcd_mult_distrib2 [of m k n]) | 
| 145 | apply simp | |
| 146 | apply (erule_tac t = m in ssubst) | |
| 147 | apply simp | |
| 148 | done | |
| 149 | ||
| 27556 | 150 | lemma relprime_dvd_mult_iff: "gcd k n = 1 ==> (k dvd m * n) = (k dvd m)" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27568diff
changeset | 151 | by (auto intro: relprime_dvd_mult dvd_mult2) | 
| 21256 | 152 | |
| 27556 | 153 | lemma gcd_mult_cancel: "gcd k n = 1 ==> gcd (k * m) n = gcd m n" | 
| 21256 | 154 | apply (rule dvd_anti_sym) | 
| 155 | apply (rule gcd_greatest) | |
| 156 | apply (rule_tac n = k in relprime_dvd_mult) | |
| 157 | apply (simp add: gcd_assoc) | |
| 158 | apply (simp add: gcd_commute) | |
| 159 | apply (simp_all add: mult_commute) | |
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27568diff
changeset | 160 | apply (blast intro: dvd_mult) | 
| 21256 | 161 | done | 
| 162 | ||
| 163 | ||
| 164 | text {* \medskip Addition laws *}
 | |
| 165 | ||
| 27676 | 166 | lemma gcd_add1 [simp, algebra]: "gcd (m + n) n = gcd m n" | 
| 167 | by (cases "n = 0") (auto simp add: gcd_non_0) | |
| 21256 | 168 | |
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 169 | lemma gcd_add2 [simp, algebra]: "gcd m (m + n) = gcd m n" | 
| 21256 | 170 | proof - | 
| 27556 | 171 | have "gcd m (m + n) = gcd (m + n) m" by (rule gcd_commute) | 
| 172 | also have "... = gcd (n + m) m" by (simp add: add_commute) | |
| 173 | also have "... = gcd n m" by simp | |
| 174 | also have "... = gcd m n" by (rule gcd_commute) | |
| 21256 | 175 | finally show ?thesis . | 
| 176 | qed | |
| 177 | ||
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 178 | lemma gcd_add2' [simp, algebra]: "gcd m (n + m) = gcd m n" | 
| 21256 | 179 | apply (subst add_commute) | 
| 180 | apply (rule gcd_add2) | |
| 181 | done | |
| 182 | ||
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 183 | lemma gcd_add_mult[algebra]: "gcd m (k * m + n) = gcd m n" | 
| 21263 | 184 | by (induct k) (simp_all add: add_assoc) | 
| 21256 | 185 | |
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 186 | lemma gcd_dvd_prod: "gcd m n dvd m * n" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 187 | using mult_dvd_mono [of 1] by auto | 
| 22027 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 188 | |
| 22367 | 189 | text {*
 | 
| 190 | \medskip Division by gcd yields rrelatively primes. | |
| 191 | *} | |
| 22027 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 192 | |
| 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 193 | lemma div_gcd_relprime: | 
| 22367 | 194 | assumes nz: "a \<noteq> 0 \<or> b \<noteq> 0" | 
| 27556 | 195 | shows "gcd (a div gcd a b) (b div gcd a b) = 1" | 
| 22367 | 196 | proof - | 
| 27556 | 197 | let ?g = "gcd a b" | 
| 22027 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 198 | let ?a' = "a div ?g" | 
| 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 199 | let ?b' = "b div ?g" | 
| 27556 | 200 | let ?g' = "gcd ?a' ?b'" | 
| 22027 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 201 | have dvdg: "?g dvd a" "?g dvd b" by simp_all | 
| 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 202 | have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by simp_all | 
| 22367 | 203 | from dvdg dvdg' obtain ka kb ka' kb' where | 
| 204 | kab: "a = ?g * ka" "b = ?g * kb" "?a' = ?g' * ka'" "?b' = ?g' * kb'" | |
| 22027 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 205 | unfolding dvd_def by blast | 
| 22367 | 206 | then have "?g * ?a' = (?g * ?g') * ka'" "?g * ?b' = (?g * ?g') * kb'" by simp_all | 
| 207 | then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b" | |
| 208 | by (auto simp add: dvd_mult_div_cancel [OF dvdg(1)] | |
| 209 | dvd_mult_div_cancel [OF dvdg(2)] dvd_def) | |
| 22027 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 210 | have "?g \<noteq> 0" using nz by (simp add: gcd_zero) | 
| 22367 | 211 | then have gp: "?g > 0" by simp | 
| 212 | from gcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" . | |
| 213 | with dvd_mult_cancel1 [OF gp] show "?g' = 1" by simp | |
| 22027 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 214 | qed | 
| 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 215 | |
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 216 | |
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 217 | lemma gcd_unique: "d dvd a\<and>d dvd b \<and> (\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d) \<longleftrightarrow> d = gcd a b" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 218 | proof(auto) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 219 | assume H: "d dvd a" "d dvd b" "\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 220 | from H(3)[rule_format] gcd_dvd1[of a b] gcd_dvd2[of a b] | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 221 | have th: "gcd a b dvd d" by blast | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 222 | from dvd_anti_sym[OF th gcd_greatest[OF H(1,2)]] show "d = gcd a b" by blast | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 223 | qed | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 224 | |
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 225 | lemma gcd_eq: assumes H: "\<forall>d. d dvd x \<and> d dvd y \<longleftrightarrow> d dvd u \<and> d dvd v" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 226 | shows "gcd x y = gcd u v" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 227 | proof- | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 228 | from H have "\<forall>d. d dvd x \<and> d dvd y \<longleftrightarrow> d dvd gcd u v" by simp | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 229 | with gcd_unique[of "gcd u v" x y] show ?thesis by auto | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 230 | qed | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 231 | |
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 232 | lemma ind_euclid: | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 233 | assumes c: " \<forall>a b. P (a::nat) b \<longleftrightarrow> P b a" and z: "\<forall>a. P a 0" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 234 | and add: "\<forall>a b. P a b \<longrightarrow> P a (a + b)" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 235 | shows "P a b" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 236 | proof(induct n\<equiv>"a+b" arbitrary: a b rule: nat_less_induct) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 237 | fix n a b | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 238 | assume H: "\<forall>m < n. \<forall>a b. m = a + b \<longrightarrow> P a b" "n = a + b" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 239 | have "a = b \<or> a < b \<or> b < a" by arith | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 240 |   moreover {assume eq: "a= b"
 | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 241 | from add[rule_format, OF z[rule_format, of a]] have "P a b" using eq by simp} | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 242 | moreover | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 243 |   {assume lt: "a < b"
 | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 244 | hence "a + b - a < n \<or> a = 0" using H(2) by arith | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 245 | moreover | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 246 |     {assume "a =0" with z c have "P a b" by blast }
 | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 247 | moreover | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 248 |     {assume ab: "a + b - a < n"
 | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 249 | have th0: "a + b - a = a + (b - a)" using lt by arith | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 250 | from add[rule_format, OF H(1)[rule_format, OF ab th0]] | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 251 | have "P a b" by (simp add: th0[symmetric])} | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 252 | ultimately have "P a b" by blast} | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 253 | moreover | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 254 |   {assume lt: "a > b"
 | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 255 | hence "b + a - b < n \<or> b = 0" using H(2) by arith | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 256 | moreover | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 257 |     {assume "b =0" with z c have "P a b" by blast }
 | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 258 | moreover | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 259 |     {assume ab: "b + a - b < n"
 | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 260 | have th0: "b + a - b = b + (a - b)" using lt by arith | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 261 | from add[rule_format, OF H(1)[rule_format, OF ab th0]] | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 262 | have "P b a" by (simp add: th0[symmetric]) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 263 | hence "P a b" using c by blast } | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 264 | ultimately have "P a b" by blast} | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 265 | ultimately show "P a b" by blast | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 266 | qed | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 267 | |
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 268 | lemma bezout_lemma: | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 269 | assumes ex: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x = b * y + d \<or> b * x = a * y + d)" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 270 | shows "\<exists>d x y. d dvd a \<and> d dvd a + b \<and> (a * x = (a + b) * y + d \<or> (a + b) * x = a * y + d)" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 271 | using ex | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 272 | apply clarsimp | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 273 | apply (rule_tac x="d" in exI, simp add: dvd_add) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 274 | apply (case_tac "a * x = b * y + d" , simp_all) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 275 | apply (rule_tac x="x + y" in exI) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 276 | apply (rule_tac x="y" in exI) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 277 | apply algebra | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 278 | apply (rule_tac x="x" in exI) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 279 | apply (rule_tac x="x + y" in exI) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 280 | apply algebra | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 281 | done | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 282 | |
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 283 | lemma bezout_add: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x = b * y + d \<or> b * x = a * y + d)" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 284 | apply(induct a b rule: ind_euclid) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 285 | apply blast | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 286 | apply clarify | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 287 | apply (rule_tac x="a" in exI, simp add: dvd_add) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 288 | apply clarsimp | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 289 | apply (rule_tac x="d" in exI) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 290 | apply (case_tac "a * x = b * y + d", simp_all add: dvd_add) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 291 | apply (rule_tac x="x+y" in exI) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 292 | apply (rule_tac x="y" in exI) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 293 | apply algebra | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 294 | apply (rule_tac x="x" in exI) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 295 | apply (rule_tac x="x+y" in exI) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 296 | apply algebra | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 297 | done | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 298 | |
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 299 | lemma bezout: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x - b * y = d \<or> b * x - a * y = d)" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 300 | using bezout_add[of a b] | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 301 | apply clarsimp | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 302 | apply (rule_tac x="d" in exI, simp) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 303 | apply (rule_tac x="x" in exI) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 304 | apply (rule_tac x="y" in exI) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 305 | apply auto | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 306 | done | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 307 | |
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 308 | |
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 309 | text {* We can get a stronger version with a nonzeroness assumption. *}
 | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 310 | lemma divides_le: "m dvd n ==> m <= n \<or> n = (0::nat)" by (auto simp add: dvd_def) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 311 | |
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 312 | lemma bezout_add_strong: assumes nz: "a \<noteq> (0::nat)" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 313 | shows "\<exists>d x y. d dvd a \<and> d dvd b \<and> a * x = b * y + d" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 314 | proof- | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 315 | from nz have ap: "a > 0" by simp | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 316 | from bezout_add[of a b] | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 317 | have "(\<exists>d x y. d dvd a \<and> d dvd b \<and> a * x = b * y + d) \<or> (\<exists>d x y. d dvd a \<and> d dvd b \<and> b * x = a * y + d)" by blast | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 318 | moreover | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 319 |  {fix d x y assume H: "d dvd a" "d dvd b" "a * x = b * y + d"
 | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 320 | from H have ?thesis by blast } | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 321 | moreover | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 322 |  {fix d x y assume H: "d dvd a" "d dvd b" "b * x = a * y + d"
 | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 323 |    {assume b0: "b = 0" with H  have ?thesis by simp}
 | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 324 | moreover | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 325 |    {assume b: "b \<noteq> 0" hence bp: "b > 0" by simp
 | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 326 | from divides_le[OF H(2)] b have "d < b \<or> d = b" using le_less by blast | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 327 | moreover | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 328 |      {assume db: "d=b"
 | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 329 | from prems have ?thesis apply simp | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 330 | apply (rule exI[where x = b], simp) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 331 | apply (rule exI[where x = b]) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 332 | by (rule exI[where x = "a - 1"], simp add: diff_mult_distrib2)} | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 333 | moreover | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 334 |     {assume db: "d < b" 
 | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 335 | 	{assume "x=0" hence ?thesis  using prems by simp }
 | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 336 | moreover | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 337 | 	{assume x0: "x \<noteq> 0" hence xp: "x > 0" by simp
 | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 338 | |
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 339 | from db have "d \<le> b - 1" by simp | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 340 | hence "d*b \<le> b*(b - 1)" by simp | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 341 | with xp mult_mono[of "1" "x" "d*b" "b*(b - 1)"] | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 342 | have dble: "d*b \<le> x*b*(b - 1)" using bp by simp | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 343 | from H (3) have "a * ((b - 1) * y) + d * (b - 1 + 1) = d + x*b*(b - 1)" by algebra | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 344 | hence "a * ((b - 1) * y) = d + x*b*(b - 1) - d*b" using bp by simp | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 345 | hence "a * ((b - 1) * y) = d + (x*b*(b - 1) - d*b)" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 346 | by (simp only: diff_add_assoc[OF dble, of d, symmetric]) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 347 | hence "a * ((b - 1) * y) = b*(x*(b - 1) - d) + d" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 348 | by (simp only: diff_mult_distrib2 add_commute mult_ac) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 349 | hence ?thesis using H(1,2) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 350 | apply - | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 351 | apply (rule exI[where x=d], simp) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 352 | apply (rule exI[where x="(b - 1) * y"]) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 353 | by (rule exI[where x="x*(b - 1) - d"], simp)} | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 354 | ultimately have ?thesis by blast} | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 355 | ultimately have ?thesis by blast} | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 356 | ultimately have ?thesis by blast} | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 357 | ultimately show ?thesis by blast | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 358 | qed | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 359 | |
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 360 | |
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 361 | lemma bezout_gcd: "\<exists>x y. a * x - b * y = gcd a b \<or> b * x - a * y = gcd a b" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 362 | proof- | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 363 | let ?g = "gcd a b" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 364 | from bezout[of a b] obtain d x y where d: "d dvd a" "d dvd b" "a * x - b * y = d \<or> b * x - a * y = d" by blast | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 365 | from d(1,2) have "d dvd ?g" by simp | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 366 | then obtain k where k: "?g = d*k" unfolding dvd_def by blast | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 367 | from d(3) have "(a * x - b * y)*k = d*k \<or> (b * x - a * y)*k = d*k" by blast | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 368 | hence "a * x * k - b * y*k = d*k \<or> b * x * k - a * y*k = d*k" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 369 | by (algebra add: diff_mult_distrib) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 370 | hence "a * (x * k) - b * (y*k) = ?g \<or> b * (x * k) - a * (y*k) = ?g" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 371 | by (simp add: k mult_assoc) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 372 | thus ?thesis by blast | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 373 | qed | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 374 | |
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 375 | lemma bezout_gcd_strong: assumes a: "a \<noteq> 0" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 376 | shows "\<exists>x y. a * x = b * y + gcd a b" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 377 | proof- | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 378 | let ?g = "gcd a b" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 379 | from bezout_add_strong[OF a, of b] | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 380 | obtain d x y where d: "d dvd a" "d dvd b" "a * x = b * y + d" by blast | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 381 | from d(1,2) have "d dvd ?g" by simp | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 382 | then obtain k where k: "?g = d*k" unfolding dvd_def by blast | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 383 | from d(3) have "a * x * k = (b * y + d) *k " by algebra | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 384 | hence "a * (x * k) = b * (y*k) + ?g" by (algebra add: k) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 385 | thus ?thesis by blast | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 386 | qed | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 387 | |
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 388 | lemma gcd_mult_distrib: "gcd(a * c) (b * c) = c * gcd a b" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 389 | by(simp add: gcd_mult_distrib2 mult_commute) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 390 | |
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 391 | lemma gcd_bezout: "(\<exists>x y. a * x - b * y = d \<or> b * x - a * y = d) \<longleftrightarrow> gcd a b dvd d" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 392 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 393 | proof- | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 394 | let ?g = "gcd a b" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 395 |   {assume H: ?rhs then obtain k where k: "d = ?g*k" unfolding dvd_def by blast
 | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 396 | from bezout_gcd[of a b] obtain x y where xy: "a * x - b * y = ?g \<or> b * x - a * y = ?g" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 397 | by blast | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 398 | hence "(a * x - b * y)*k = ?g*k \<or> (b * x - a * y)*k = ?g*k" by auto | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 399 | hence "a * x*k - b * y*k = ?g*k \<or> b * x * k - a * y*k = ?g*k" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 400 | by (simp only: diff_mult_distrib) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 401 | hence "a * (x*k) - b * (y*k) = d \<or> b * (x * k) - a * (y*k) = d" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 402 | by (simp add: k[symmetric] mult_assoc) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 403 | hence ?lhs by blast} | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 404 | moreover | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 405 |   {fix x y assume H: "a * x - b * y = d \<or> b * x - a * y = d"
 | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 406 | have dv: "?g dvd a*x" "?g dvd b * y" "?g dvd b*x" "?g dvd a * y" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 407 | using dvd_mult2[OF gcd_dvd1[of a b]] dvd_mult2[OF gcd_dvd2[of a b]] by simp_all | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 408 | from dvd_diff[OF dv(1,2)] dvd_diff[OF dv(3,4)] H | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 409 | have ?rhs by auto} | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 410 | ultimately show ?thesis by blast | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 411 | qed | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 412 | |
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 413 | lemma gcd_bezout_sum: assumes H:"a * x + b * y = d" shows "gcd a b dvd d" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 414 | proof- | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 415 | let ?g = "gcd a b" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 416 | have dv: "?g dvd a*x" "?g dvd b * y" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 417 | using dvd_mult2[OF gcd_dvd1[of a b]] dvd_mult2[OF gcd_dvd2[of a b]] by simp_all | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 418 | from dvd_add[OF dv] H | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 419 | show ?thesis by auto | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 420 | qed | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 421 | |
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 422 | lemma gcd_mult': "gcd b (a * b) = b" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 423 | by (simp add: gcd_mult mult_commute[of a b]) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 424 | |
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 425 | lemma gcd_add: "gcd(a + b) b = gcd a b" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 426 | "gcd(b + a) b = gcd a b" "gcd a (a + b) = gcd a b" "gcd a (b + a) = gcd a b" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 427 | apply (simp_all add: gcd_add1) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 428 | by (simp add: gcd_commute gcd_add1) | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 429 | |
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 430 | lemma gcd_sub: "b <= a ==> gcd(a - b) b = gcd a b" "a <= b ==> gcd a (b - a) = gcd a b" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 431 | proof- | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 432 |   {fix a b assume H: "b \<le> (a::nat)"
 | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 433 | hence th: "a - b + b = a" by arith | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 434 | from gcd_add(1)[of "a - b" b] th have "gcd(a - b) b = gcd a b" by simp} | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 435 | note th = this | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 436 | {
 | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 437 | assume ab: "b \<le> a" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 438 | from th[OF ab] show "gcd (a - b) b = gcd a b" by blast | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 439 | next | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 440 | assume ab: "a \<le> b" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 441 | from th[OF ab] show "gcd a (b - a) = gcd a b" | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 442 | by (simp add: gcd_commute)} | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 443 | qed | 
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 444 | |
| 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 445 | |
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 446 | subsection {* LCM defined by GCD *}
 | 
| 22367 | 447 | |
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 448 | |
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 449 | definition | 
| 27556 | 450 | lcm :: "nat \<Rightarrow> nat \<Rightarrow> nat" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 451 | where | 
| 27568 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 452 | lcm_def: "lcm m n = m * n div gcd m n" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 453 | |
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 454 | lemma prod_gcd_lcm: | 
| 27556 | 455 | "m * n = gcd m n * lcm m n" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 456 | unfolding lcm_def by (simp add: dvd_mult_div_cancel [OF gcd_dvd_prod]) | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 457 | |
| 27556 | 458 | lemma lcm_0 [simp]: "lcm m 0 = 0" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 459 | unfolding lcm_def by simp | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 460 | |
| 27556 | 461 | lemma lcm_1 [simp]: "lcm m 1 = m" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 462 | unfolding lcm_def by simp | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 463 | |
| 27556 | 464 | lemma lcm_0_left [simp]: "lcm 0 n = 0" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 465 | unfolding lcm_def by simp | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 466 | |
| 27556 | 467 | lemma lcm_1_left [simp]: "lcm 1 m = m" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 468 | unfolding lcm_def by simp | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 469 | |
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 470 | lemma dvd_pos: | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 471 | fixes n m :: nat | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 472 | assumes "n > 0" and "m dvd n" | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 473 | shows "m > 0" | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 474 | using assms by (cases m) auto | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 475 | |
| 23951 | 476 | lemma lcm_least: | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 477 | assumes "m dvd k" and "n dvd k" | 
| 27556 | 478 | shows "lcm m n dvd k" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 479 | proof (cases k) | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 480 | case 0 then show ?thesis by auto | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 481 | next | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 482 | case (Suc _) then have pos_k: "k > 0" by auto | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 483 | from assms dvd_pos [OF this] have pos_mn: "m > 0" "n > 0" by auto | 
| 27556 | 484 | with gcd_zero [of m n] have pos_gcd: "gcd m n > 0" by simp | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 485 | from assms obtain p where k_m: "k = m * p" using dvd_def by blast | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 486 | from assms obtain q where k_n: "k = n * q" using dvd_def by blast | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 487 | from pos_k k_m have pos_p: "p > 0" by auto | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 488 | from pos_k k_n have pos_q: "q > 0" by auto | 
| 27556 | 489 | have "k * k * gcd q p = k * gcd (k * q) (k * p)" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 490 | by (simp add: mult_ac gcd_mult_distrib2) | 
| 27556 | 491 | also have "\<dots> = k * gcd (m * p * q) (n * q * p)" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 492 | by (simp add: k_m [symmetric] k_n [symmetric]) | 
| 27556 | 493 | also have "\<dots> = k * p * q * gcd m n" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 494 | by (simp add: mult_ac gcd_mult_distrib2) | 
| 27556 | 495 | finally have "(m * p) * (n * q) * gcd q p = k * p * q * gcd m n" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 496 | by (simp only: k_m [symmetric] k_n [symmetric]) | 
| 27556 | 497 | then have "p * q * m * n * gcd q p = p * q * k * gcd m n" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 498 | by (simp add: mult_ac) | 
| 27556 | 499 | with pos_p pos_q have "m * n * gcd q p = k * gcd m n" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 500 | by simp | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 501 | with prod_gcd_lcm [of m n] | 
| 27556 | 502 | have "lcm m n * gcd q p * gcd m n = k * gcd m n" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 503 | by (simp add: mult_ac) | 
| 27556 | 504 | with pos_gcd have "lcm m n * gcd q p = k" by simp | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 505 | then show ?thesis using dvd_def by auto | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 506 | qed | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 507 | |
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 508 | lemma lcm_dvd1 [iff]: | 
| 27556 | 509 | "m dvd lcm m n" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 510 | proof (cases m) | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 511 | case 0 then show ?thesis by simp | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 512 | next | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 513 | case (Suc _) | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 514 | then have mpos: "m > 0" by simp | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 515 | show ?thesis | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 516 | proof (cases n) | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 517 | case 0 then show ?thesis by simp | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 518 | next | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 519 | case (Suc _) | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 520 | then have npos: "n > 0" by simp | 
| 27556 | 521 | have "gcd m n dvd n" by simp | 
| 522 | then obtain k where "n = gcd m n * k" using dvd_def by auto | |
| 523 | then have "m * n div gcd m n = m * (gcd m n * k) div gcd m n" by (simp add: mult_ac) | |
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 524 | also have "\<dots> = m * k" using mpos npos gcd_zero by simp | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 525 | finally show ?thesis by (simp add: lcm_def) | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 526 | qed | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 527 | qed | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 528 | |
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 529 | lemma lcm_dvd2 [iff]: | 
| 27556 | 530 | "n dvd lcm m n" | 
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 531 | proof (cases n) | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 532 | case 0 then show ?thesis by simp | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 533 | next | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 534 | case (Suc _) | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 535 | then have npos: "n > 0" by simp | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 536 | show ?thesis | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 537 | proof (cases m) | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 538 | case 0 then show ?thesis by simp | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 539 | next | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 540 | case (Suc _) | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 541 | then have mpos: "m > 0" by simp | 
| 27556 | 542 | have "gcd m n dvd m" by simp | 
| 543 | then obtain k where "m = gcd m n * k" using dvd_def by auto | |
| 544 | then have "m * n div gcd m n = (gcd m n * k) * n div gcd m n" by (simp add: mult_ac) | |
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 545 | also have "\<dots> = n * k" using mpos npos gcd_zero by simp | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 546 | finally show ?thesis by (simp add: lcm_def) | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 547 | qed | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 548 | qed | 
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 549 | |
| 27568 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 550 | lemma gcd_add1_eq: "gcd (m + k) k = gcd (m + k) m" | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 551 | by (simp add: gcd_commute) | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 552 | |
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 553 | lemma gcd_diff2: "m \<le> n ==> gcd n (n - m) = gcd n m" | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 554 | apply (subgoal_tac "n = m + (n - m)") | 
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 555 | apply (erule ssubst, rule gcd_add1_eq, simp) | 
| 27568 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 556 | done | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 557 | |
| 23687 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 558 | |
| 
06884f7ffb18
extended - convers now basic lcm properties also
 haftmann parents: 
23431diff
changeset | 559 | subsection {* GCD and LCM on integers *}
 | 
| 22367 | 560 | |
| 561 | definition | |
| 27556 | 562 | zgcd :: "int \<Rightarrow> int \<Rightarrow> int" where | 
| 563 | "zgcd i j = int (gcd (nat (abs i)) (nat (abs j)))" | |
| 22367 | 564 | |
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 565 | lemma zgcd_zdvd1 [iff,simp, algebra]: "zgcd i j dvd i" | 
| 27556 | 566 | by (simp add: zgcd_def int_dvd_iff) | 
| 22027 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 567 | |
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 568 | lemma zgcd_zdvd2 [iff,simp, algebra]: "zgcd i j dvd j" | 
| 27556 | 569 | by (simp add: zgcd_def int_dvd_iff) | 
| 22027 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 570 | |
| 27556 | 571 | lemma zgcd_pos: "zgcd i j \<ge> 0" | 
| 572 | by (simp add: zgcd_def) | |
| 22367 | 573 | |
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 574 | lemma zgcd0 [simp,algebra]: "(zgcd i j = 0) = (i = 0 \<and> j = 0)" | 
| 27556 | 575 | by (simp add: zgcd_def gcd_zero) arith | 
| 22027 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 576 | |
| 27556 | 577 | lemma zgcd_commute: "zgcd i j = zgcd j i" | 
| 578 | unfolding zgcd_def by (simp add: gcd_commute) | |
| 22367 | 579 | |
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 580 | lemma zgcd_zminus [simp, algebra]: "zgcd (- i) j = zgcd i j" | 
| 27556 | 581 | unfolding zgcd_def by simp | 
| 22367 | 582 | |
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 583 | lemma zgcd_zminus2 [simp, algebra]: "zgcd i (- j) = zgcd i j" | 
| 27556 | 584 | unfolding zgcd_def by simp | 
| 22367 | 585 | |
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 586 | (* should be solved by algebra*) | 
| 27556 | 587 | lemma zrelprime_dvd_mult: "zgcd i j = 1 \<Longrightarrow> i dvd k * j \<Longrightarrow> i dvd k" | 
| 588 | unfolding zgcd_def | |
| 22367 | 589 | proof - | 
| 27556 | 590 | assume "int (gcd (nat \<bar>i\<bar>) (nat \<bar>j\<bar>)) = 1" "i dvd k * j" | 
| 591 | then have g: "gcd (nat \<bar>i\<bar>) (nat \<bar>j\<bar>) = 1" by simp | |
| 22367 | 592 | from `i dvd k * j` obtain h where h: "k*j = i*h" unfolding dvd_def by blast | 
| 22027 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 593 | have th: "nat \<bar>i\<bar> dvd nat \<bar>k\<bar> * nat \<bar>j\<bar>" | 
| 22367 | 594 | unfolding dvd_def | 
| 595 | by (rule_tac x= "nat \<bar>h\<bar>" in exI, simp add: h nat_abs_mult_distrib [symmetric]) | |
| 596 | from relprime_dvd_mult [OF g th] obtain h' where h': "nat \<bar>k\<bar> = nat \<bar>i\<bar> * h'" | |
| 22027 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 597 | unfolding dvd_def by blast | 
| 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 598 | from h' have "int (nat \<bar>k\<bar>) = int (nat \<bar>i\<bar> * h')" by simp | 
| 23431 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23365diff
changeset | 599 | then have "\<bar>k\<bar> = \<bar>i\<bar> * int h'" by (simp add: int_mult) | 
| 22027 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 600 | then show ?thesis | 
| 22367 | 601 | apply (subst zdvd_abs1 [symmetric]) | 
| 602 | apply (subst zdvd_abs2 [symmetric]) | |
| 22027 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 603 | apply (unfold dvd_def) | 
| 22367 | 604 | apply (rule_tac x = "int h'" in exI, simp) | 
| 22027 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 605 | done | 
| 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 606 | qed | 
| 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 607 | |
| 27556 | 608 | lemma int_nat_abs: "int (nat (abs x)) = abs x" by arith | 
| 22367 | 609 | |
| 27556 | 610 | lemma zgcd_greatest: | 
| 22367 | 611 | assumes "k dvd m" and "k dvd n" | 
| 27556 | 612 | shows "k dvd zgcd m n" | 
| 22367 | 613 | proof - | 
| 22027 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 614 | let ?k' = "nat \<bar>k\<bar>" | 
| 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 615 | let ?m' = "nat \<bar>m\<bar>" | 
| 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 616 | let ?n' = "nat \<bar>n\<bar>" | 
| 22367 | 617 | from `k dvd m` and `k dvd n` have dvd': "?k' dvd ?m'" "?k' dvd ?n'" | 
| 22027 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 618 | unfolding zdvd_int by (simp_all only: int_nat_abs zdvd_abs1 zdvd_abs2) | 
| 27556 | 619 | from gcd_greatest [OF dvd'] have "int (nat \<bar>k\<bar>) dvd zgcd m n" | 
| 620 | unfolding zgcd_def by (simp only: zdvd_int) | |
| 621 | then have "\<bar>k\<bar> dvd zgcd m n" by (simp only: int_nat_abs) | |
| 622 | then show "k dvd zgcd m n" by (simp add: zdvd_abs1) | |
| 22027 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 623 | qed | 
| 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 624 | |
| 27556 | 625 | lemma div_zgcd_relprime: | 
| 22367 | 626 | assumes nz: "a \<noteq> 0 \<or> b \<noteq> 0" | 
| 27556 | 627 | shows "zgcd (a div (zgcd a b)) (b div (zgcd a b)) = 1" | 
| 22367 | 628 | proof - | 
| 25112 | 629 | from nz have nz': "nat \<bar>a\<bar> \<noteq> 0 \<or> nat \<bar>b\<bar> \<noteq> 0" by arith | 
| 27556 | 630 | let ?g = "zgcd a b" | 
| 22027 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 631 | let ?a' = "a div ?g" | 
| 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 632 | let ?b' = "b div ?g" | 
| 27556 | 633 | let ?g' = "zgcd ?a' ?b'" | 
| 27568 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 634 | have dvdg: "?g dvd a" "?g dvd b" by (simp_all add: zgcd_zdvd1 zgcd_zdvd2) | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 635 | have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by (simp_all add: zgcd_zdvd1 zgcd_zdvd2) | 
| 22367 | 636 | from dvdg dvdg' obtain ka kb ka' kb' where | 
| 637 | kab: "a = ?g*ka" "b = ?g*kb" "?a' = ?g'*ka'" "?b' = ?g' * kb'" | |
| 22027 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 638 | unfolding dvd_def by blast | 
| 22367 | 639 | then have "?g* ?a' = (?g * ?g') * ka'" "?g* ?b' = (?g * ?g') * kb'" by simp_all | 
| 640 | then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b" | |
| 641 | by (auto simp add: zdvd_mult_div_cancel [OF dvdg(1)] | |
| 642 | zdvd_mult_div_cancel [OF dvdg(2)] dvd_def) | |
| 22027 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 643 | have "?g \<noteq> 0" using nz by simp | 
| 27556 | 644 | then have gp: "?g \<noteq> 0" using zgcd_pos[where i="a" and j="b"] by arith | 
| 645 | from zgcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" . | |
| 22367 | 646 | with zdvd_mult_cancel1 [OF gp] have "\<bar>?g'\<bar> = 1" by simp | 
| 27556 | 647 | with zgcd_pos show "?g' = 1" by simp | 
| 22027 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 648 | qed | 
| 
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
 chaieb parents: 
21404diff
changeset | 649 | |
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 650 | lemma zgcd_0 [simp, algebra]: "zgcd m 0 = abs m" | 
| 27568 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 651 | by (simp add: zgcd_def abs_if) | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 652 | |
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 653 | lemma zgcd_0_left [simp, algebra]: "zgcd 0 m = abs m" | 
| 27568 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 654 | by (simp add: zgcd_def abs_if) | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 655 | |
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 656 | lemma zgcd_non_0: "0 < n ==> zgcd m n = zgcd n (m mod n)" | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 657 | apply (frule_tac b = n and a = m in pos_mod_sign) | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 658 | apply (simp del: pos_mod_sign add: zgcd_def abs_if nat_mod_distrib) | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 659 | apply (auto simp add: gcd_non_0 nat_mod_distrib [symmetric] zmod_zminus1_eq_if) | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 660 | apply (frule_tac a = m in pos_mod_bound) | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 661 | apply (simp del: pos_mod_bound add: nat_diff_distrib gcd_diff2 nat_le_eq_zle) | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 662 | done | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 663 | |
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 664 | lemma zgcd_eq: "zgcd m n = zgcd n (m mod n)" | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 665 | apply (case_tac "n = 0", simp add: DIVISION_BY_ZERO) | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 666 | apply (auto simp add: linorder_neq_iff zgcd_non_0) | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 667 | apply (cut_tac m = "-m" and n = "-n" in zgcd_non_0, auto) | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 668 | done | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 669 | |
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 670 | lemma zgcd_1 [simp, algebra]: "zgcd m 1 = 1" | 
| 27568 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 671 | by (simp add: zgcd_def abs_if) | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 672 | |
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 673 | lemma zgcd_0_1_iff [simp, algebra]: "zgcd 0 m = 1 \<longleftrightarrow> \<bar>m\<bar> = 1" | 
| 27568 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 674 | by (simp add: zgcd_def abs_if) | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 675 | |
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 676 | lemma zgcd_greatest_iff[algebra]: "k dvd zgcd m n = (k dvd m \<and> k dvd n)" | 
| 27568 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 677 | by (simp add: zgcd_def abs_if int_dvd_iff dvd_int_iff nat_dvd_iff) | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 678 | |
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 679 | lemma zgcd_1_left [simp, algebra]: "zgcd 1 m = 1" | 
| 27568 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 680 | by (simp add: zgcd_def gcd_1_left) | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 681 | |
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 682 | lemma zgcd_assoc: "zgcd (zgcd k m) n = zgcd k (zgcd m n)" | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 683 | by (simp add: zgcd_def gcd_assoc) | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 684 | |
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 685 | lemma zgcd_left_commute: "zgcd k (zgcd m n) = zgcd m (zgcd k n)" | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 686 | apply (rule zgcd_commute [THEN trans]) | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 687 | apply (rule zgcd_assoc [THEN trans]) | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 688 | apply (rule zgcd_commute [THEN arg_cong]) | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 689 | done | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 690 | |
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 691 | lemmas zgcd_ac = zgcd_assoc zgcd_commute zgcd_left_commute | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 692 |   -- {* addition is an AC-operator *}
 | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 693 | |
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 694 | lemma zgcd_zmult_distrib2: "0 \<le> k ==> k * zgcd m n = zgcd (k * m) (k * n)" | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 695 | by (simp del: minus_mult_right [symmetric] | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 696 | add: minus_mult_right nat_mult_distrib zgcd_def abs_if | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 697 | mult_less_0_iff gcd_mult_distrib2 [symmetric] zmult_int [symmetric]) | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 698 | |
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 699 | lemma zgcd_zmult_distrib2_abs: "zgcd (k * m) (k * n) = abs k * zgcd m n" | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 700 | by (simp add: abs_if zgcd_zmult_distrib2) | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 701 | |
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 702 | lemma zgcd_self [simp]: "0 \<le> m ==> zgcd m m = m" | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 703 | by (cut_tac k = m and m = 1 and n = 1 in zgcd_zmult_distrib2, simp_all) | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 704 | |
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 705 | lemma zgcd_zmult_eq_self [simp]: "0 \<le> k ==> zgcd k (k * n) = k" | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 706 | by (cut_tac k = k and m = 1 and n = n in zgcd_zmult_distrib2, simp_all) | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 707 | |
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 708 | lemma zgcd_zmult_eq_self2 [simp]: "0 \<le> k ==> zgcd (k * n) k = k" | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 709 | by (cut_tac k = k and m = n and n = 1 in zgcd_zmult_distrib2, simp_all) | 
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 710 | |
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 711 | |
| 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 712 | definition "zlcm i j = int (lcm(nat(abs i)) (nat(abs j)))" | 
| 23244 
1630951f0512
added lcm, ilcm (lcm for integers) and some lemmas about them;
 chaieb parents: 
22367diff
changeset | 713 | |
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 714 | lemma dvd_zlcm_self1[simp, algebra]: "i dvd zlcm i j" | 
| 27556 | 715 | by(simp add:zlcm_def dvd_int_iff) | 
| 23983 | 716 | |
| 27669 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 chaieb parents: 
27651diff
changeset | 717 | lemma dvd_zlcm_self2[simp, algebra]: "j dvd zlcm i j" | 
| 27556 | 718 | by(simp add:zlcm_def dvd_int_iff) | 
| 23983 | 719 | |
| 23244 
1630951f0512
added lcm, ilcm (lcm for integers) and some lemmas about them;
 chaieb parents: 
22367diff
changeset | 720 | |
| 27556 | 721 | lemma dvd_imp_dvd_zlcm1: | 
| 722 | assumes "k dvd i" shows "k dvd (zlcm i j)" | |
| 23983 | 723 | proof - | 
| 724 | have "nat(abs k) dvd nat(abs i)" using `k dvd i` | |
| 23994 | 725 | by(simp add:int_dvd_iff[symmetric] dvd_int_iff[symmetric] zdvd_abs1) | 
| 27556 | 726 | thus ?thesis by(simp add:zlcm_def dvd_int_iff)(blast intro: dvd_trans) | 
| 23983 | 727 | qed | 
| 728 | ||
| 27556 | 729 | lemma dvd_imp_dvd_zlcm2: | 
| 730 | assumes "k dvd j" shows "k dvd (zlcm i j)" | |
| 23983 | 731 | proof - | 
| 732 | have "nat(abs k) dvd nat(abs j)" using `k dvd j` | |
| 23994 | 733 | by(simp add:int_dvd_iff[symmetric] dvd_int_iff[symmetric] zdvd_abs1) | 
| 27556 | 734 | thus ?thesis by(simp add:zlcm_def dvd_int_iff)(blast intro: dvd_trans) | 
| 23983 | 735 | qed | 
| 736 | ||
| 23994 | 737 | |
| 23244 
1630951f0512
added lcm, ilcm (lcm for integers) and some lemmas about them;
 chaieb parents: 
22367diff
changeset | 738 | lemma zdvd_self_abs1: "(d::int) dvd (abs d)" | 
| 
1630951f0512
added lcm, ilcm (lcm for integers) and some lemmas about them;
 chaieb parents: 
22367diff
changeset | 739 | by (case_tac "d <0", simp_all) | 
| 
1630951f0512
added lcm, ilcm (lcm for integers) and some lemmas about them;
 chaieb parents: 
22367diff
changeset | 740 | |
| 
1630951f0512
added lcm, ilcm (lcm for integers) and some lemmas about them;
 chaieb parents: 
22367diff
changeset | 741 | lemma zdvd_self_abs2: "(abs (d::int)) dvd d" | 
| 
1630951f0512
added lcm, ilcm (lcm for integers) and some lemmas about them;
 chaieb parents: 
22367diff
changeset | 742 | by (case_tac "d<0", simp_all) | 
| 
1630951f0512
added lcm, ilcm (lcm for integers) and some lemmas about them;
 chaieb parents: 
22367diff
changeset | 743 | |
| 
1630951f0512
added lcm, ilcm (lcm for integers) and some lemmas about them;
 chaieb parents: 
22367diff
changeset | 744 | (* lcm a b is positive for positive a and b *) | 
| 
1630951f0512
added lcm, ilcm (lcm for integers) and some lemmas about them;
 chaieb parents: 
22367diff
changeset | 745 | |
| 
1630951f0512
added lcm, ilcm (lcm for integers) and some lemmas about them;
 chaieb parents: 
22367diff
changeset | 746 | lemma lcm_pos: | 
| 
1630951f0512
added lcm, ilcm (lcm for integers) and some lemmas about them;
 chaieb parents: 
22367diff
changeset | 747 | assumes mpos: "m > 0" | 
| 27568 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 748 | and npos: "n>0" | 
| 27556 | 749 | shows "lcm m n > 0" | 
| 23244 
1630951f0512
added lcm, ilcm (lcm for integers) and some lemmas about them;
 chaieb parents: 
22367diff
changeset | 750 | proof(rule ccontr, simp add: lcm_def gcd_zero) | 
| 27568 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 751 | assume h:"m*n div gcd m n = 0" | 
| 27556 | 752 | from mpos npos have "gcd m n \<noteq> 0" using gcd_zero by simp | 
| 753 | hence gcdp: "gcd m n > 0" by simp | |
| 23244 
1630951f0512
added lcm, ilcm (lcm for integers) and some lemmas about them;
 chaieb parents: 
22367diff
changeset | 754 | with h | 
| 27556 | 755 | have "m*n < gcd m n" | 
| 756 | by (cases "m * n < gcd m n") (auto simp add: div_if[OF gcdp, where m="m*n"]) | |
| 23244 
1630951f0512
added lcm, ilcm (lcm for integers) and some lemmas about them;
 chaieb parents: 
22367diff
changeset | 757 | moreover | 
| 27556 | 758 | have "gcd m n dvd m" by simp | 
| 759 | with mpos dvd_imp_le have t1:"gcd m n \<le> m" by simp | |
| 27568 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 chaieb parents: 
27556diff
changeset | 760 | with npos have t1:"gcd m n *n \<le> m*n" by simp | 
| 27556 | 761 | have "gcd m n \<le> gcd m n*n" using npos by simp | 
| 762 | with t1 have "gcd m n \<le> m*n" by arith | |
| 23244 
1630951f0512
added lcm, ilcm (lcm for integers) and some lemmas about them;
 chaieb parents: 
22367diff
changeset | 763 | ultimately show "False" by simp | 
| 
1630951f0512
added lcm, ilcm (lcm for integers) and some lemmas about them;
 chaieb parents: 
22367diff
changeset | 764 | qed | 
| 
1630951f0512
added lcm, ilcm (lcm for integers) and some lemmas about them;
 chaieb parents: 
22367diff
changeset | 765 | |
| 27556 | 766 | lemma zlcm_pos: | 
| 23983 | 767 | assumes anz: "a \<noteq> 0" | 
| 768 | and bnz: "b \<noteq> 0" | |
| 27556 | 769 | shows "0 < zlcm a b" | 
| 23244 
1630951f0512
added lcm, ilcm (lcm for integers) and some lemmas about them;
 chaieb parents: 
22367diff
changeset | 770 | proof- | 
| 
1630951f0512
added lcm, ilcm (lcm for integers) and some lemmas about them;
 chaieb parents: 
22367diff
changeset | 771 | let ?na = "nat (abs a)" | 
| 
1630951f0512
added lcm, ilcm (lcm for integers) and some lemmas about them;
 chaieb parents: 
22367diff
changeset | 772 | let ?nb = "nat (abs b)" | 
| 23983 | 773 | have nap: "?na >0" using anz by simp | 
| 774 | have nbp: "?nb >0" using bnz by simp | |
| 27556 | 775 | have "0 < lcm ?na ?nb" by (rule lcm_pos[OF nap nbp]) | 
| 776 | thus ?thesis by (simp add: zlcm_def) | |
| 23244 
1630951f0512
added lcm, ilcm (lcm for integers) and some lemmas about them;
 chaieb parents: 
22367diff
changeset | 777 | qed | 
| 
1630951f0512
added lcm, ilcm (lcm for integers) and some lemmas about them;
 chaieb parents: 
22367diff
changeset | 778 | |
| 28562 | 779 | lemma zgcd_code [code]: | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27568diff
changeset | 780 | "zgcd k l = \<bar>if l = 0 then k else zgcd l (\<bar>k\<bar> mod \<bar>l\<bar>)\<bar>" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27568diff
changeset | 781 | by (simp add: zgcd_def gcd.simps [of "nat \<bar>k\<bar>"] nat_mod_distrib) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27568diff
changeset | 782 | |
| 21256 | 783 | end |