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