| author | blanchet |
| Mon, 20 Jan 2014 18:25:44 +0100 | |
| changeset 55077 | 4cf280104b85 |
| parent 54228 | 229282d53781 |
| child 55130 | 70db8d380d62 |
| permissions | -rw-r--r-- |
| 32479 | 1 |
(* Authors: Christophe Tabacznyj, Lawrence C. Paulson, Amine Chaieb, |
| 31798 | 2 |
Thomas M. Rasmussen, Jeremy Avigad, Tobias Nipkow |
| 31706 | 3 |
|
4 |
||
| 32479 | 5 |
This file deals with properties of primes. Definitions and lemmas are |
6 |
proved uniformly for the natural numbers and integers. |
|
| 31706 | 7 |
|
8 |
This file combines and revises a number of prior developments. |
|
9 |
||
10 |
The original theories "GCD" and "Primes" were by Christophe Tabacznyj |
|
11 |
and Lawrence C. Paulson, based on \cite{davenport92}. They introduced
|
|
12 |
gcd, lcm, and prime for the natural numbers. |
|
13 |
||
14 |
The original theory "IntPrimes" was by Thomas M. Rasmussen, and |
|
15 |
extended gcd, lcm, primes to the integers. Amine Chaieb provided |
|
16 |
another extension of the notions to the integers, and added a number |
|
17 |
of results to "Primes" and "GCD". IntPrimes also defined and developed |
|
18 |
the congruence relations on the integers. The notion was extended to |
|
| 33718 | 19 |
the natural numbers by Chaieb. |
| 31706 | 20 |
|
|
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
21 |
Jeremy Avigad combined all of these, made everything uniform for the |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
22 |
natural numbers and the integers, and added a number of new theorems. |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
23 |
|
| 31798 | 24 |
Tobias Nipkow cleaned up a lot. |
| 21256 | 25 |
*) |
26 |
||
| 31706 | 27 |
|
| 32479 | 28 |
header {* Primes *}
|
| 21256 | 29 |
|
| 32479 | 30 |
theory Primes |
|
37294
a2a8216999a2
absolute import -- must work with Main.thy / HOL-Proofs
haftmann
parents:
35644
diff
changeset
|
31 |
imports "~~/src/HOL/GCD" |
| 31706 | 32 |
begin |
33 |
||
34 |
class prime = one + |
|
| 44872 | 35 |
fixes prime :: "'a \<Rightarrow> bool" |
| 31706 | 36 |
|
37 |
instantiation nat :: prime |
|
38 |
begin |
|
| 21256 | 39 |
|
| 44872 | 40 |
definition prime_nat :: "nat \<Rightarrow> bool" |
41 |
where "prime_nat p = (1 < p \<and> (\<forall>m. m dvd p --> m = 1 \<or> m = p))" |
|
| 31706 | 42 |
|
| 44872 | 43 |
instance .. |
|
23687
06884f7ffb18
extended - convers now basic lcm properties also
haftmann
parents:
23431
diff
changeset
|
44 |
|
| 31706 | 45 |
end |
46 |
||
47 |
instantiation int :: prime |
|
48 |
begin |
|
49 |
||
| 44872 | 50 |
definition prime_int :: "int \<Rightarrow> bool" |
51 |
where "prime_int p = prime (nat p)" |
|
| 31706 | 52 |
|
| 44872 | 53 |
instance .. |
| 31706 | 54 |
|
55 |
end |
|
56 |
||
57 |
||
58 |
subsection {* Set up Transfer *}
|
|
59 |
||
| 32479 | 60 |
lemma transfer_nat_int_prime: |
| 31706 | 61 |
"(x::int) >= 0 \<Longrightarrow> prime (nat x) = prime x" |
| 44872 | 62 |
unfolding gcd_int_def lcm_int_def prime_int_def by auto |
|
23687
06884f7ffb18
extended - convers now basic lcm properties also
haftmann
parents:
23431
diff
changeset
|
63 |
|
| 35644 | 64 |
declare transfer_morphism_nat_int[transfer add return: |
| 32479 | 65 |
transfer_nat_int_prime] |
| 31706 | 66 |
|
| 44872 | 67 |
lemma transfer_int_nat_prime: "prime (int x) = prime x" |
68 |
unfolding gcd_int_def lcm_int_def prime_int_def by auto |
|
| 31706 | 69 |
|
| 35644 | 70 |
declare transfer_morphism_int_nat[transfer add return: |
| 32479 | 71 |
transfer_int_nat_prime] |
| 31798 | 72 |
|
| 21256 | 73 |
|
| 32479 | 74 |
subsection {* Primes *}
|
| 31706 | 75 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
76 |
lemma prime_odd_nat: "prime (p::nat) \<Longrightarrow> p > 2 \<Longrightarrow> odd p" |
| 54228 | 77 |
apply (auto simp add: prime_nat_def even_def dvd_eq_mod_eq_0) |
78 |
apply (metis dvd_eq_mod_eq_0 even_Suc even_def mod_by_1 nat_dvd_not_less not_mod_2_eq_0_eq_1 zero_less_numeral) |
|
79 |
done |
|
| 31706 | 80 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
81 |
lemma prime_odd_int: "prime (p::int) \<Longrightarrow> p > 2 \<Longrightarrow> odd p" |
| 31706 | 82 |
unfolding prime_int_def |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
83 |
apply (frule prime_odd_nat) |
| 31706 | 84 |
apply (auto simp add: even_nat_def) |
| 44872 | 85 |
done |
| 31706 | 86 |
|
| 31992 | 87 |
(* FIXME Is there a better way to handle these, rather than making them elim rules? *) |
|
22027
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
chaieb
parents:
21404
diff
changeset
|
88 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
89 |
lemma prime_ge_0_nat [elim]: "prime (p::nat) \<Longrightarrow> p >= 0" |
| 44872 | 90 |
unfolding prime_nat_def by auto |
|
22027
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
chaieb
parents:
21404
diff
changeset
|
91 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
92 |
lemma prime_gt_0_nat [elim]: "prime (p::nat) \<Longrightarrow> p > 0" |
| 44872 | 93 |
unfolding prime_nat_def by auto |
| 22367 | 94 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
95 |
lemma prime_ge_1_nat [elim]: "prime (p::nat) \<Longrightarrow> p >= 1" |
| 44872 | 96 |
unfolding prime_nat_def by auto |
|
22027
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
chaieb
parents:
21404
diff
changeset
|
97 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
98 |
lemma prime_gt_1_nat [elim]: "prime (p::nat) \<Longrightarrow> p > 1" |
| 44872 | 99 |
unfolding prime_nat_def by auto |
| 22367 | 100 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
101 |
lemma prime_ge_Suc_0_nat [elim]: "prime (p::nat) \<Longrightarrow> p >= Suc 0" |
| 44872 | 102 |
unfolding prime_nat_def by auto |
| 22367 | 103 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
104 |
lemma prime_gt_Suc_0_nat [elim]: "prime (p::nat) \<Longrightarrow> p > Suc 0" |
| 44872 | 105 |
unfolding prime_nat_def by auto |
| 31706 | 106 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
107 |
lemma prime_ge_2_nat [elim]: "prime (p::nat) \<Longrightarrow> p >= 2" |
| 44872 | 108 |
unfolding prime_nat_def by auto |
| 31706 | 109 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
110 |
lemma prime_ge_0_int [elim]: "prime (p::int) \<Longrightarrow> p >= 0" |
| 44872 | 111 |
unfolding prime_int_def prime_nat_def by auto |
| 22367 | 112 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
113 |
lemma prime_gt_0_int [elim]: "prime (p::int) \<Longrightarrow> p > 0" |
| 44872 | 114 |
unfolding prime_int_def prime_nat_def by auto |
| 31706 | 115 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
116 |
lemma prime_ge_1_int [elim]: "prime (p::int) \<Longrightarrow> p >= 1" |
| 44872 | 117 |
unfolding prime_int_def prime_nat_def by auto |
|
22027
e4a08629c4bd
A few lemmas about relative primes when dividing trough gcd
chaieb
parents:
21404
diff
changeset
|
118 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
119 |
lemma prime_gt_1_int [elim]: "prime (p::int) \<Longrightarrow> p > 1" |
| 44872 | 120 |
unfolding prime_int_def prime_nat_def by auto |
| 31706 | 121 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
122 |
lemma prime_ge_2_int [elim]: "prime (p::int) \<Longrightarrow> p >= 2" |
| 44872 | 123 |
unfolding prime_int_def prime_nat_def by auto |
| 22367 | 124 |
|
| 31706 | 125 |
|
126 |
lemma prime_int_altdef: "prime (p::int) = (1 < p \<and> (\<forall>m \<ge> 0. m dvd p \<longrightarrow> |
|
127 |
m = 1 \<or> m = p))" |
|
128 |
using prime_nat_def [transferred] |
|
| 44872 | 129 |
apply (cases "p >= 0") |
130 |
apply blast |
|
131 |
apply (auto simp add: prime_ge_0_int) |
|
132 |
done |
|
| 31706 | 133 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
134 |
lemma prime_imp_coprime_nat: "prime (p::nat) \<Longrightarrow> \<not> p dvd n \<Longrightarrow> coprime p n" |
| 31706 | 135 |
apply (unfold prime_nat_def) |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
136 |
apply (metis gcd_dvd1_nat gcd_dvd2_nat) |
| 31706 | 137 |
done |
138 |
||
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
139 |
lemma prime_imp_coprime_int: "prime (p::int) \<Longrightarrow> \<not> p dvd n \<Longrightarrow> coprime p n" |
| 31706 | 140 |
apply (unfold prime_int_altdef) |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
141 |
apply (metis gcd_dvd1_int gcd_dvd2_int gcd_ge_0_int) |
|
27568
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents:
27556
diff
changeset
|
142 |
done |
|
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents:
27556
diff
changeset
|
143 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
144 |
lemma prime_dvd_mult_nat: "prime (p::nat) \<Longrightarrow> p dvd m * n \<Longrightarrow> p dvd m \<or> p dvd n" |
|
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
145 |
by (blast intro: coprime_dvd_mult_nat prime_imp_coprime_nat) |
| 31706 | 146 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
147 |
lemma prime_dvd_mult_int: "prime (p::int) \<Longrightarrow> p dvd m * n \<Longrightarrow> p dvd m \<or> p dvd n" |
|
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
148 |
by (blast intro: coprime_dvd_mult_int prime_imp_coprime_int) |
| 31706 | 149 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
150 |
lemma prime_dvd_mult_eq_nat [simp]: "prime (p::nat) \<Longrightarrow> |
| 31706 | 151 |
p dvd m * n = (p dvd m \<or> p dvd n)" |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
152 |
by (rule iffI, rule prime_dvd_mult_nat, auto) |
|
27568
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents:
27556
diff
changeset
|
153 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
154 |
lemma prime_dvd_mult_eq_int [simp]: "prime (p::int) \<Longrightarrow> |
| 31706 | 155 |
p dvd m * n = (p dvd m \<or> p dvd n)" |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
156 |
by (rule iffI, rule prime_dvd_mult_int, auto) |
|
27568
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents:
27556
diff
changeset
|
157 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
158 |
lemma not_prime_eq_prod_nat: "(n::nat) > 1 \<Longrightarrow> ~ prime n \<Longrightarrow> |
| 31706 | 159 |
EX m k. n = m * k & 1 < m & m < n & 1 < k & k < n" |
160 |
unfolding prime_nat_def dvd_def apply auto |
|
| 44872 | 161 |
by (metis mult_commute linorder_neq_iff linorder_not_le mult_1 |
162 |
n_less_n_mult_m one_le_mult_iff less_imp_le_nat) |
|
|
27568
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents:
27556
diff
changeset
|
163 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
164 |
lemma not_prime_eq_prod_int: "(n::int) > 1 \<Longrightarrow> ~ prime n \<Longrightarrow> |
| 31706 | 165 |
EX m k. n = m * k & 1 < m & m < n & 1 < k & k < n" |
166 |
unfolding prime_int_altdef dvd_def |
|
167 |
apply auto |
|
| 44872 | 168 |
by (metis div_mult_self1_is_id div_mult_self2_is_id |
169 |
int_div_less_self int_one_le_iff_zero_less zero_less_mult_pos less_le) |
|
|
27568
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents:
27556
diff
changeset
|
170 |
|
|
53598
2bd8d459ebae
remove unneeded assumption from prime_dvd_power lemmas;
huffman
parents:
47108
diff
changeset
|
171 |
lemma prime_dvd_power_nat: "prime (p::nat) \<Longrightarrow> p dvd x^n \<Longrightarrow> p dvd x" |
|
2bd8d459ebae
remove unneeded assumption from prime_dvd_power lemmas;
huffman
parents:
47108
diff
changeset
|
172 |
by (induct n) auto |
|
27568
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents:
27556
diff
changeset
|
173 |
|
|
53598
2bd8d459ebae
remove unneeded assumption from prime_dvd_power lemmas;
huffman
parents:
47108
diff
changeset
|
174 |
lemma prime_dvd_power_int: "prime (p::int) \<Longrightarrow> p dvd x^n \<Longrightarrow> p dvd x" |
|
2bd8d459ebae
remove unneeded assumption from prime_dvd_power lemmas;
huffman
parents:
47108
diff
changeset
|
175 |
apply (induct n) |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
176 |
apply (frule prime_ge_0_int) |
| 31706 | 177 |
apply auto |
| 44872 | 178 |
done |
| 31706 | 179 |
|
|
53598
2bd8d459ebae
remove unneeded assumption from prime_dvd_power lemmas;
huffman
parents:
47108
diff
changeset
|
180 |
lemma prime_dvd_power_nat_iff: "prime (p::nat) \<Longrightarrow> n > 0 \<Longrightarrow> |
|
2bd8d459ebae
remove unneeded assumption from prime_dvd_power lemmas;
huffman
parents:
47108
diff
changeset
|
181 |
p dvd x^n \<longleftrightarrow> p dvd x" |
|
2bd8d459ebae
remove unneeded assumption from prime_dvd_power lemmas;
huffman
parents:
47108
diff
changeset
|
182 |
by (cases n) (auto elim: prime_dvd_power_nat) |
|
2bd8d459ebae
remove unneeded assumption from prime_dvd_power lemmas;
huffman
parents:
47108
diff
changeset
|
183 |
|
|
2bd8d459ebae
remove unneeded assumption from prime_dvd_power lemmas;
huffman
parents:
47108
diff
changeset
|
184 |
lemma prime_dvd_power_int_iff: "prime (p::int) \<Longrightarrow> n > 0 \<Longrightarrow> |
|
2bd8d459ebae
remove unneeded assumption from prime_dvd_power lemmas;
huffman
parents:
47108
diff
changeset
|
185 |
p dvd x^n \<longleftrightarrow> p dvd x" |
|
2bd8d459ebae
remove unneeded assumption from prime_dvd_power lemmas;
huffman
parents:
47108
diff
changeset
|
186 |
by (cases n) (auto elim: prime_dvd_power_int) |
|
2bd8d459ebae
remove unneeded assumption from prime_dvd_power lemmas;
huffman
parents:
47108
diff
changeset
|
187 |
|
|
2bd8d459ebae
remove unneeded assumption from prime_dvd_power lemmas;
huffman
parents:
47108
diff
changeset
|
188 |
|
| 44872 | 189 |
subsubsection {* Make prime naively executable *}
|
| 32007 | 190 |
|
191 |
lemma zero_not_prime_nat [simp]: "~prime (0::nat)" |
|
192 |
by (simp add: prime_nat_def) |
|
193 |
||
194 |
lemma zero_not_prime_int [simp]: "~prime (0::int)" |
|
195 |
by (simp add: prime_int_def) |
|
196 |
||
197 |
lemma one_not_prime_nat [simp]: "~prime (1::nat)" |
|
198 |
by (simp add: prime_nat_def) |
|
199 |
||
200 |
lemma Suc_0_not_prime_nat [simp]: "~prime (Suc 0)" |
|
201 |
by (simp add: prime_nat_def One_nat_def) |
|
202 |
||
203 |
lemma one_not_prime_int [simp]: "~prime (1::int)" |
|
204 |
by (simp add: prime_int_def) |
|
205 |
||
| 37607 | 206 |
lemma prime_nat_code [code]: |
| 44872 | 207 |
"prime (p::nat) \<longleftrightarrow> p > 1 \<and> (\<forall>n \<in> {1<..<p}. ~ n dvd p)"
|
208 |
apply (simp add: Ball_def) |
|
209 |
apply (metis less_not_refl prime_nat_def dvd_triv_right not_prime_eq_prod_nat) |
|
210 |
done |
|
| 32007 | 211 |
|
212 |
lemma prime_nat_simp: |
|
| 44872 | 213 |
"prime (p::nat) \<longleftrightarrow> p > 1 \<and> (\<forall>n \<in> set [2..<p]. \<not> n dvd p)" |
214 |
by (auto simp add: prime_nat_code) |
|
| 32007 | 215 |
|
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45605
diff
changeset
|
216 |
lemmas prime_nat_simp_numeral [simp] = prime_nat_simp [of "numeral m"] for m |
| 32007 | 217 |
|
| 37607 | 218 |
lemma prime_int_code [code]: |
219 |
"prime (p::int) \<longleftrightarrow> p > 1 \<and> (\<forall>n \<in> {1<..<p}. ~ n dvd p)" (is "?L = ?R")
|
|
| 32007 | 220 |
proof |
| 44872 | 221 |
assume "?L" |
222 |
then show "?R" |
|
| 44821 | 223 |
by (clarsimp simp: prime_gt_1_int) (metis int_one_le_iff_zero_less prime_int_altdef less_le) |
| 32007 | 224 |
next |
| 44872 | 225 |
assume "?R" |
226 |
then show "?L" by (clarsimp simp: Ball_def) (metis dvdI not_prime_eq_prod_int) |
|
| 32007 | 227 |
qed |
228 |
||
| 44872 | 229 |
lemma prime_int_simp: "prime (p::int) \<longleftrightarrow> p > 1 \<and> (\<forall>n \<in> set [2..p - 1]. ~ n dvd p)" |
230 |
by (auto simp add: prime_int_code) |
|
| 32007 | 231 |
|
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45605
diff
changeset
|
232 |
lemmas prime_int_simp_numeral [simp] = prime_int_simp [of "numeral m"] for m |
| 32007 | 233 |
|
234 |
lemma two_is_prime_nat [simp]: "prime (2::nat)" |
|
| 44872 | 235 |
by simp |
| 32007 | 236 |
|
237 |
lemma two_is_prime_int [simp]: "prime (2::int)" |
|
| 44872 | 238 |
by simp |
| 32007 | 239 |
|
| 32111 | 240 |
text{* A bit of regression testing: *}
|
241 |
||
| 44872 | 242 |
lemma "prime(97::nat)" by simp |
243 |
lemma "prime(97::int)" by simp |
|
244 |
lemma "prime(997::nat)" by eval |
|
245 |
lemma "prime(997::int)" by eval |
|
| 32111 | 246 |
|
| 32007 | 247 |
|
248 |
lemma prime_imp_power_coprime_nat: "prime (p::nat) \<Longrightarrow> ~ p dvd a \<Longrightarrow> coprime a (p^m)" |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
249 |
apply (rule coprime_exp_nat) |
|
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
250 |
apply (subst gcd_commute_nat) |
|
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
251 |
apply (erule (1) prime_imp_coprime_nat) |
| 44872 | 252 |
done |
|
27568
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents:
27556
diff
changeset
|
253 |
|
| 32007 | 254 |
lemma prime_imp_power_coprime_int: "prime (p::int) \<Longrightarrow> ~ p dvd a \<Longrightarrow> coprime a (p^m)" |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
255 |
apply (rule coprime_exp_int) |
|
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
256 |
apply (subst gcd_commute_int) |
|
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
257 |
apply (erule (1) prime_imp_coprime_int) |
| 44872 | 258 |
done |
|
27568
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents:
27556
diff
changeset
|
259 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
260 |
lemma primes_coprime_nat: "prime (p::nat) \<Longrightarrow> prime q \<Longrightarrow> p \<noteq> q \<Longrightarrow> coprime p q" |
|
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
261 |
apply (rule prime_imp_coprime_nat, assumption) |
| 44872 | 262 |
apply (unfold prime_nat_def) |
263 |
apply auto |
|
264 |
done |
|
|
27568
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents:
27556
diff
changeset
|
265 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
266 |
lemma primes_coprime_int: "prime (p::int) \<Longrightarrow> prime q \<Longrightarrow> p \<noteq> q \<Longrightarrow> coprime p q" |
|
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
267 |
apply (rule prime_imp_coprime_int, assumption) |
| 40461 | 268 |
apply (unfold prime_int_altdef) |
| 44821 | 269 |
apply (metis int_one_le_iff_zero_less less_le) |
| 44872 | 270 |
done |
|
27568
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents:
27556
diff
changeset
|
271 |
|
| 44872 | 272 |
lemma primes_imp_powers_coprime_nat: |
273 |
"prime (p::nat) \<Longrightarrow> prime q \<Longrightarrow> p ~= q \<Longrightarrow> coprime (p^m) (q^n)" |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
274 |
by (rule coprime_exp2_nat, rule primes_coprime_nat) |
|
27568
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents:
27556
diff
changeset
|
275 |
|
| 44872 | 276 |
lemma primes_imp_powers_coprime_int: |
277 |
"prime (p::int) \<Longrightarrow> prime q \<Longrightarrow> p ~= q \<Longrightarrow> coprime (p^m) (q^n)" |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
278 |
by (rule coprime_exp2_int, rule primes_coprime_int) |
|
27568
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents:
27556
diff
changeset
|
279 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
280 |
lemma prime_factor_nat: "n \<noteq> (1::nat) \<Longrightarrow> \<exists> p. prime p \<and> p dvd n" |
| 31706 | 281 |
apply (induct n rule: nat_less_induct) |
282 |
apply (case_tac "n = 0") |
|
| 44872 | 283 |
using two_is_prime_nat |
284 |
apply blast |
|
285 |
apply (metis One_nat_def dvd.order_trans dvd_refl less_Suc0 linorder_neqE_nat |
|
286 |
nat_dvd_not_less neq0_conv prime_nat_def) |
|
287 |
done |
|
|
23244
1630951f0512
added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents:
22367
diff
changeset
|
288 |
|
| 31706 | 289 |
(* An Isar version: |
290 |
||
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
291 |
lemma prime_factor_b_nat: |
| 31706 | 292 |
fixes n :: nat |
293 |
assumes "n \<noteq> 1" |
|
294 |
shows "\<exists>p. prime p \<and> p dvd n" |
|
| 23983 | 295 |
|
| 31706 | 296 |
using `n ~= 1` |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
297 |
proof (induct n rule: less_induct_nat) |
| 31706 | 298 |
fix n :: nat |
299 |
assume "n ~= 1" and |
|
300 |
ih: "\<forall>m<n. m \<noteq> 1 \<longrightarrow> (\<exists>p. prime p \<and> p dvd m)" |
|
| 44872 | 301 |
then show "\<exists>p. prime p \<and> p dvd n" |
| 31706 | 302 |
proof - |
303 |
{
|
|
304 |
assume "n = 0" |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
305 |
moreover note two_is_prime_nat |
| 31706 | 306 |
ultimately have ?thesis |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
307 |
by (auto simp del: two_is_prime_nat) |
| 31706 | 308 |
} |
309 |
moreover |
|
310 |
{
|
|
311 |
assume "prime n" |
|
| 44872 | 312 |
then have ?thesis by auto |
| 31706 | 313 |
} |
314 |
moreover |
|
315 |
{
|
|
316 |
assume "n ~= 0" and "~ prime n" |
|
317 |
with `n ~= 1` have "n > 1" by auto |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
318 |
with `~ prime n` and not_prime_eq_prod_nat obtain m k where |
| 31706 | 319 |
"n = m * k" and "1 < m" and "m < n" by blast |
320 |
with ih obtain p where "prime p" and "p dvd m" by blast |
|
321 |
with `n = m * k` have ?thesis by auto |
|
322 |
} |
|
323 |
ultimately show ?thesis by blast |
|
324 |
qed |
|
| 23983 | 325 |
qed |
326 |
||
| 31706 | 327 |
*) |
328 |
||
329 |
text {* One property of coprimality is easier to prove via prime factors. *}
|
|
330 |
||
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
331 |
lemma prime_divprod_pow_nat: |
| 31706 | 332 |
assumes p: "prime (p::nat)" and ab: "coprime a b" and pab: "p^n dvd a * b" |
333 |
shows "p^n dvd a \<or> p^n dvd b" |
|
334 |
proof- |
|
| 44872 | 335 |
{ assume "n = 0 \<or> a = 1 \<or> b = 1" with pab have ?thesis
|
| 31706 | 336 |
apply (cases "n=0", simp_all) |
| 44872 | 337 |
apply (cases "a=1", simp_all) |
338 |
done } |
|
| 31706 | 339 |
moreover |
| 44872 | 340 |
{ assume n: "n \<noteq> 0" and a: "a\<noteq>1" and b: "b\<noteq>1"
|
341 |
then obtain m where m: "n = Suc m" by (cases n) auto |
|
342 |
from n have "p dvd p^n" apply (intro dvd_power) apply auto done |
|
| 31706 | 343 |
also note pab |
344 |
finally have pab': "p dvd a * b". |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
345 |
from prime_dvd_mult_nat[OF p pab'] |
| 31706 | 346 |
have "p dvd a \<or> p dvd b" . |
347 |
moreover |
|
| 33946 | 348 |
{ assume pa: "p dvd a"
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
349 |
from coprime_common_divisor_nat [OF ab, OF pa] p have "\<not> p dvd b" by auto |
| 31706 | 350 |
with p have "coprime b p" |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
351 |
by (subst gcd_commute_nat, intro prime_imp_coprime_nat) |
| 44872 | 352 |
then have pnb: "coprime (p^n) b" |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
353 |
by (subst gcd_commute_nat, rule coprime_exp_nat) |
| 33946 | 354 |
from coprime_dvd_mult_nat[OF pnb pab] have ?thesis by blast } |
| 31706 | 355 |
moreover |
| 33946 | 356 |
{ assume pb: "p dvd b"
|
| 31706 | 357 |
have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute) |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
358 |
from coprime_common_divisor_nat [OF ab, of p] pb p have "\<not> p dvd a" |
| 31706 | 359 |
by auto |
360 |
with p have "coprime a p" |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
361 |
by (subst gcd_commute_nat, intro prime_imp_coprime_nat) |
| 44872 | 362 |
then have pna: "coprime (p^n) a" |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31814
diff
changeset
|
363 |
by (subst gcd_commute_nat, rule coprime_exp_nat) |
| 33946 | 364 |
from coprime_dvd_mult_nat[OF pna pnba] have ?thesis by blast } |
| 44872 | 365 |
ultimately have ?thesis by blast } |
| 31706 | 366 |
ultimately show ?thesis by blast |
| 23983 | 367 |
qed |
368 |
||
| 44872 | 369 |
|
|
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
370 |
subsection {* Infinitely many primes *}
|
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
371 |
|
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
372 |
lemma next_prime_bound: "\<exists>(p::nat). prime p \<and> n < p \<and> p <= fact n + 1" |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
373 |
proof- |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
374 |
have f1: "fact n + 1 \<noteq> 1" using fact_ge_one_nat [of n] by arith |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
375 |
from prime_factor_nat [OF f1] |
| 44872 | 376 |
obtain p where "prime p" and "p dvd fact n + 1" by auto |
377 |
then have "p \<le> fact n + 1" apply (intro dvd_imp_le) apply auto done |
|
378 |
{ assume "p \<le> n"
|
|
|
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
379 |
from `prime p` have "p \<ge> 1" |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
380 |
by (cases p, simp_all) |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
381 |
with `p <= n` have "p dvd fact n" |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
382 |
by (intro dvd_fact_nat) |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
383 |
with `p dvd fact n + 1` have "p dvd fact n + 1 - fact n" |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
384 |
by (rule dvd_diff_nat) |
| 44872 | 385 |
then have "p dvd 1" by simp |
386 |
then have "p <= 1" by auto |
|
|
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
387 |
moreover from `prime p` have "p > 1" by auto |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
388 |
ultimately have False by auto} |
| 44872 | 389 |
then have "n < p" by presburger |
|
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
390 |
with `prime p` and `p <= fact n + 1` show ?thesis by auto |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
391 |
qed |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
392 |
|
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
393 |
lemma bigger_prime: "\<exists>p. prime p \<and> p > (n::nat)" |
| 44872 | 394 |
using next_prime_bound by auto |
|
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
395 |
|
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
396 |
lemma primes_infinite: "\<not> (finite {(p::nat). prime p})"
|
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
397 |
proof |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
398 |
assume "finite {(p::nat). prime p}"
|
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
399 |
with Max_ge have "(EX b. (ALL x : {(p::nat). prime p}. x <= b))"
|
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
400 |
by auto |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
401 |
then obtain b where "ALL (x::nat). prime x \<longrightarrow> x <= b" |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
402 |
by auto |
| 44872 | 403 |
with bigger_prime [of b] show False |
404 |
by auto |
|
|
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
405 |
qed |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
31952
diff
changeset
|
406 |
|
| 21256 | 407 |
end |