| author | haftmann | 
| Fri, 18 Jul 2008 18:25:53 +0200 | |
| changeset 27651 | 16a26996c30e | 
| parent 27567 | e3fe9a327c63 | 
| child 27670 | 3b5425dead98 | 
| permissions | -rw-r--r-- | 
| 11368 | 1  | 
(* Title: HOL/Library/Primes.thy  | 
| 11363 | 2  | 
ID: $Id$  | 
| 27106 | 3  | 
Author: Amine Chaieb, Christophe Tabacznyj and Lawrence C Paulson  | 
| 11363 | 4  | 
Copyright 1996 University of Cambridge  | 
5  | 
*)  | 
|
6  | 
||
| 16762 | 7  | 
header {* Primality on nat *}
 | 
| 11363 | 8  | 
|
| 15131 | 9  | 
theory Primes  | 
| 27487 | 10  | 
imports Plain "~~/src/HOL/ATP_Linkup" GCD Parity  | 
| 15131 | 11  | 
begin  | 
| 11363 | 12  | 
|
| 19086 | 13  | 
definition  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21256 
diff
changeset
 | 
14  | 
coprime :: "nat => nat => bool" where  | 
| 27567 | 15  | 
"coprime m n \<longleftrightarrow> gcd m n = 1"  | 
| 11363 | 16  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21256 
diff
changeset
 | 
17  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21256 
diff
changeset
 | 
18  | 
prime :: "nat \<Rightarrow> bool" where  | 
| 27106 | 19  | 
[code func del]: "prime p \<longleftrightarrow> (1 < p \<and> (\<forall>m. m dvd p --> m = 1 \<or> m = p))"  | 
| 11363 | 20  | 
|
21  | 
||
| 16762 | 22  | 
lemma two_is_prime: "prime 2"  | 
23  | 
apply (auto simp add: prime_def)  | 
|
24  | 
apply (case_tac m)  | 
|
25  | 
apply (auto dest!: dvd_imp_le)  | 
|
| 11363 | 26  | 
done  | 
27  | 
||
| 27556 | 28  | 
lemma prime_imp_relprime: "prime p ==> \<not> p dvd n ==> gcd p n = 1"  | 
| 11363 | 29  | 
apply (auto simp add: prime_def)  | 
| 23839 | 30  | 
apply (metis One_nat_def gcd_dvd1 gcd_dvd2)  | 
| 11363 | 31  | 
done  | 
32  | 
||
33  | 
text {*
 | 
|
34  | 
This theorem leads immediately to a proof of the uniqueness of  | 
|
35  | 
  factorization.  If @{term p} divides a product of primes then it is
 | 
|
36  | 
one of those primes.  | 
|
37  | 
*}  | 
|
38  | 
||
| 16663 | 39  | 
lemma prime_dvd_mult: "prime p ==> p dvd m * n ==> p dvd m \<or> p dvd n"  | 
| 12739 | 40  | 
by (blast intro: relprime_dvd_mult prime_imp_relprime)  | 
| 11363 | 41  | 
|
| 16663 | 42  | 
lemma prime_dvd_square: "prime p ==> p dvd m^Suc (Suc 0) ==> p dvd m"  | 
| 12739 | 43  | 
by (auto dest: prime_dvd_mult)  | 
44  | 
||
| 16663 | 45  | 
lemma prime_dvd_power_two: "prime p ==> p dvd m\<twosuperior> ==> p dvd m"  | 
| 
14353
 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 
paulson 
parents: 
13187 
diff
changeset
 | 
46  | 
by (rule prime_dvd_square) (simp_all add: power2_eq_square)  | 
| 11368 | 47  | 
|
| 26125 | 48  | 
|
49  | 
lemma exp_eq_1:"(x::nat)^n = 1 \<longleftrightarrow> x = 1 \<or> n = 0" by (induct n, auto)  | 
|
50  | 
lemma exp_mono_lt: "(x::nat) ^ (Suc n) < y ^ (Suc n) \<longleftrightarrow> x < y"  | 
|
51  | 
using power_less_imp_less_base[of x "Suc n" y] power_strict_mono[of x y "Suc n"]  | 
|
52  | 
by auto  | 
|
53  | 
lemma exp_mono_le: "(x::nat) ^ (Suc n) \<le> y ^ (Suc n) \<longleftrightarrow> x \<le> y"  | 
|
54  | 
by (simp only: linorder_not_less[symmetric] exp_mono_lt)  | 
|
55  | 
||
56  | 
lemma exp_mono_eq: "(x::nat) ^ Suc n = y ^ Suc n \<longleftrightarrow> x = y"  | 
|
57  | 
using power_inject_base[of x n y] by auto  | 
|
58  | 
||
59  | 
||
60  | 
lemma even_square: assumes e: "even (n::nat)" shows "\<exists>x. n ^ 2 = 4*x"  | 
|
61  | 
proof-  | 
|
62  | 
from e have "2 dvd n" by presburger  | 
|
63  | 
then obtain k where k: "n = 2*k" using dvd_def by auto  | 
|
64  | 
hence "n^2 = 4* (k^2)" by (simp add: power2_eq_square)  | 
|
65  | 
thus ?thesis by blast  | 
|
66  | 
qed  | 
|
67  | 
||
68  | 
lemma odd_square: assumes e: "odd (n::nat)" shows "\<exists>x. n ^ 2 = 4*x + 1"  | 
|
69  | 
proof-  | 
|
70  | 
from e have np: "n > 0" by presburger  | 
|
71  | 
from e have "2 dvd (n - 1)" by presburger  | 
|
72  | 
then obtain k where "n - 1 = 2*k" using dvd_def by auto  | 
|
73  | 
hence k: "n = 2*k + 1" using e by presburger  | 
|
74  | 
hence "n^2 = 4* (k^2 + k) + 1" by algebra  | 
|
75  | 
thus ?thesis by blast  | 
|
76  | 
qed  | 
|
77  | 
||
78  | 
lemma diff_square: "(x::nat)^2 - y^2 = (x+y)*(x - y)"  | 
|
79  | 
proof-  | 
|
80  | 
have "x \<le> y \<or> y \<le> x" by (rule nat_le_linear)  | 
|
81  | 
moreover  | 
|
82  | 
  {assume le: "x \<le> y"
 | 
|
83  | 
hence "x ^2 \<le> y^2" by (simp only: numeral_2_eq_2 exp_mono_le Let_def)  | 
|
84  | 
with le have ?thesis by simp }  | 
|
85  | 
moreover  | 
|
86  | 
  {assume le: "y \<le> x"
 | 
|
87  | 
hence le2: "y ^2 \<le> x^2" by (simp only: numeral_2_eq_2 exp_mono_le Let_def)  | 
|
88  | 
from le have "\<exists>z. y + z = x" by presburger  | 
|
89  | 
then obtain z where z: "x = y + z" by blast  | 
|
90  | 
from le2 have "\<exists>z. x^2 = y^2 + z" by presburger  | 
|
91  | 
then obtain z2 where z2: "x^2 = y^2 + z2" by blast  | 
|
92  | 
from z z2 have ?thesis apply simp by algebra }  | 
|
93  | 
ultimately show ?thesis by blast  | 
|
94  | 
qed  | 
|
95  | 
||
| 26144 | 96  | 
text {* Elementary theory of divisibility *}
 | 
| 26125 | 97  | 
lemma divides_ge: "(a::nat) dvd b \<Longrightarrow> b = 0 \<or> a \<le> b" unfolding dvd_def by auto  | 
98  | 
lemma divides_antisym: "(x::nat) dvd y \<and> y dvd x \<longleftrightarrow> x = y"  | 
|
99  | 
using dvd_anti_sym[of x y] by auto  | 
|
100  | 
||
101  | 
lemma divides_add_revr: assumes da: "(d::nat) dvd a" and dab:"d dvd (a + b)"  | 
|
102  | 
shows "d dvd b"  | 
|
103  | 
proof-  | 
|
104  | 
from da obtain k where k:"a = d*k" by (auto simp add: dvd_def)  | 
|
105  | 
from dab obtain k' where k': "a + b = d*k'" by (auto simp add: dvd_def)  | 
|
106  | 
from k k' have "b = d *(k' - k)" by (simp add : diff_mult_distrib2)  | 
|
107  | 
thus ?thesis unfolding dvd_def by blast  | 
|
108  | 
qed  | 
|
109  | 
||
110  | 
declare nat_mult_dvd_cancel_disj[presburger]  | 
|
111  | 
lemma nat_mult_dvd_cancel_disj'[presburger]:  | 
|
112  | 
"(m\<Colon>nat)*k dvd n*k \<longleftrightarrow> k = 0 \<or> m dvd n" unfolding mult_commute[of m k] mult_commute[of n k] by presburger  | 
|
113  | 
||
114  | 
lemma divides_mul_l: "(a::nat) dvd b ==> (c * a) dvd (c * b)"  | 
|
115  | 
by presburger  | 
|
116  | 
||
117  | 
lemma divides_mul_r: "(a::nat) dvd b ==> (a * c) dvd (b * c)" by presburger  | 
|
118  | 
lemma divides_cases: "(n::nat) dvd m ==> m = 0 \<or> m = n \<or> 2 * n <= m"  | 
|
119  | 
by (auto simp add: dvd_def)  | 
|
120  | 
lemma divides_le: "m dvd n ==> m <= n \<or> n = (0::nat)" by (auto simp add: dvd_def)  | 
|
121  | 
||
122  | 
lemma divides_div_not: "(x::nat) = (q * n) + r \<Longrightarrow> 0 < r \<Longrightarrow> r < n ==> ~(n dvd x)"  | 
|
123  | 
proof(auto simp add: dvd_def)  | 
|
124  | 
fix k assume H: "0 < r" "r < n" "q * n + r = n * k"  | 
|
125  | 
from H(3) have r: "r = n* (k -q)" by(simp add: diff_mult_distrib2 mult_commute)  | 
|
126  | 
  {assume "k - q = 0" with r H(1) have False by simp}
 | 
|
127  | 
moreover  | 
|
128  | 
  {assume "k - q \<noteq> 0" with r have "r \<ge> n" by auto
 | 
|
129  | 
with H(2) have False by simp}  | 
|
130  | 
ultimately show False by blast  | 
|
131  | 
qed  | 
|
132  | 
lemma divides_exp: "(x::nat) dvd y ==> x ^ n dvd y ^ n"  | 
|
133  | 
by (auto simp add: power_mult_distrib dvd_def)  | 
|
134  | 
||
135  | 
lemma divides_exp2: "n \<noteq> 0 \<Longrightarrow> (x::nat) ^ n dvd y \<Longrightarrow> x dvd y"  | 
|
136  | 
by (induct n ,auto simp add: dvd_def)  | 
|
137  | 
||
138  | 
fun fact :: "nat \<Rightarrow> nat" where  | 
|
139  | 
"fact 0 = 1"  | 
|
140  | 
| "fact (Suc n) = Suc n * fact n"  | 
|
141  | 
||
142  | 
lemma fact_lt: "0 < fact n" by(induct n, simp_all)  | 
|
143  | 
lemma fact_le: "fact n \<ge> 1" using fact_lt[of n] by simp  | 
|
144  | 
lemma fact_mono: assumes le: "m \<le> n" shows "fact m \<le> fact n"  | 
|
145  | 
proof-  | 
|
146  | 
from le have "\<exists>i. n = m+i" by presburger  | 
|
147  | 
then obtain i where i: "n = m+i" by blast  | 
|
148  | 
have "fact m \<le> fact (m + i)"  | 
|
149  | 
proof(induct m)  | 
|
150  | 
case 0 thus ?case using fact_le[of i] by simp  | 
|
151  | 
next  | 
|
152  | 
case (Suc m)  | 
|
153  | 
have "fact (Suc m) = Suc m * fact m" by simp  | 
|
154  | 
have th1: "Suc m \<le> Suc (m + i)" by simp  | 
|
155  | 
from mult_le_mono[of "Suc m" "Suc (m+i)" "fact m" "fact (m+i)", OF th1 Suc.hyps]  | 
|
156  | 
show ?case by simp  | 
|
157  | 
qed  | 
|
158  | 
thus ?thesis using i by simp  | 
|
159  | 
qed  | 
|
160  | 
||
161  | 
lemma divides_fact: "1 <= p \<Longrightarrow> p <= n ==> p dvd fact n"  | 
|
162  | 
proof(induct n arbitrary: p)  | 
|
163  | 
case 0 thus ?case by simp  | 
|
164  | 
next  | 
|
165  | 
case (Suc n p)  | 
|
166  | 
from Suc.prems have "p = Suc n \<or> p \<le> n" by presburger  | 
|
167  | 
moreover  | 
|
168  | 
  {assume "p = Suc n" hence ?case  by (simp only: fact.simps dvd_triv_left)}
 | 
|
169  | 
moreover  | 
|
170  | 
  {assume "p \<le> n"
 | 
|
171  | 
with Suc.prems(1) Suc.hyps have th: "p dvd fact n" by simp  | 
|
172  | 
from dvd_mult[OF th] have ?case by (simp only: fact.simps) }  | 
|
173  | 
ultimately show ?case by blast  | 
|
174  | 
qed  | 
|
175  | 
||
176  | 
declare dvd_triv_left[presburger]  | 
|
177  | 
declare dvd_triv_right[presburger]  | 
|
178  | 
lemma divides_rexp:  | 
|
179  | 
"x dvd y \<Longrightarrow> (x::nat) dvd (y^(Suc n))" by (simp add: dvd_mult2[of x y])  | 
|
180  | 
||
| 26144 | 181  | 
text {* The Bezout theorem is a bit ugly for N; it'd be easier for Z *}
 | 
| 26125 | 182  | 
lemma ind_euclid:  | 
183  | 
assumes c: " \<forall>a b. P (a::nat) b \<longleftrightarrow> P b a" and z: "\<forall>a. P a 0"  | 
|
184  | 
and add: "\<forall>a b. P a b \<longrightarrow> P a (a + b)"  | 
|
185  | 
shows "P a b"  | 
|
186  | 
proof(induct n\<equiv>"a+b" arbitrary: a b rule: nat_less_induct)  | 
|
187  | 
fix n a b  | 
|
188  | 
assume H: "\<forall>m < n. \<forall>a b. m = a + b \<longrightarrow> P a b" "n = a + b"  | 
|
189  | 
have "a = b \<or> a < b \<or> b < a" by arith  | 
|
190  | 
  moreover {assume eq: "a= b"
 | 
|
191  | 
from add[rule_format, OF z[rule_format, of a]] have "P a b" using eq by simp}  | 
|
192  | 
moreover  | 
|
193  | 
  {assume lt: "a < b"
 | 
|
194  | 
hence "a + b - a < n \<or> a = 0" using H(2) by arith  | 
|
195  | 
moreover  | 
|
196  | 
    {assume "a =0" with z c have "P a b" by blast }
 | 
|
197  | 
moreover  | 
|
198  | 
    {assume ab: "a + b - a < n"
 | 
|
199  | 
have th0: "a + b - a = a + (b - a)" using lt by arith  | 
|
200  | 
from add[rule_format, OF H(1)[rule_format, OF ab th0]]  | 
|
201  | 
have "P a b" by (simp add: th0[symmetric])}  | 
|
202  | 
ultimately have "P a b" by blast}  | 
|
203  | 
moreover  | 
|
204  | 
  {assume lt: "a > b"
 | 
|
205  | 
hence "b + a - b < n \<or> b = 0" using H(2) by arith  | 
|
206  | 
moreover  | 
|
207  | 
    {assume "b =0" with z c have "P a b" by blast }
 | 
|
208  | 
moreover  | 
|
209  | 
    {assume ab: "b + a - b < n"
 | 
|
210  | 
have th0: "b + a - b = b + (a - b)" using lt by arith  | 
|
211  | 
from add[rule_format, OF H(1)[rule_format, OF ab th0]]  | 
|
212  | 
have "P b a" by (simp add: th0[symmetric])  | 
|
213  | 
hence "P a b" using c by blast }  | 
|
214  | 
ultimately have "P a b" by blast}  | 
|
215  | 
ultimately show "P a b" by blast  | 
|
216  | 
qed  | 
|
217  | 
||
218  | 
lemma bezout_lemma:  | 
|
219  | 
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)"  | 
|
220  | 
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)"  | 
|
221  | 
using ex  | 
|
222  | 
apply clarsimp  | 
|
223  | 
apply (rule_tac x="d" in exI, simp add: dvd_add)  | 
|
224  | 
apply (case_tac "a * x = b * y + d" , simp_all)  | 
|
225  | 
apply (rule_tac x="x + y" in exI)  | 
|
226  | 
apply (rule_tac x="y" in exI)  | 
|
227  | 
apply algebra  | 
|
228  | 
apply (rule_tac x="x" in exI)  | 
|
229  | 
apply (rule_tac x="x + y" in exI)  | 
|
230  | 
apply algebra  | 
|
231  | 
done  | 
|
232  | 
||
233  | 
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)"  | 
|
234  | 
apply(induct a b rule: ind_euclid)  | 
|
235  | 
apply blast  | 
|
236  | 
apply clarify  | 
|
237  | 
apply (rule_tac x="a" in exI, simp add: dvd_add)  | 
|
238  | 
apply clarsimp  | 
|
239  | 
apply (rule_tac x="d" in exI)  | 
|
240  | 
apply (case_tac "a * x = b * y + d", simp_all add: dvd_add)  | 
|
241  | 
apply (rule_tac x="x+y" in exI)  | 
|
242  | 
apply (rule_tac x="y" in exI)  | 
|
243  | 
apply algebra  | 
|
244  | 
apply (rule_tac x="x" in exI)  | 
|
245  | 
apply (rule_tac x="x+y" in exI)  | 
|
246  | 
apply algebra  | 
|
247  | 
done  | 
|
248  | 
||
249  | 
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)"  | 
|
250  | 
using bezout_add[of a b]  | 
|
251  | 
apply clarsimp  | 
|
252  | 
apply (rule_tac x="d" in exI, simp)  | 
|
253  | 
apply (rule_tac x="x" in exI)  | 
|
254  | 
apply (rule_tac x="y" in exI)  | 
|
255  | 
apply auto  | 
|
256  | 
done  | 
|
257  | 
||
| 26144 | 258  | 
text {* We can get a stronger version with a nonzeroness assumption. *}
 | 
| 26125 | 259  | 
|
260  | 
lemma bezout_add_strong: assumes nz: "a \<noteq> (0::nat)"  | 
|
261  | 
shows "\<exists>d x y. d dvd a \<and> d dvd b \<and> a * x = b * y + d"  | 
|
262  | 
proof-  | 
|
263  | 
from nz have ap: "a > 0" by simp  | 
|
264  | 
from bezout_add[of a b]  | 
|
265  | 
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  | 
|
266  | 
moreover  | 
|
267  | 
 {fix d x y assume H: "d dvd a" "d dvd b" "a * x = b * y + d"
 | 
|
268  | 
from H have ?thesis by blast }  | 
|
269  | 
moreover  | 
|
270  | 
 {fix d x y assume H: "d dvd a" "d dvd b" "b * x = a * y + d"
 | 
|
271  | 
   {assume b0: "b = 0" with H  have ?thesis by simp}
 | 
|
272  | 
moreover  | 
|
273  | 
   {assume b: "b \<noteq> 0" hence bp: "b > 0" by simp
 | 
|
274  | 
from divides_le[OF H(2)] b have "d < b \<or> d = b" using le_less by blast  | 
|
275  | 
moreover  | 
|
276  | 
     {assume db: "d=b"
 | 
|
277  | 
from prems have ?thesis apply simp  | 
|
278  | 
apply (rule exI[where x = b], simp)  | 
|
279  | 
apply (rule exI[where x = b])  | 
|
280  | 
by (rule exI[where x = "a - 1"], simp add: diff_mult_distrib2)}  | 
|
281  | 
moreover  | 
|
282  | 
    {assume db: "d < b" 
 | 
|
283  | 
	{assume "x=0" hence ?thesis  using prems by simp }
 | 
|
284  | 
moreover  | 
|
285  | 
	{assume x0: "x \<noteq> 0" hence xp: "x > 0" by simp
 | 
|
286  | 
||
287  | 
from db have "d \<le> b - 1" by simp  | 
|
288  | 
hence "d*b \<le> b*(b - 1)" by simp  | 
|
289  | 
with xp mult_mono[of "1" "x" "d*b" "b*(b - 1)"]  | 
|
290  | 
have dble: "d*b \<le> x*b*(b - 1)" using bp by simp  | 
|
291  | 
from H (3) have "d + (b - 1) * (b*x) = d + (b - 1) * (a*y + d)" by simp  | 
|
292  | 
hence "d + (b - 1) * a * y + (b - 1) * d = d + (b - 1) * b * x"  | 
|
293  | 
by (simp only: mult_assoc right_distrib)  | 
|
294  | 
hence "a * ((b - 1) * y) + d * (b - 1 + 1) = d + x*b*(b - 1)" by algebra  | 
|
295  | 
hence "a * ((b - 1) * y) = d + x*b*(b - 1) - d*b" using bp by simp  | 
|
296  | 
hence "a * ((b - 1) * y) = d + (x*b*(b - 1) - d*b)"  | 
|
297  | 
by (simp only: diff_add_assoc[OF dble, of d, symmetric])  | 
|
298  | 
hence "a * ((b - 1) * y) = b*(x*(b - 1) - d) + d"  | 
|
299  | 
by (simp only: diff_mult_distrib2 add_commute mult_ac)  | 
|
300  | 
hence ?thesis using H(1,2)  | 
|
301  | 
apply -  | 
|
302  | 
apply (rule exI[where x=d], simp)  | 
|
303  | 
apply (rule exI[where x="(b - 1) * y"])  | 
|
304  | 
by (rule exI[where x="x*(b - 1) - d"], simp)}  | 
|
305  | 
ultimately have ?thesis by blast}  | 
|
306  | 
ultimately have ?thesis by blast}  | 
|
307  | 
ultimately have ?thesis by blast}  | 
|
308  | 
ultimately show ?thesis by blast  | 
|
309  | 
qed  | 
|
310  | 
||
| 26144 | 311  | 
text {* Greatest common divisor. *}
 | 
| 27556 | 312  | 
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"  | 
| 26125 | 313  | 
proof(auto)  | 
314  | 
assume H: "d dvd a" "d dvd b" "\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d"  | 
|
315  | 
from H(3)[rule_format] gcd_dvd1[of a b] gcd_dvd2[of a b]  | 
|
| 27556 | 316  | 
have th: "gcd a b dvd d" by blast  | 
| 27567 | 317  | 
from dvd_anti_sym[OF th gcd_greatest[OF H(1,2)]] show "d = gcd a b" by blast  | 
| 26125 | 318  | 
qed  | 
319  | 
||
320  | 
lemma gcd_eq: assumes H: "\<forall>d. d dvd x \<and> d dvd y \<longleftrightarrow> d dvd u \<and> d dvd v"  | 
|
| 27556 | 321  | 
shows "gcd x y = gcd u v"  | 
| 26125 | 322  | 
proof-  | 
| 27556 | 323  | 
from H have "\<forall>d. d dvd x \<and> d dvd y \<longleftrightarrow> d dvd gcd u v" by simp  | 
324  | 
with gcd_unique[of "gcd u v" x y] show ?thesis by auto  | 
|
| 26125 | 325  | 
qed  | 
326  | 
||
| 27556 | 327  | 
lemma bezout_gcd: "\<exists>x y. a * x - b * y = gcd a b \<or> b * x - a * y = gcd a b"  | 
| 26125 | 328  | 
proof-  | 
| 27556 | 329  | 
let ?g = "gcd a b"  | 
| 26125 | 330  | 
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  | 
331  | 
from d(1,2) have "d dvd ?g" by simp  | 
|
332  | 
then obtain k where k: "?g = d*k" unfolding dvd_def by blast  | 
|
333  | 
from d(3) have "(a * x - b * y)*k = d*k \<or> (b * x - a * y)*k = d*k" by blast  | 
|
334  | 
hence "a * x * k - b * y*k = d*k \<or> b * x * k - a * y*k = d*k"  | 
|
335  | 
by (simp only: diff_mult_distrib)  | 
|
336  | 
hence "a * (x * k) - b * (y*k) = ?g \<or> b * (x * k) - a * (y*k) = ?g"  | 
|
337  | 
by (simp add: k mult_assoc)  | 
|
338  | 
thus ?thesis by blast  | 
|
339  | 
qed  | 
|
340  | 
||
341  | 
lemma bezout_gcd_strong: assumes a: "a \<noteq> 0"  | 
|
| 27556 | 342  | 
shows "\<exists>x y. a * x = b * y + gcd a b"  | 
| 26125 | 343  | 
proof-  | 
| 27556 | 344  | 
let ?g = "gcd a b"  | 
| 26125 | 345  | 
from bezout_add_strong[OF a, of b]  | 
346  | 
obtain d x y where d: "d dvd a" "d dvd b" "a * x = b * y + d" by blast  | 
|
347  | 
from d(1,2) have "d dvd ?g" by simp  | 
|
348  | 
then obtain k where k: "?g = d*k" unfolding dvd_def by blast  | 
|
349  | 
from d(3) have "a * x * k = (b * y + d) *k " by auto  | 
|
350  | 
hence "a * (x * k) = b * (y*k) + ?g" by (algebra add: k)  | 
|
351  | 
thus ?thesis by blast  | 
|
352  | 
qed  | 
|
353  | 
||
| 27567 | 354  | 
lemma gcd_mult_distrib: "gcd(a * c) (b * c) = c * gcd a b"  | 
| 26125 | 355  | 
by(simp add: gcd_mult_distrib2 mult_commute)  | 
356  | 
||
| 27556 | 357  | 
lemma gcd_bezout: "(\<exists>x y. a * x - b * y = d \<or> b * x - a * y = d) \<longleftrightarrow> gcd a b dvd d"  | 
| 26125 | 358  | 
(is "?lhs \<longleftrightarrow> ?rhs")  | 
359  | 
proof-  | 
|
| 27556 | 360  | 
let ?g = "gcd a b"  | 
| 26125 | 361  | 
  {assume H: ?rhs then obtain k where k: "d = ?g*k" unfolding dvd_def by blast
 | 
362  | 
from bezout_gcd[of a b] obtain x y where xy: "a * x - b * y = ?g \<or> b * x - a * y = ?g"  | 
|
363  | 
by blast  | 
|
364  | 
hence "(a * x - b * y)*k = ?g*k \<or> (b * x - a * y)*k = ?g*k" by auto  | 
|
365  | 
hence "a * x*k - b * y*k = ?g*k \<or> b * x * k - a * y*k = ?g*k"  | 
|
366  | 
by (simp only: diff_mult_distrib)  | 
|
367  | 
hence "a * (x*k) - b * (y*k) = d \<or> b * (x * k) - a * (y*k) = d"  | 
|
368  | 
by (simp add: k[symmetric] mult_assoc)  | 
|
369  | 
hence ?lhs by blast}  | 
|
370  | 
moreover  | 
|
371  | 
  {fix x y assume H: "a * x - b * y = d \<or> b * x - a * y = d"
 | 
|
372  | 
have dv: "?g dvd a*x" "?g dvd b * y" "?g dvd b*x" "?g dvd a * y"  | 
|
373  | 
using dvd_mult2[OF gcd_dvd1[of a b]] dvd_mult2[OF gcd_dvd2[of a b]] by simp_all  | 
|
374  | 
from dvd_diff[OF dv(1,2)] dvd_diff[OF dv(3,4)] H  | 
|
375  | 
have ?rhs by auto}  | 
|
376  | 
ultimately show ?thesis by blast  | 
|
377  | 
qed  | 
|
378  | 
||
| 27556 | 379  | 
lemma gcd_bezout_sum: assumes H:"a * x + b * y = d" shows "gcd a b dvd d"  | 
| 26125 | 380  | 
proof-  | 
| 27556 | 381  | 
let ?g = "gcd a b"  | 
| 26125 | 382  | 
have dv: "?g dvd a*x" "?g dvd b * y"  | 
383  | 
using dvd_mult2[OF gcd_dvd1[of a b]] dvd_mult2[OF gcd_dvd2[of a b]] by simp_all  | 
|
384  | 
from dvd_add[OF dv] H  | 
|
385  | 
show ?thesis by auto  | 
|
386  | 
qed  | 
|
387  | 
||
| 27556 | 388  | 
lemma gcd_mult': "gcd b (a * b) = b"  | 
| 26125 | 389  | 
by (simp add: gcd_mult mult_commute[of a b])  | 
390  | 
||
| 27567 | 391  | 
lemma gcd_add: "gcd(a + b) b = gcd a b"  | 
392  | 
"gcd(b + a) b = gcd a b" "gcd a (a + b) = gcd a b" "gcd a (b + a) = gcd a b"  | 
|
| 26125 | 393  | 
apply (simp_all add: gcd_add1)  | 
394  | 
by (simp add: gcd_commute gcd_add1)  | 
|
395  | 
||
| 27567 | 396  | 
lemma gcd_sub: "b <= a ==> gcd(a - b) b = gcd a b" "a <= b ==> gcd a (b - a) = gcd a b"  | 
| 26125 | 397  | 
proof-  | 
398  | 
  {fix a b assume H: "b \<le> (a::nat)"
 | 
|
399  | 
hence th: "a - b + b = a" by arith  | 
|
| 27567 | 400  | 
from gcd_add(1)[of "a - b" b] th have "gcd(a - b) b = gcd a b" by simp}  | 
| 26125 | 401  | 
note th = this  | 
402  | 
{
 | 
|
403  | 
assume ab: "b \<le> a"  | 
|
| 27567 | 404  | 
from th[OF ab] show "gcd (a - b) b = gcd a b" by blast  | 
| 26125 | 405  | 
next  | 
406  | 
assume ab: "a \<le> b"  | 
|
| 27556 | 407  | 
from th[OF ab] show "gcd a (b - a) = gcd a b"  | 
| 26125 | 408  | 
by (simp add: gcd_commute)}  | 
409  | 
qed  | 
|
410  | 
||
| 26144 | 411  | 
text {* Coprimality *}
 | 
| 26125 | 412  | 
|
413  | 
lemma coprime: "coprime a b \<longleftrightarrow> (\<forall>d. d dvd a \<and> d dvd b \<longleftrightarrow> d = 1)"  | 
|
414  | 
using gcd_unique[of 1 a b, simplified] by (auto simp add: coprime_def)  | 
|
415  | 
lemma coprime_commute: "coprime a b \<longleftrightarrow> coprime b a" by (simp add: coprime_def gcd_commute)  | 
|
416  | 
||
417  | 
lemma coprime_bezout: "coprime a b \<longleftrightarrow> (\<exists>x y. a * x - b * y = 1 \<or> b * x - a * y = 1)"  | 
|
418  | 
using coprime_def gcd_bezout by auto  | 
|
419  | 
||
420  | 
lemma coprime_divprod: "d dvd a * b \<Longrightarrow> coprime d a \<Longrightarrow> d dvd b"  | 
|
421  | 
using relprime_dvd_mult_iff[of d a b] by (auto simp add: coprime_def mult_commute)  | 
|
422  | 
||
423  | 
lemma coprime_1[simp]: "coprime a 1" by (simp add: coprime_def)  | 
|
424  | 
lemma coprime_1'[simp]: "coprime 1 a" by (simp add: coprime_def)  | 
|
425  | 
lemma coprime_Suc0[simp]: "coprime a (Suc 0)" by (simp add: coprime_def)  | 
|
426  | 
lemma coprime_Suc0'[simp]: "coprime (Suc 0) a" by (simp add: coprime_def)  | 
|
427  | 
||
428  | 
lemma gcd_coprime:  | 
|
| 27556 | 429  | 
assumes z: "gcd a b \<noteq> 0" and a: "a = a' * gcd a b" and b: "b = b' * gcd a b"  | 
| 26125 | 430  | 
shows "coprime a' b'"  | 
431  | 
proof-  | 
|
| 27556 | 432  | 
let ?g = "gcd a b"  | 
| 26125 | 433  | 
  {assume bz: "a = 0" from b bz z a have ?thesis by (simp add: gcd_zero coprime_def)}
 | 
434  | 
moreover  | 
|
435  | 
  {assume az: "a\<noteq> 0" 
 | 
|
436  | 
from z have z': "?g > 0" by simp  | 
|
437  | 
from bezout_gcd_strong[OF az, of b]  | 
|
438  | 
obtain x y where xy: "a*x = b*y + ?g" by blast  | 
|
439  | 
from xy a b have "?g * a'*x = ?g * (b'*y + 1)" by (simp add: ring_simps)  | 
|
440  | 
hence "?g * (a'*x) = ?g * (b'*y + 1)" by (simp add: mult_assoc)  | 
|
441  | 
hence "a'*x = (b'*y + 1)"  | 
|
442  | 
by (simp only: nat_mult_eq_cancel1[OF z'])  | 
|
443  | 
hence "a'*x - b'*y = 1" by simp  | 
|
444  | 
with coprime_bezout[of a' b'] have ?thesis by auto}  | 
|
445  | 
ultimately show ?thesis by blast  | 
|
446  | 
qed  | 
|
447  | 
lemma coprime_0: "coprime d 0 \<longleftrightarrow> d = 1" by (simp add: coprime_def)  | 
|
448  | 
lemma coprime_mul: assumes da: "coprime d a" and db: "coprime d b"  | 
|
449  | 
shows "coprime d (a * b)"  | 
|
450  | 
proof-  | 
|
| 27556 | 451  | 
from da have th: "gcd a d = 1" by (simp add: coprime_def gcd_commute)  | 
| 27567 | 452  | 
from gcd_mult_cancel[of a d b, OF th] db[unfolded coprime_def] have "gcd d (a*b) = 1"  | 
| 26125 | 453  | 
by (simp add: gcd_commute)  | 
454  | 
thus ?thesis unfolding coprime_def .  | 
|
455  | 
qed  | 
|
456  | 
lemma coprime_lmul2: assumes dab: "coprime d (a * b)" shows "coprime d b"  | 
|
457  | 
using prems unfolding coprime_bezout  | 
|
458  | 
apply clarsimp  | 
|
459  | 
apply (case_tac "d * x - a * b * y = Suc 0 ", simp_all)  | 
|
460  | 
apply (rule_tac x="x" in exI)  | 
|
461  | 
apply (rule_tac x="a*y" in exI)  | 
|
462  | 
apply (simp add: mult_ac)  | 
|
463  | 
apply (rule_tac x="a*x" in exI)  | 
|
464  | 
apply (rule_tac x="y" in exI)  | 
|
465  | 
apply (simp add: mult_ac)  | 
|
466  | 
done  | 
|
467  | 
||
468  | 
lemma coprime_rmul2: "coprime d (a * b) \<Longrightarrow> coprime d a"  | 
|
469  | 
unfolding coprime_bezout  | 
|
470  | 
apply clarsimp  | 
|
471  | 
apply (case_tac "d * x - a * b * y = Suc 0 ", simp_all)  | 
|
472  | 
apply (rule_tac x="x" in exI)  | 
|
473  | 
apply (rule_tac x="b*y" in exI)  | 
|
474  | 
apply (simp add: mult_ac)  | 
|
475  | 
apply (rule_tac x="b*x" in exI)  | 
|
476  | 
apply (rule_tac x="y" in exI)  | 
|
477  | 
apply (simp add: mult_ac)  | 
|
478  | 
done  | 
|
479  | 
lemma coprime_mul_eq: "coprime d (a * b) \<longleftrightarrow> coprime d a \<and> coprime d b"  | 
|
480  | 
using coprime_rmul2[of d a b] coprime_lmul2[of d a b] coprime_mul[of d a b]  | 
|
481  | 
by blast  | 
|
482  | 
||
483  | 
lemma gcd_coprime_exists:  | 
|
| 27556 | 484  | 
assumes nz: "gcd a b \<noteq> 0"  | 
485  | 
shows "\<exists>a' b'. a = a' * gcd a b \<and> b = b' * gcd a b \<and> coprime a' b'"  | 
|
| 26125 | 486  | 
proof-  | 
| 27556 | 487  | 
let ?g = "gcd a b"  | 
| 26125 | 488  | 
from gcd_dvd1[of a b] gcd_dvd2[of a b]  | 
489  | 
obtain a' b' where "a = ?g*a'" "b = ?g*b'" unfolding dvd_def by blast  | 
|
490  | 
hence ab': "a = a'*?g" "b = b'*?g" by algebra+  | 
|
491  | 
from ab' gcd_coprime[OF nz ab'] show ?thesis by blast  | 
|
492  | 
qed  | 
|
493  | 
||
494  | 
lemma coprime_exp: "coprime d a ==> coprime d (a^n)"  | 
|
495  | 
by(induct n, simp_all add: coprime_mul)  | 
|
496  | 
||
497  | 
lemma coprime_exp_imp: "coprime a b ==> coprime (a ^n) (b ^n)"  | 
|
498  | 
by (induct n, simp_all add: coprime_mul_eq coprime_commute coprime_exp)  | 
|
499  | 
lemma coprime_refl[simp]: "coprime n n \<longleftrightarrow> n = 1" by (simp add: coprime_def)  | 
|
500  | 
lemma coprime_plus1[simp]: "coprime (n + 1) n"  | 
|
501  | 
apply (simp add: coprime_bezout)  | 
|
502  | 
apply (rule exI[where x=1])  | 
|
503  | 
apply (rule exI[where x=1])  | 
|
504  | 
apply simp  | 
|
505  | 
done  | 
|
506  | 
lemma coprime_minus1: "n \<noteq> 0 ==> coprime (n - 1) n"  | 
|
507  | 
using coprime_plus1[of "n - 1"] coprime_commute[of "n - 1" n] by auto  | 
|
508  | 
||
| 27556 | 509  | 
lemma bezout_gcd_pow: "\<exists>x y. a ^n * x - b ^ n * y = gcd a b ^ n \<or> b ^ n * x - a ^ n * y = gcd a b ^ n"  | 
| 26125 | 510  | 
proof-  | 
| 27556 | 511  | 
let ?g = "gcd a b"  | 
| 26125 | 512  | 
  {assume z: "?g = 0" hence ?thesis 
 | 
513  | 
apply (cases n, simp)  | 
|
514  | 
apply arith  | 
|
515  | 
apply (simp only: z power_0_Suc)  | 
|
516  | 
apply (rule exI[where x=0])  | 
|
517  | 
apply (rule exI[where x=0])  | 
|
518  | 
by simp}  | 
|
519  | 
moreover  | 
|
520  | 
  {assume z: "?g \<noteq> 0"
 | 
|
521  | 
from gcd_dvd1[of a b] gcd_dvd2[of a b] obtain a' b' where  | 
|
522  | 
ab': "a = a'*?g" "b = b'*?g" unfolding dvd_def by (auto simp add: mult_ac)  | 
|
523  | 
hence ab'': "?g*a' = a" "?g * b' = b" by algebra+  | 
|
524  | 
from coprime_exp_imp[OF gcd_coprime[OF z ab'], unfolded coprime_bezout, of n]  | 
|
525  | 
obtain x y where "a'^n * x - b'^n * y = 1 \<or> b'^n * x - a'^n * y = 1" by blast  | 
|
526  | 
hence "?g^n * (a'^n * x - b'^n * y) = ?g^n \<or> ?g^n*(b'^n * x - a'^n * y) = ?g^n"  | 
|
527  | 
using z by auto  | 
|
528  | 
then have "a^n * x - b^n * y = ?g^n \<or> b^n * x - a^n * y = ?g^n"  | 
|
529  | 
using z ab'' by (simp only: power_mult_distrib[symmetric]  | 
|
530  | 
diff_mult_distrib2 mult_assoc[symmetric])  | 
|
531  | 
hence ?thesis by blast }  | 
|
532  | 
ultimately show ?thesis by blast  | 
|
533  | 
qed  | 
|
| 27567 | 534  | 
|
535  | 
lemma gcd_exp: "gcd (a^n) (b^n) = gcd a b^n"  | 
|
| 26125 | 536  | 
proof-  | 
| 27556 | 537  | 
let ?g = "gcd (a^n) (b^n)"  | 
| 27567 | 538  | 
let ?gn = "gcd a b^n"  | 
| 26125 | 539  | 
  {fix e assume H: "e dvd a^n" "e dvd b^n"
 | 
540  | 
from bezout_gcd_pow[of a n b] obtain x y  | 
|
541  | 
where xy: "a ^ n * x - b ^ n * y = ?gn \<or> b ^ n * x - a ^ n * y = ?gn" by blast  | 
|
542  | 
from dvd_diff [OF dvd_mult2[OF H(1), of x] dvd_mult2[OF H(2), of y]]  | 
|
543  | 
dvd_diff [OF dvd_mult2[OF H(2), of x] dvd_mult2[OF H(1), of y]] xy  | 
|
| 27556 | 544  | 
have "e dvd ?gn" by (cases "a ^ n * x - b ^ n * y = gcd a b ^ n", simp_all)}  | 
| 26125 | 545  | 
hence th: "\<forall>e. e dvd a^n \<and> e dvd b^n \<longrightarrow> e dvd ?gn" by blast  | 
546  | 
from divides_exp[OF gcd_dvd1[of a b], of n] divides_exp[OF gcd_dvd2[of a b], of n] th  | 
|
547  | 
gcd_unique have "?gn = ?g" by blast thus ?thesis by simp  | 
|
548  | 
qed  | 
|
549  | 
||
550  | 
lemma coprime_exp2: "coprime (a ^ Suc n) (b^ Suc n) \<longleftrightarrow> coprime a b"  | 
|
551  | 
by (simp only: coprime_def gcd_exp exp_eq_1) simp  | 
|
552  | 
||
553  | 
lemma division_decomp: assumes dc: "(a::nat) dvd b * c"  | 
|
554  | 
shows "\<exists>b' c'. a = b' * c' \<and> b' dvd b \<and> c' dvd c"  | 
|
555  | 
proof-  | 
|
| 27556 | 556  | 
let ?g = "gcd a b"  | 
| 26125 | 557  | 
  {assume "?g = 0" with dc have ?thesis apply (simp add: gcd_zero)
 | 
558  | 
apply (rule exI[where x="0"])  | 
|
559  | 
by (rule exI[where x="c"], simp)}  | 
|
560  | 
moreover  | 
|
561  | 
  {assume z: "?g \<noteq> 0"
 | 
|
562  | 
from gcd_coprime_exists[OF z]  | 
|
563  | 
obtain a' b' where ab': "a = a' * ?g" "b = b' * ?g" "coprime a' b'" by blast  | 
|
564  | 
from gcd_dvd2[of a b] have thb: "?g dvd b" .  | 
|
565  | 
from ab'(1) have "a' dvd a" unfolding dvd_def by blast  | 
|
566  | 
with dc have th0: "a' dvd b*c" using dvd_trans[of a' a "b*c"] by simp  | 
|
567  | 
from dc ab'(1,2) have "a'*?g dvd (b'*?g) *c" by auto  | 
|
568  | 
hence "?g*a' dvd ?g * (b' * c)" by (simp add: mult_assoc)  | 
|
569  | 
with z have th_1: "a' dvd b'*c" by simp  | 
|
570  | 
from coprime_divprod[OF th_1 ab'(3)] have thc: "a' dvd c" .  | 
|
571  | 
from ab' have "a = ?g*a'" by algebra  | 
|
572  | 
with thb thc have ?thesis by blast }  | 
|
573  | 
ultimately show ?thesis by blast  | 
|
574  | 
qed  | 
|
575  | 
||
576  | 
lemma nat_power_eq_0_iff: "(m::nat) ^ n = 0 \<longleftrightarrow> n \<noteq> 0 \<and> m = 0" by (induct n, auto)  | 
|
577  | 
||
578  | 
lemma divides_rev: assumes ab: "(a::nat) ^ n dvd b ^n" and n:"n \<noteq> 0" shows "a dvd b"  | 
|
579  | 
proof-  | 
|
| 27556 | 580  | 
let ?g = "gcd a b"  | 
| 26125 | 581  | 
from n obtain m where m: "n = Suc m" by (cases n, simp_all)  | 
582  | 
  {assume "?g = 0" with ab n have ?thesis by (simp add: gcd_zero)}
 | 
|
583  | 
moreover  | 
|
584  | 
  {assume z: "?g \<noteq> 0"
 | 
|
585  | 
hence zn: "?g ^ n \<noteq> 0" using n by (simp add: neq0_conv)  | 
|
586  | 
from gcd_coprime_exists[OF z]  | 
|
587  | 
obtain a' b' where ab': "a = a' * ?g" "b = b' * ?g" "coprime a' b'" by blast  | 
|
588  | 
from ab have "(a' * ?g) ^ n dvd (b' * ?g)^n" by (simp add: ab'(1,2)[symmetric])  | 
|
589  | 
hence "?g^n*a'^n dvd ?g^n *b'^n" by (simp only: power_mult_distrib mult_commute)  | 
|
590  | 
with zn z n have th0:"a'^n dvd b'^n" by (auto simp add: nat_power_eq_0_iff)  | 
|
591  | 
have "a' dvd a'^n" by (simp add: m)  | 
|
592  | 
with th0 have "a' dvd b'^n" using dvd_trans[of a' "a'^n" "b'^n"] by simp  | 
|
593  | 
hence th1: "a' dvd b'^m * b'" by (simp add: m mult_commute)  | 
|
594  | 
from coprime_divprod[OF th1 coprime_exp[OF ab'(3), of m]]  | 
|
595  | 
have "a' dvd b'" .  | 
|
596  | 
hence "a'*?g dvd b'*?g" by simp  | 
|
597  | 
with ab'(1,2) have ?thesis by simp }  | 
|
598  | 
ultimately show ?thesis by blast  | 
|
599  | 
qed  | 
|
600  | 
||
601  | 
lemma divides_mul: assumes mr: "m dvd r" and nr: "n dvd r" and mn:"coprime m n"  | 
|
602  | 
shows "m * n dvd r"  | 
|
603  | 
proof-  | 
|
604  | 
from mr nr obtain m' n' where m': "r = m*m'" and n': "r = n*n'"  | 
|
605  | 
unfolding dvd_def by blast  | 
|
606  | 
from mr n' have "m dvd n'*n" by (simp add: mult_commute)  | 
|
607  | 
hence "m dvd n'" using relprime_dvd_mult_iff[OF mn[unfolded coprime_def]] by simp  | 
|
608  | 
then obtain k where k: "n' = m*k" unfolding dvd_def by blast  | 
|
609  | 
from n' k show ?thesis unfolding dvd_def by auto  | 
|
610  | 
qed  | 
|
611  | 
||
| 26144 | 612  | 
|
613  | 
text {* A binary form of the Chinese Remainder Theorem. *}
 | 
|
| 26125 | 614  | 
|
615  | 
lemma chinese_remainder: assumes ab: "coprime a b" and a:"a \<noteq> 0" and b:"b \<noteq> 0"  | 
|
616  | 
shows "\<exists>x q1 q2. x = u + q1 * a \<and> x = v + q2 * b"  | 
|
617  | 
proof-  | 
|
618  | 
from bezout_add_strong[OF a, of b] bezout_add_strong[OF b, of a]  | 
|
619  | 
obtain d1 x1 y1 d2 x2 y2 where dxy1: "d1 dvd a" "d1 dvd b" "a * x1 = b * y1 + d1"  | 
|
620  | 
and dxy2: "d2 dvd b" "d2 dvd a" "b * x2 = a * y2 + d2" by blast  | 
|
621  | 
from gcd_unique[of 1 a b, simplified ab[unfolded coprime_def], simplified]  | 
|
622  | 
dxy1(1,2) dxy2(1,2) have d12: "d1 = 1" "d2 =1" by auto  | 
|
623  | 
let ?x = "v * a * x1 + u * b * x2"  | 
|
624  | 
let ?q1 = "v * x1 + u * y2"  | 
|
625  | 
let ?q2 = "v * y1 + u * x2"  | 
|
626  | 
from dxy2(3)[simplified d12] dxy1(3)[simplified d12]  | 
|
627  | 
have "?x = u + ?q1 * a" "?x = v + ?q2 * b" by algebra+  | 
|
628  | 
thus ?thesis by blast  | 
|
629  | 
qed  | 
|
630  | 
||
| 26144 | 631  | 
text {* Primality *}
 | 
632  | 
||
633  | 
text {* A few useful theorems about primes *}
 | 
|
| 26125 | 634  | 
|
635  | 
lemma prime_0[simp]: "~prime 0" by (simp add: prime_def)  | 
|
636  | 
lemma prime_1[simp]: "~ prime 1" by (simp add: prime_def)  | 
|
637  | 
lemma prime_Suc0[simp]: "~ prime (Suc 0)" by (simp add: prime_def)  | 
|
638  | 
||
639  | 
lemma prime_ge_2: "prime p ==> p \<ge> 2" by (simp add: prime_def)  | 
|
640  | 
lemma prime_factor: assumes n: "n \<noteq> 1" shows "\<exists> p. prime p \<and> p dvd n"  | 
|
641  | 
using n  | 
|
642  | 
proof(induct n rule: nat_less_induct)  | 
|
643  | 
fix n  | 
|
644  | 
assume H: "\<forall>m<n. m \<noteq> 1 \<longrightarrow> (\<exists>p. prime p \<and> p dvd m)" "n \<noteq> 1"  | 
|
645  | 
let ?ths = "\<exists>p. prime p \<and> p dvd n"  | 
|
646  | 
  {assume "n=0" hence ?ths using two_is_prime by auto}
 | 
|
647  | 
moreover  | 
|
648  | 
  {assume nz: "n\<noteq>0" 
 | 
|
649  | 
    {assume "prime n" hence ?ths by - (rule exI[where x="n"], simp)}
 | 
|
650  | 
moreover  | 
|
651  | 
    {assume n: "\<not> prime n"
 | 
|
652  | 
with nz H(2)  | 
|
653  | 
obtain k where k:"k dvd n" "k \<noteq> 1" "k \<noteq> n" by (auto simp add: prime_def)  | 
|
654  | 
from dvd_imp_le[OF k(1)] nz k(3) have kn: "k < n" by simp  | 
|
655  | 
from H(1)[rule_format, OF kn k(2)] obtain p where p: "prime p" "p dvd k" by blast  | 
|
656  | 
from dvd_trans[OF p(2) k(1)] p(1) have ?ths by blast}  | 
|
657  | 
ultimately have ?ths by blast}  | 
|
658  | 
ultimately show ?ths by blast  | 
|
659  | 
qed  | 
|
660  | 
||
661  | 
lemma prime_factor_lt: assumes p: "prime p" and n: "n \<noteq> 0" and npm:"n = p * m"  | 
|
662  | 
shows "m < n"  | 
|
663  | 
proof-  | 
|
664  | 
  {assume "m=0" with n have ?thesis by simp}
 | 
|
665  | 
moreover  | 
|
666  | 
  {assume m: "m \<noteq> 0"
 | 
|
667  | 
from npm have mn: "m dvd n" unfolding dvd_def by auto  | 
|
668  | 
from npm m have "n \<noteq> m" using p by auto  | 
|
669  | 
with dvd_imp_le[OF mn] n have ?thesis by simp}  | 
|
670  | 
ultimately show ?thesis by blast  | 
|
671  | 
qed  | 
|
672  | 
||
673  | 
lemma euclid_bound: "\<exists>p. prime p \<and> n < p \<and> p <= Suc (fact n)"  | 
|
674  | 
proof-  | 
|
675  | 
have f1: "fact n + 1 \<noteq> 1" using fact_le[of n] by arith  | 
|
676  | 
from prime_factor[OF f1] obtain p where p: "prime p" "p dvd fact n + 1" by blast  | 
|
677  | 
from dvd_imp_le[OF p(2)] have pfn: "p \<le> fact n + 1" by simp  | 
|
678  | 
  {assume np: "p \<le> n"
 | 
|
679  | 
from p(1) have p1: "p \<ge> 1" by (cases p, simp_all)  | 
|
680  | 
from divides_fact[OF p1 np] have pfn': "p dvd fact n" .  | 
|
681  | 
from divides_add_revr[OF pfn' p(2)] p(1) have False by simp}  | 
|
682  | 
hence "n < p" by arith  | 
|
683  | 
with p(1) pfn show ?thesis by auto  | 
|
684  | 
qed  | 
|
685  | 
||
686  | 
lemma euclid: "\<exists>p. prime p \<and> p > n" using euclid_bound by auto  | 
|
687  | 
lemma primes_infinite: "\<not> (finite {p. prime p})"
 | 
|
688  | 
proof (auto simp add: finite_conv_nat_seg_image)  | 
|
689  | 
fix n f  | 
|
690  | 
  assume H: "Collect prime = f ` {i. i < (n::nat)}"
 | 
|
691  | 
let ?P = "Collect prime"  | 
|
692  | 
let ?m = "Max ?P"  | 
|
693  | 
  have P0: "?P \<noteq> {}" using two_is_prime by auto
 | 
|
694  | 
from H have fP: "finite ?P" using finite_conv_nat_seg_image by blast  | 
|
| 
26757
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26159 
diff
changeset
 | 
695  | 
from Max_in [OF fP P0] have "?m \<in> ?P" .  | 
| 
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26159 
diff
changeset
 | 
696  | 
from Max_ge [OF fP] have contr: "\<forall> p. prime p \<longrightarrow> p \<le> ?m" by blast  | 
| 
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26159 
diff
changeset
 | 
697  | 
from euclid [of ?m] obtain q where q: "prime q" "q > ?m" by blast  | 
| 26125 | 698  | 
with contr show False by auto  | 
699  | 
qed  | 
|
700  | 
||
701  | 
lemma coprime_prime: assumes ab: "coprime a b"  | 
|
702  | 
shows "~(prime p \<and> p dvd a \<and> p dvd b)"  | 
|
703  | 
proof  | 
|
704  | 
assume "prime p \<and> p dvd a \<and> p dvd b"  | 
|
705  | 
thus False using ab gcd_greatest[of p a b] by (simp add: coprime_def)  | 
|
706  | 
qed  | 
|
707  | 
lemma coprime_prime_eq: "coprime a b \<longleftrightarrow> (\<forall>p. ~(prime p \<and> p dvd a \<and> p dvd b))"  | 
|
708  | 
(is "?lhs = ?rhs")  | 
|
709  | 
proof-  | 
|
710  | 
  {assume "?lhs" with coprime_prime  have ?rhs by blast}
 | 
|
711  | 
moreover  | 
|
712  | 
  {assume r: "?rhs" and c: "\<not> ?lhs"
 | 
|
713  | 
then obtain g where g: "g\<noteq>1" "g dvd a" "g dvd b" unfolding coprime_def by blast  | 
|
714  | 
from prime_factor[OF g(1)] obtain p where p: "prime p" "p dvd g" by blast  | 
|
715  | 
from dvd_trans [OF p(2) g(2)] dvd_trans [OF p(2) g(3)]  | 
|
716  | 
have "p dvd a" "p dvd b" . with p(1) r have False by blast}  | 
|
717  | 
ultimately show ?thesis by blast  | 
|
718  | 
qed  | 
|
719  | 
||
720  | 
lemma prime_coprime: assumes p: "prime p"  | 
|
721  | 
shows "n = 1 \<or> p dvd n \<or> coprime p n"  | 
|
722  | 
using p prime_imp_relprime[of p n] by (auto simp add: coprime_def)  | 
|
723  | 
||
724  | 
lemma prime_coprime_strong: "prime p \<Longrightarrow> p dvd n \<or> coprime p n"  | 
|
725  | 
using prime_coprime[of p n] by auto  | 
|
726  | 
||
727  | 
declare coprime_0[simp]  | 
|
728  | 
||
729  | 
lemma coprime_0'[simp]: "coprime 0 d \<longleftrightarrow> d = 1" by (simp add: coprime_commute[of 0 d])  | 
|
730  | 
lemma coprime_bezout_strong: assumes ab: "coprime a b" and b: "b \<noteq> 1"  | 
|
731  | 
shows "\<exists>x y. a * x = b * y + 1"  | 
|
732  | 
proof-  | 
|
733  | 
from ab b have az: "a \<noteq> 0" by - (rule ccontr, auto)  | 
|
734  | 
from bezout_gcd_strong[OF az, of b] ab[unfolded coprime_def]  | 
|
735  | 
show ?thesis by auto  | 
|
736  | 
qed  | 
|
737  | 
||
738  | 
lemma bezout_prime: assumes p: "prime p" and pa: "\<not> p dvd a"  | 
|
739  | 
shows "\<exists>x y. a*x = p*y + 1"  | 
|
740  | 
proof-  | 
|
741  | 
from p have p1: "p \<noteq> 1" using prime_1 by blast  | 
|
742  | 
from prime_coprime[OF p, of a] p1 pa have ap: "coprime a p"  | 
|
743  | 
by (auto simp add: coprime_commute)  | 
|
744  | 
from coprime_bezout_strong[OF ap p1] show ?thesis .  | 
|
745  | 
qed  | 
|
746  | 
lemma prime_divprod: assumes p: "prime p" and pab: "p dvd a*b"  | 
|
747  | 
shows "p dvd a \<or> p dvd b"  | 
|
748  | 
proof-  | 
|
749  | 
  {assume "a=1" hence ?thesis using pab by simp }
 | 
|
750  | 
moreover  | 
|
751  | 
  {assume "p dvd a" hence ?thesis by blast}
 | 
|
752  | 
moreover  | 
|
753  | 
  {assume pa: "coprime p a" from coprime_divprod[OF pab pa]  have ?thesis .. }
 | 
|
754  | 
ultimately show ?thesis using prime_coprime[OF p, of a] by blast  | 
|
755  | 
qed  | 
|
756  | 
||
757  | 
lemma prime_divprod_eq: assumes p: "prime p"  | 
|
758  | 
shows "p dvd a*b \<longleftrightarrow> p dvd a \<or> p dvd b"  | 
|
759  | 
using p prime_divprod dvd_mult dvd_mult2 by auto  | 
|
760  | 
||
761  | 
lemma prime_divexp: assumes p:"prime p" and px: "p dvd x^n"  | 
|
762  | 
shows "p dvd x"  | 
|
763  | 
using px  | 
|
764  | 
proof(induct n)  | 
|
765  | 
case 0 thus ?case by simp  | 
|
766  | 
next  | 
|
767  | 
case (Suc n)  | 
|
768  | 
hence th: "p dvd x*x^n" by simp  | 
|
769  | 
  {assume H: "p dvd x^n"
 | 
|
770  | 
from Suc.hyps[OF H] have ?case .}  | 
|
771  | 
with prime_divprod[OF p th] show ?case by blast  | 
|
772  | 
qed  | 
|
773  | 
||
774  | 
lemma prime_divexp_n: "prime p \<Longrightarrow> p dvd x^n \<Longrightarrow> p^n dvd x^n"  | 
|
775  | 
using prime_divexp[of p x n] divides_exp[of p x n] by blast  | 
|
776  | 
||
777  | 
lemma coprime_prime_dvd_ex: assumes xy: "\<not>coprime x y"  | 
|
778  | 
shows "\<exists>p. prime p \<and> p dvd x \<and> p dvd y"  | 
|
779  | 
proof-  | 
|
780  | 
from xy[unfolded coprime_def] obtain g where g: "g \<noteq> 1" "g dvd x" "g dvd y"  | 
|
781  | 
by blast  | 
|
782  | 
from prime_factor[OF g(1)] obtain p where p: "prime p" "p dvd g" by blast  | 
|
783  | 
from g(2,3) dvd_trans[OF p(2)] p(1) show ?thesis by auto  | 
|
784  | 
qed  | 
|
785  | 
lemma coprime_sos: assumes xy: "coprime x y"  | 
|
786  | 
shows "coprime (x * y) (x^2 + y^2)"  | 
|
787  | 
proof-  | 
|
788  | 
  {assume c: "\<not> coprime (x * y) (x^2 + y^2)"
 | 
|
789  | 
from coprime_prime_dvd_ex[OF c] obtain p  | 
|
790  | 
where p: "prime p" "p dvd x*y" "p dvd x^2 + y^2" by blast  | 
|
791  | 
    {assume px: "p dvd x"
 | 
|
| 
27651
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27567 
diff
changeset
 | 
792  | 
from dvd_mult[OF px, of x] p(3)  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27567 
diff
changeset
 | 
793  | 
obtain r s where "x * x = p * r" and "x^2 + y^2 = p * s"  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27567 
diff
changeset
 | 
794  | 
by (auto elim!: dvdE)  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27567 
diff
changeset
 | 
795  | 
then have "y^2 = p * (s - r)"  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27567 
diff
changeset
 | 
796  | 
by (auto simp add: power2_eq_square diff_mult_distrib2)  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27567 
diff
changeset
 | 
797  | 
then have "p dvd y^2" ..  | 
| 26125 | 798  | 
with prime_divexp[OF p(1), of y 2] have py: "p dvd y" .  | 
799  | 
from p(1) px py xy[unfolded coprime, rule_format, of p] prime_1  | 
|
800  | 
have False by simp }  | 
|
801  | 
moreover  | 
|
802  | 
    {assume py: "p dvd y"
 | 
|
| 
27651
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27567 
diff
changeset
 | 
803  | 
from dvd_mult[OF py, of y] p(3)  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27567 
diff
changeset
 | 
804  | 
obtain r s where "y * y = p * r" and "x^2 + y^2 = p * s"  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27567 
diff
changeset
 | 
805  | 
by (auto elim!: dvdE)  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27567 
diff
changeset
 | 
806  | 
then have "x^2 = p * (s - r)"  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27567 
diff
changeset
 | 
807  | 
by (auto simp add: power2_eq_square diff_mult_distrib2)  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27567 
diff
changeset
 | 
808  | 
then have "p dvd x^2" ..  | 
| 26125 | 809  | 
with prime_divexp[OF p(1), of x 2] have px: "p dvd x" .  | 
810  | 
from p(1) px py xy[unfolded coprime, rule_format, of p] prime_1  | 
|
811  | 
have False by simp }  | 
|
812  | 
ultimately have False using prime_divprod[OF p(1,2)] by blast}  | 
|
813  | 
thus ?thesis by blast  | 
|
814  | 
qed  | 
|
815  | 
||
816  | 
lemma distinct_prime_coprime: "prime p \<Longrightarrow> prime q \<Longrightarrow> p \<noteq> q \<Longrightarrow> coprime p q"  | 
|
817  | 
unfolding prime_def coprime_prime_eq by blast  | 
|
818  | 
||
819  | 
lemma prime_coprime_lt: assumes p: "prime p" and x: "0 < x" and xp: "x < p"  | 
|
820  | 
shows "coprime x p"  | 
|
821  | 
proof-  | 
|
822  | 
  {assume c: "\<not> coprime x p"
 | 
|
823  | 
then obtain g where g: "g \<noteq> 1" "g dvd x" "g dvd p" unfolding coprime_def by blast  | 
|
824  | 
from dvd_imp_le[OF g(2)] x xp have gp: "g < p" by arith  | 
|
825  | 
from g(2) x have "g \<noteq> 0" by - (rule ccontr, simp)  | 
|
826  | 
with g gp p[unfolded prime_def] have False by blast}  | 
|
827  | 
thus ?thesis by blast  | 
|
828  | 
qed  | 
|
829  | 
||
830  | 
lemma even_dvd[simp]: "even (n::nat) \<longleftrightarrow> 2 dvd n" by presburger  | 
|
831  | 
lemma prime_odd: "prime p \<Longrightarrow> p = 2 \<or> odd p" unfolding prime_def by auto  | 
|
832  | 
||
| 26144 | 833  | 
|
834  | 
text {* One property of coprimality is easier to prove via prime factors. *}
 | 
|
| 26125 | 835  | 
|
836  | 
lemma prime_divprod_pow:  | 
|
837  | 
assumes p: "prime p" and ab: "coprime a b" and pab: "p^n dvd a * b"  | 
|
838  | 
shows "p^n dvd a \<or> p^n dvd b"  | 
|
839  | 
proof-  | 
|
840  | 
  {assume "n = 0 \<or> a = 1 \<or> b = 1" with pab have ?thesis 
 | 
|
841  | 
apply (cases "n=0", simp_all)  | 
|
842  | 
apply (cases "a=1", simp_all) done}  | 
|
843  | 
moreover  | 
|
844  | 
  {assume n: "n \<noteq> 0" and a: "a\<noteq>1" and b: "b\<noteq>1" 
 | 
|
845  | 
then obtain m where m: "n = Suc m" by (cases n, auto)  | 
|
846  | 
from divides_exp2[OF n pab] have pab': "p dvd a*b" .  | 
|
847  | 
from prime_divprod[OF p pab']  | 
|
848  | 
have "p dvd a \<or> p dvd b" .  | 
|
849  | 
moreover  | 
|
850  | 
    {assume pa: "p dvd a"
 | 
|
851  | 
have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute)  | 
|
852  | 
from coprime_prime[OF ab, of p] p pa have "\<not> p dvd b" by blast  | 
|
853  | 
with prime_coprime[OF p, of b] b  | 
|
854  | 
have cpb: "coprime b p" using coprime_commute by blast  | 
|
855  | 
from coprime_exp[OF cpb] have pnb: "coprime (p^n) b"  | 
|
856  | 
by (simp add: coprime_commute)  | 
|
857  | 
from coprime_divprod[OF pnba pnb] have ?thesis by blast }  | 
|
858  | 
moreover  | 
|
859  | 
    {assume pb: "p dvd b"
 | 
|
860  | 
have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute)  | 
|
861  | 
from coprime_prime[OF ab, of p] p pb have "\<not> p dvd a" by blast  | 
|
862  | 
with prime_coprime[OF p, of a] a  | 
|
863  | 
have cpb: "coprime a p" using coprime_commute by blast  | 
|
864  | 
from coprime_exp[OF cpb] have pnb: "coprime (p^n) a"  | 
|
865  | 
by (simp add: coprime_commute)  | 
|
866  | 
from coprime_divprod[OF pab pnb] have ?thesis by blast }  | 
|
867  | 
ultimately have ?thesis by blast}  | 
|
868  | 
ultimately show ?thesis by blast  | 
|
869  | 
qed  | 
|
870  | 
||
871  | 
lemma nat_mult_eq_one: "(n::nat) * m = 1 \<longleftrightarrow> n = 1 \<and> m = 1" (is "?lhs \<longleftrightarrow> ?rhs")  | 
|
872  | 
proof  | 
|
873  | 
assume H: "?lhs"  | 
|
874  | 
hence "n dvd 1" "m dvd 1" unfolding dvd_def by (auto simp add: mult_commute)  | 
|
875  | 
thus ?rhs by auto  | 
|
876  | 
next  | 
|
877  | 
assume ?rhs then show ?lhs by auto  | 
|
878  | 
qed  | 
|
879  | 
||
880  | 
lemma power_Suc0[simp]: "Suc 0 ^ n = Suc 0"  | 
|
881  | 
unfolding One_nat_def[symmetric] power_one ..  | 
|
882  | 
lemma coprime_pow: assumes ab: "coprime a b" and abcn: "a * b = c ^n"  | 
|
883  | 
shows "\<exists>r s. a = r^n \<and> b = s ^n"  | 
|
884  | 
using ab abcn  | 
|
885  | 
proof(induct c arbitrary: a b rule: nat_less_induct)  | 
|
886  | 
fix c a b  | 
|
887  | 
assume H: "\<forall>m<c. \<forall>a b. coprime a b \<longrightarrow> a * b = m ^ n \<longrightarrow> (\<exists>r s. a = r ^ n \<and> b = s ^ n)" "coprime a b" "a * b = c ^ n"  | 
|
888  | 
let ?ths = "\<exists>r s. a = r^n \<and> b = s ^n"  | 
|
889  | 
  {assume n: "n = 0"
 | 
|
890  | 
with H(3) power_one have "a*b = 1" by simp  | 
|
891  | 
hence "a = 1 \<and> b = 1" by simp  | 
|
892  | 
hence ?ths  | 
|
893  | 
apply -  | 
|
894  | 
apply (rule exI[where x=1])  | 
|
895  | 
apply (rule exI[where x=1])  | 
|
896  | 
using power_one[of n]  | 
|
897  | 
by simp}  | 
|
898  | 
moreover  | 
|
899  | 
  {assume n: "n \<noteq> 0" then obtain m where m: "n = Suc m" by (cases n, auto)
 | 
|
900  | 
    {assume c: "c = 0"
 | 
|
901  | 
with H(3) m H(2) have ?ths apply simp  | 
|
902  | 
apply (cases "a=0", simp_all)  | 
|
903  | 
apply (rule exI[where x="0"], simp)  | 
|
904  | 
apply (rule exI[where x="0"], simp)  | 
|
905  | 
done}  | 
|
906  | 
moreover  | 
|
907  | 
    {assume "c=1" with H(3) power_one have "a*b = 1" by simp 
 | 
|
908  | 
hence "a = 1 \<and> b = 1" by simp  | 
|
909  | 
hence ?ths  | 
|
910  | 
apply -  | 
|
911  | 
apply (rule exI[where x=1])  | 
|
912  | 
apply (rule exI[where x=1])  | 
|
913  | 
using power_one[of n]  | 
|
914  | 
by simp}  | 
|
915  | 
moreover  | 
|
916  | 
  {assume c: "c\<noteq>1" "c \<noteq> 0"
 | 
|
917  | 
from prime_factor[OF c(1)] obtain p where p: "prime p" "p dvd c" by blast  | 
|
918  | 
from prime_divprod_pow[OF p(1) H(2), unfolded H(3), OF divides_exp[OF p(2), of n]]  | 
|
919  | 
have pnab: "p ^ n dvd a \<or> p^n dvd b" .  | 
|
920  | 
from p(2) obtain l where l: "c = p*l" unfolding dvd_def by blast  | 
|
921  | 
have pn0: "p^n \<noteq> 0" using n prime_ge_2 [OF p(1)] by (simp add: neq0_conv)  | 
|
922  | 
    {assume pa: "p^n dvd a"
 | 
|
923  | 
then obtain k where k: "a = p^n * k" unfolding dvd_def by blast  | 
|
924  | 
from l have "l dvd c" by auto  | 
|
925  | 
with dvd_imp_le[of l c] c have "l \<le> c" by auto  | 
|
926  | 
      moreover {assume "l = c" with l c  have "p = 1" by simp with p have False by simp}
 | 
|
927  | 
ultimately have lc: "l < c" by arith  | 
|
928  | 
from coprime_lmul2 [OF H(2)[unfolded k coprime_commute[of "p^n*k" b]]]  | 
|
929  | 
have kb: "coprime k b" by (simp add: coprime_commute)  | 
|
930  | 
from H(3) l k pn0 have kbln: "k * b = l ^ n"  | 
|
931  | 
by (auto simp add: power_mult_distrib)  | 
|
932  | 
from H(1)[rule_format, OF lc kb kbln]  | 
|
933  | 
obtain r s where rs: "k = r ^n" "b = s^n" by blast  | 
|
934  | 
from k rs(1) have "a = (p*r)^n" by (simp add: power_mult_distrib)  | 
|
935  | 
with rs(2) have ?ths by blast }  | 
|
936  | 
moreover  | 
|
937  | 
    {assume pb: "p^n dvd b"
 | 
|
938  | 
then obtain k where k: "b = p^n * k" unfolding dvd_def by blast  | 
|
939  | 
from l have "l dvd c" by auto  | 
|
940  | 
with dvd_imp_le[of l c] c have "l \<le> c" by auto  | 
|
941  | 
      moreover {assume "l = c" with l c  have "p = 1" by simp with p have False by simp}
 | 
|
942  | 
ultimately have lc: "l < c" by arith  | 
|
943  | 
from coprime_lmul2 [OF H(2)[unfolded k coprime_commute[of "p^n*k" a]]]  | 
|
944  | 
have kb: "coprime k a" by (simp add: coprime_commute)  | 
|
945  | 
from H(3) l k pn0 n have kbln: "k * a = l ^ n"  | 
|
946  | 
by (simp add: power_mult_distrib mult_commute)  | 
|
947  | 
from H(1)[rule_format, OF lc kb kbln]  | 
|
948  | 
obtain r s where rs: "k = r ^n" "a = s^n" by blast  | 
|
949  | 
from k rs(1) have "b = (p*r)^n" by (simp add: power_mult_distrib)  | 
|
950  | 
with rs(2) have ?ths by blast }  | 
|
951  | 
ultimately have ?ths using pnab by blast}  | 
|
952  | 
ultimately have ?ths by blast}  | 
|
953  | 
ultimately show ?ths by blast  | 
|
954  | 
qed  | 
|
955  | 
||
| 26144 | 956  | 
text {* More useful lemmas. *}
 | 
| 26125 | 957  | 
lemma prime_product:  | 
| 
27651
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27567 
diff
changeset
 | 
958  | 
assumes "prime (p * q)"  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27567 
diff
changeset
 | 
959  | 
shows "p = 1 \<or> q = 1"  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27567 
diff
changeset
 | 
960  | 
proof -  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27567 
diff
changeset
 | 
961  | 
from assms have  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27567 
diff
changeset
 | 
962  | 
"1 < p * q" and P: "\<And>m. m dvd p * q \<Longrightarrow> m = 1 \<or> m = p * q"  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27567 
diff
changeset
 | 
963  | 
unfolding prime_def by auto  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27567 
diff
changeset
 | 
964  | 
from `1 < p * q` have "p \<noteq> 0" by (cases p) auto  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27567 
diff
changeset
 | 
965  | 
then have Q: "p = p * q \<longleftrightarrow> q = 1" by auto  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27567 
diff
changeset
 | 
966  | 
have "p dvd p * q" by simp  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27567 
diff
changeset
 | 
967  | 
then have "p = 1 \<or> p = p * q" by (rule P)  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27567 
diff
changeset
 | 
968  | 
then show ?thesis by (simp add: Q)  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27567 
diff
changeset
 | 
969  | 
qed  | 
| 26125 | 970  | 
|
971  | 
lemma prime_exp: "prime (p^n) \<longleftrightarrow> prime p \<and> n = 1"  | 
|
972  | 
proof(induct n)  | 
|
973  | 
case 0 thus ?case by simp  | 
|
974  | 
next  | 
|
975  | 
case (Suc n)  | 
|
976  | 
  {assume "p = 0" hence ?case by simp}
 | 
|
977  | 
moreover  | 
|
978  | 
  {assume "p=1" hence ?case by simp}
 | 
|
979  | 
moreover  | 
|
980  | 
  {assume p: "p \<noteq> 0" "p\<noteq>1"
 | 
|
981  | 
    {assume pp: "prime (p^Suc n)"
 | 
|
982  | 
hence "p = 1 \<or> p^n = 1" using prime_product[of p "p^n"] by simp  | 
|
983  | 
with p have n: "n = 0"  | 
|
984  | 
by (simp only: exp_eq_1 ) simp  | 
|
985  | 
with pp have "prime p \<and> Suc n = 1" by simp}  | 
|
986  | 
moreover  | 
|
987  | 
    {assume n: "prime p \<and> Suc n = 1" hence "prime (p^Suc n)" by simp}
 | 
|
988  | 
ultimately have ?case by blast}  | 
|
989  | 
ultimately show ?case by blast  | 
|
990  | 
qed  | 
|
991  | 
||
992  | 
lemma prime_power_mult:  | 
|
993  | 
assumes p: "prime p" and xy: "x * y = p ^ k"  | 
|
994  | 
shows "\<exists>i j. x = p ^i \<and> y = p^ j"  | 
|
995  | 
using xy  | 
|
996  | 
proof(induct k arbitrary: x y)  | 
|
997  | 
case 0 thus ?case apply simp by (rule exI[where x="0"], simp)  | 
|
998  | 
next  | 
|
999  | 
case (Suc k x y)  | 
|
1000  | 
from Suc.prems have pxy: "p dvd x*y" by auto  | 
|
1001  | 
from prime_divprod[OF p pxy] have pxyc: "p dvd x \<or> p dvd y" .  | 
|
1002  | 
from p have p0: "p \<noteq> 0" by - (rule ccontr, simp)  | 
|
1003  | 
  {assume px: "p dvd x"
 | 
|
1004  | 
then obtain d where d: "x = p*d" unfolding dvd_def by blast  | 
|
1005  | 
from Suc.prems d have "p*d*y = p^Suc k" by simp  | 
|
1006  | 
hence th: "d*y = p^k" using p0 by simp  | 
|
1007  | 
from Suc.hyps[OF th] obtain i j where ij: "d = p^i" "y = p^j" by blast  | 
|
1008  | 
with d have "x = p^Suc i" by simp  | 
|
1009  | 
with ij(2) have ?case by blast}  | 
|
1010  | 
moreover  | 
|
1011  | 
  {assume px: "p dvd y"
 | 
|
1012  | 
then obtain d where d: "y = p*d" unfolding dvd_def by blast  | 
|
1013  | 
from Suc.prems d have "p*d*x = p^Suc k" by (simp add: mult_commute)  | 
|
1014  | 
hence th: "d*x = p^k" using p0 by simp  | 
|
1015  | 
from Suc.hyps[OF th] obtain i j where ij: "d = p^i" "x = p^j" by blast  | 
|
1016  | 
with d have "y = p^Suc i" by simp  | 
|
1017  | 
with ij(2) have ?case by blast}  | 
|
1018  | 
ultimately show ?case using pxyc by blast  | 
|
1019  | 
qed  | 
|
1020  | 
||
1021  | 
lemma prime_power_exp: assumes p: "prime p" and n:"n \<noteq> 0"  | 
|
1022  | 
and xn: "x^n = p^k" shows "\<exists>i. x = p^i"  | 
|
1023  | 
using n xn  | 
|
1024  | 
proof(induct n arbitrary: k)  | 
|
1025  | 
case 0 thus ?case by simp  | 
|
1026  | 
next  | 
|
1027  | 
case (Suc n k) hence th: "x*x^n = p^k" by simp  | 
|
1028  | 
  {assume "n = 0" with prems have ?case apply simp 
 | 
|
1029  | 
by (rule exI[where x="k"],simp)}  | 
|
1030  | 
moreover  | 
|
1031  | 
  {assume n: "n \<noteq> 0"
 | 
|
1032  | 
from prime_power_mult[OF p th]  | 
|
1033  | 
obtain i j where ij: "x = p^i" "x^n = p^j"by blast  | 
|
1034  | 
from Suc.hyps[OF n ij(2)] have ?case .}  | 
|
1035  | 
ultimately show ?case by blast  | 
|
1036  | 
qed  | 
|
1037  | 
||
1038  | 
lemma divides_primepow: assumes p: "prime p"  | 
|
1039  | 
shows "d dvd p^k \<longleftrightarrow> (\<exists> i. i \<le> k \<and> d = p ^i)"  | 
|
1040  | 
proof  | 
|
1041  | 
assume H: "d dvd p^k" then obtain e where e: "d*e = p^k"  | 
|
1042  | 
unfolding dvd_def apply (auto simp add: mult_commute) by blast  | 
|
1043  | 
from prime_power_mult[OF p e] obtain i j where ij: "d = p^i" "e=p^j" by blast  | 
|
1044  | 
from prime_ge_2[OF p] have p1: "p > 1" by arith  | 
|
1045  | 
from e ij have "p^(i + j) = p^k" by (simp add: power_add)  | 
|
1046  | 
hence "i + j = k" using power_inject_exp[of p "i+j" k, OF p1] by simp  | 
|
1047  | 
hence "i \<le> k" by arith  | 
|
1048  | 
with ij(1) show "\<exists>i\<le>k. d = p ^ i" by blast  | 
|
1049  | 
next  | 
|
1050  | 
  {fix i assume H: "i \<le> k" "d = p^i"
 | 
|
1051  | 
hence "\<exists>j. k = i + j" by arith  | 
|
1052  | 
then obtain j where j: "k = i + j" by blast  | 
|
1053  | 
hence "p^k = p^j*d" using H(2) by (simp add: power_add)  | 
|
1054  | 
hence "d dvd p^k" unfolding dvd_def by auto}  | 
|
1055  | 
thus "\<exists>i\<le>k. d = p ^ i \<Longrightarrow> d dvd p ^ k" by blast  | 
|
1056  | 
qed  | 
|
1057  | 
||
1058  | 
lemma coprime_divisors: "d dvd a \<Longrightarrow> e dvd b \<Longrightarrow> coprime a b \<Longrightarrow> coprime d e"  | 
|
1059  | 
by (auto simp add: dvd_def coprime)  | 
|
1060  | 
||
| 26159 | 1061  | 
declare power_Suc0[simp del]  | 
1062  | 
declare even_dvd[simp del]  | 
|
| 
26757
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26159 
diff
changeset
 | 
1063  | 
|
| 11363 | 1064  | 
end  |