| author | wenzelm |
| Tue, 07 Nov 2017 15:50:36 +0100 | |
| changeset 67024 | 72d37a2e9cca |
| parent 66954 | 0230af0f3c59 |
| child 67051 | e7e54a0b9197 |
| permissions | -rw-r--r-- |
| 41959 | 1 |
(* Title: HOL/Number_Theory/Cong.thy |
| 66380 | 2 |
Author: Christophe Tabacznyj |
3 |
Author: Lawrence C. Paulson |
|
4 |
Author: Amine Chaieb |
|
5 |
Author: Thomas M. Rasmussen |
|
6 |
Author: Jeremy Avigad |
|
| 31719 | 7 |
|
8 |
Defines congruence (notation: [x = y] (mod z)) for natural numbers and |
|
9 |
integers. |
|
10 |
||
11 |
This file combines and revises a number of prior developments. |
|
12 |
||
13 |
The original theories "GCD" and "Primes" were by Christophe Tabacznyj |
|
| 58623 | 14 |
and Lawrence C. Paulson, based on @{cite davenport92}. They introduced
|
| 31719 | 15 |
gcd, lcm, and prime for the natural numbers. |
16 |
||
17 |
The original theory "IntPrimes" was by Thomas M. Rasmussen, and |
|
18 |
extended gcd, lcm, primes to the integers. Amine Chaieb provided |
|
19 |
another extension of the notions to the integers, and added a number |
|
| 44872 | 20 |
of results to "Primes" and "GCD". |
| 31719 | 21 |
|
22 |
The original theory, "IntPrimes", by Thomas M. Rasmussen, defined and |
|
23 |
developed the congruence relations on the integers. The notion was |
|
| 33718 | 24 |
extended to the natural numbers by Chaieb. Jeremy Avigad combined |
| 31719 | 25 |
these, revised and tidied them, made the development uniform for the |
26 |
natural numbers and the integers, and added a number of new theorems. |
|
27 |
*) |
|
28 |
||
| 60526 | 29 |
section \<open>Congruence\<close> |
| 31719 | 30 |
|
31 |
theory Cong |
|
|
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
66380
diff
changeset
|
32 |
imports "HOL-Computational_Algebra.Primes" |
| 31719 | 33 |
begin |
34 |
||
| 66888 | 35 |
subsection \<open>Generic congruences\<close> |
36 |
||
37 |
context unique_euclidean_semiring |
|
| 31719 | 38 |
begin |
39 |
||
| 66888 | 40 |
definition cong :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" ("(1[_ = _] '(()mod _'))")
|
41 |
where "cong b c a \<longleftrightarrow> b mod a = c mod a" |
|
42 |
||
| 58937 | 43 |
abbreviation notcong :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" ("(1[_ \<noteq> _] '(()mod _'))")
|
| 66888 | 44 |
where "notcong b c a \<equiv> \<not> cong b c a" |
45 |
||
46 |
lemma cong_mod_left [simp]: |
|
47 |
"[b mod a = c] (mod a) \<longleftrightarrow> [b = c] (mod a)" |
|
48 |
by (simp add: cong_def) |
|
49 |
||
50 |
lemma cong_mod_right [simp]: |
|
51 |
"[b = c mod a] (mod a) \<longleftrightarrow> [b = c] (mod a)" |
|
52 |
by (simp add: cong_def) |
|
53 |
||
54 |
lemma cong_dvd_iff: |
|
55 |
"a dvd b \<longleftrightarrow> a dvd c" if "[b = c] (mod a)" |
|
56 |
using that by (auto simp: cong_def dvd_eq_mod_eq_0) |
|
57 |
||
58 |
lemma cong_0 [simp, presburger]: |
|
59 |
"[b = c] (mod 0) \<longleftrightarrow> b = c" |
|
60 |
by (simp add: cong_def) |
|
61 |
||
62 |
lemma cong_1 [simp, presburger]: |
|
63 |
"[b = c] (mod 1)" |
|
64 |
by (simp add: cong_def) |
|
65 |
||
66 |
lemma cong_refl [simp]: |
|
67 |
"[b = b] (mod a)" |
|
68 |
by (simp add: cong_def) |
|
69 |
||
70 |
lemma cong_sym: |
|
71 |
"[b = c] (mod a) \<Longrightarrow> [c = b] (mod a)" |
|
72 |
by (simp add: cong_def) |
|
73 |
||
74 |
lemma cong_sym_eq: |
|
75 |
"[b = c] (mod a) \<longleftrightarrow> [c = b] (mod a)" |
|
76 |
by (auto simp add: cong_def) |
|
77 |
||
78 |
lemma cong_trans [trans]: |
|
79 |
"[b = c] (mod a) \<Longrightarrow> [c = d] (mod a) \<Longrightarrow> [b = d] (mod a)" |
|
80 |
by (simp add: cong_def) |
|
81 |
||
82 |
lemma cong_add: |
|
83 |
"[b = c] (mod a) \<Longrightarrow> [d = e] (mod a) \<Longrightarrow> [b + d = c + e] (mod a)" |
|
84 |
by (auto simp add: cong_def intro: mod_add_cong) |
|
85 |
||
86 |
lemma cong_mult: |
|
87 |
"[b = c] (mod a) \<Longrightarrow> [d = e] (mod a) \<Longrightarrow> [b * d = c * e] (mod a)" |
|
88 |
by (auto simp add: cong_def intro: mod_mult_cong) |
|
89 |
||
90 |
lemma cong_pow: |
|
91 |
"[b = c] (mod a) \<Longrightarrow> [b ^ n = c ^ n] (mod a)" |
|
92 |
by (simp add: cong_def power_mod [symmetric, of b n a] power_mod [symmetric, of c n a]) |
|
93 |
||
94 |
lemma cong_sum: |
|
95 |
"[sum f A = sum g A] (mod a)" if "\<And>x. x \<in> A \<Longrightarrow> [f x = g x] (mod a)" |
|
96 |
using that by (induct A rule: infinite_finite_induct) (auto intro: cong_add) |
|
97 |
||
98 |
lemma cong_prod: |
|
99 |
"[prod f A = prod g A] (mod a)" if "(\<And>x. x \<in> A \<Longrightarrow> [f x = g x] (mod a))" |
|
100 |
using that by (induct A rule: infinite_finite_induct) (auto intro: cong_mult) |
|
101 |
||
102 |
lemma cong_scalar_right: |
|
103 |
"[b = c] (mod a) \<Longrightarrow> [b * d = c * d] (mod a)" |
|
104 |
by (simp add: cong_mult) |
|
105 |
||
106 |
lemma cong_scalar_left: |
|
107 |
"[b = c] (mod a) \<Longrightarrow> [d * b = d * c] (mod a)" |
|
108 |
by (simp add: cong_mult) |
|
109 |
||
110 |
lemma cong_mult_self_right: "[b * a = 0] (mod a)" |
|
111 |
by (simp add: cong_def) |
|
112 |
||
113 |
lemma cong_mult_self_left: "[a * b = 0] (mod a)" |
|
114 |
by (simp add: cong_def) |
|
115 |
||
116 |
lemma cong_0_iff: "[b = 0] (mod a) \<longleftrightarrow> a dvd b" |
|
117 |
by (simp add: cong_def dvd_eq_mod_eq_0) |
|
118 |
||
119 |
lemma mod_mult_cong_right: |
|
120 |
"[c mod (a * b) = d] (mod a) \<longleftrightarrow> [c = d] (mod a)" |
|
121 |
by (simp add: cong_def mod_mod_cancel mod_add_left_eq) |
|
122 |
||
123 |
lemma mod_mult_cong_left: |
|
124 |
"[c mod (b * a) = d] (mod a) \<longleftrightarrow> [c = d] (mod a)" |
|
125 |
using mod_mult_cong_right [of c a b d] by (simp add: ac_simps) |
|
126 |
||
127 |
end |
|
128 |
||
129 |
context unique_euclidean_ring |
|
130 |
begin |
|
131 |
||
132 |
lemma cong_diff: |
|
133 |
"[b = c] (mod a) \<Longrightarrow> [d = e] (mod a) \<Longrightarrow> [b - d = c - e] (mod a)" |
|
134 |
by (auto simp add: cong_def intro: mod_diff_cong) |
|
135 |
||
136 |
lemma cong_diff_iff_cong_0: |
|
137 |
"[b - c = 0] (mod a) \<longleftrightarrow> [b = c] (mod a)" (is "?P \<longleftrightarrow> ?Q") |
|
138 |
proof |
|
139 |
assume ?P |
|
140 |
then have "[b - c + c = 0 + c] (mod a)" |
|
141 |
by (rule cong_add) simp |
|
142 |
then show ?Q |
|
143 |
by simp |
|
144 |
next |
|
145 |
assume ?Q |
|
146 |
with cong_diff [of b c a c c] show ?P |
|
147 |
by simp |
|
148 |
qed |
|
149 |
||
150 |
lemma cong_minus_minus_iff: |
|
151 |
"[- b = - c] (mod a) \<longleftrightarrow> [b = c] (mod a)" |
|
152 |
using cong_diff_iff_cong_0 [of b c a] cong_diff_iff_cong_0 [of "- b" "- c" a] |
|
153 |
by (simp add: cong_0_iff dvd_diff_commute) |
|
154 |
||
155 |
lemma cong_modulus_minus_iff: "[b = c] (mod - a) \<longleftrightarrow> [b = c] (mod a)" |
|
156 |
using cong_diff_iff_cong_0 [of b c a] cong_diff_iff_cong_0 [of b c " -a"] |
|
157 |
by (simp add: cong_0_iff) |
|
| 31719 | 158 |
|
159 |
end |
|
160 |
||
| 66380 | 161 |
|
| 66888 | 162 |
subsection \<open>Congruences on @{typ nat} and @{typ int}\<close>
|
| 31719 | 163 |
|
| 66888 | 164 |
lemma cong_int_iff: |
165 |
"[int m = int q] (mod int n) \<longleftrightarrow> [m = q] (mod n)" |
|
166 |
by (simp add: cong_def of_nat_mod [symmetric]) |
|
167 |
||
168 |
lemma cong_Suc_0 [simp, presburger]: |
|
169 |
"[m = n] (mod Suc 0)" |
|
170 |
using cong_1 [of m n] by simp |
|
| 31719 | 171 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
172 |
lemma cong_diff_nat: |
| 66888 | 173 |
"[a - c = b - d] (mod m)" if "[a = b] (mod m)" "[c = d] (mod m)" |
| 66954 | 174 |
and "a \<ge> c" "b \<ge> d" for a b c d m :: nat |
175 |
proof - |
|
176 |
have "[c + (a - c) = d + (b - d)] (mod m)" |
|
177 |
using that by simp |
|
178 |
with \<open>[c = d] (mod m)\<close> have "[c + (a - c) = c + (b - d)] (mod m)" |
|
179 |
using mod_add_cong by (auto simp add: cong_def) fastforce |
|
180 |
then show ?thesis |
|
181 |
by (simp add: cong_def nat_mod_eq_iff) |
|
182 |
qed |
|
| 31719 | 183 |
|
| 66888 | 184 |
lemma cong_diff_iff_cong_0_nat: |
| 66954 | 185 |
"[a - b = 0] (mod m) \<longleftrightarrow> [a = b] (mod m)" if "a \<ge> b" for a b :: nat |
186 |
using that by (auto simp add: cong_def le_imp_diff_is_add dest: nat_mod_eq_lemma) |
|
| 31719 | 187 |
|
| 66954 | 188 |
lemma cong_diff_iff_cong_0_nat': |
189 |
"[nat \<bar>int a - int b\<bar> = 0] (mod m) \<longleftrightarrow> [a = b] (mod m)" |
|
190 |
proof (cases "b \<le> a") |
|
191 |
case True |
|
192 |
then show ?thesis |
|
193 |
by (simp add: nat_diff_distrib' cong_diff_iff_cong_0_nat [of b a m]) |
|
194 |
next |
|
195 |
case False |
|
196 |
then have "a \<le> b" |
|
197 |
by simp |
|
198 |
then show ?thesis |
|
199 |
by (simp add: nat_diff_distrib' cong_diff_iff_cong_0_nat [of a b m]) |
|
200 |
(auto simp add: cong_def) |
|
201 |
qed |
|
202 |
||
203 |
lemma cong_altdef_nat: |
|
204 |
"a \<ge> b \<Longrightarrow> [a = b] (mod m) \<longleftrightarrow> m dvd (a - b)" |
|
| 66380 | 205 |
for a b :: nat |
| 66888 | 206 |
by (simp add: cong_0_iff [symmetric] cong_diff_iff_cong_0_nat) |
| 31719 | 207 |
|
| 66954 | 208 |
lemma cong_altdef_nat': |
209 |
"[a = b] (mod m) \<longleftrightarrow> m dvd nat \<bar>int a - int b\<bar>" |
|
210 |
by (simp add: cong_0_iff [symmetric] cong_diff_iff_cong_0_nat') |
|
211 |
||
212 |
lemma cong_altdef_int: |
|
213 |
"[a = b] (mod m) \<longleftrightarrow> m dvd (a - b)" |
|
| 66380 | 214 |
for a b :: int |
| 66888 | 215 |
by (simp add: cong_0_iff [symmetric] cong_diff_iff_cong_0) |
| 31719 | 216 |
|
| 66888 | 217 |
lemma cong_abs_int [simp]: "[x = y] (mod abs m) \<longleftrightarrow> [x = y] (mod m)" |
| 66380 | 218 |
for x y :: int |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
219 |
by (simp add: cong_altdef_int) |
| 31719 | 220 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
221 |
lemma cong_square_int: |
| 66380 | 222 |
"prime p \<Longrightarrow> 0 < a \<Longrightarrow> [a * a = 1] (mod p) \<Longrightarrow> [a = 1] (mod p) \<or> [a = - 1] (mod p)" |
223 |
for a :: int |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
224 |
apply (simp only: cong_altdef_int) |
|
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
225 |
apply (subst prime_dvd_mult_eq_int [symmetric], assumption) |
| 36350 | 226 |
apply (auto simp add: field_simps) |
| 44872 | 227 |
done |
| 31719 | 228 |
|
| 66954 | 229 |
lemma cong_mult_rcancel_int: |
230 |
"[a * k = b * k] (mod m) \<longleftrightarrow> [a = b] (mod m)" |
|
231 |
if "coprime k m" for a k m :: int |
|
232 |
by (metis that cong_altdef_int left_diff_distrib coprime_dvd_mult_iff gcd.commute) |
|
| 31719 | 233 |
|
| 66954 | 234 |
lemma cong_mult_rcancel_nat: |
235 |
"[a * k = b * k] (mod m) \<longleftrightarrow> [a = b] (mod m)" |
|
236 |
if "coprime k m" for a k m :: nat |
|
237 |
proof - |
|
238 |
have "[a * k = b * k] (mod m) \<longleftrightarrow> m dvd nat \<bar>int (a * k) - int (b * k)\<bar>" |
|
239 |
by (simp add: cong_altdef_nat') |
|
240 |
also have "\<dots> \<longleftrightarrow> m dvd nat \<bar>(int a - int b) * int k\<bar>" |
|
241 |
by (simp add: algebra_simps) |
|
242 |
also have "\<dots> \<longleftrightarrow> m dvd nat \<bar>int a - int b\<bar> * k" |
|
243 |
by (simp add: abs_mult nat_times_as_int) |
|
244 |
also have "\<dots> \<longleftrightarrow> m dvd nat \<bar>int a - int b\<bar>" |
|
245 |
by (rule coprime_dvd_mult_iff) (use \<open>coprime k m\<close> in \<open>simp add: ac_simps\<close>) |
|
246 |
also have "\<dots> \<longleftrightarrow> [a = b] (mod m)" |
|
247 |
by (simp add: cong_altdef_nat') |
|
248 |
finally show ?thesis . |
|
249 |
qed |
|
| 31719 | 250 |
|
| 66954 | 251 |
lemma cong_mult_lcancel_int: |
252 |
"[k * a = k * b] (mod m) = [a = b] (mod m)" |
|
253 |
if "coprime k m" for a k m :: int |
|
254 |
using that by (simp add: cong_mult_rcancel_int ac_simps) |
|
255 |
||
256 |
lemma cong_mult_lcancel_nat: |
|
257 |
"[k * a = k * b] (mod m) = [a = b] (mod m)" |
|
258 |
if "coprime k m" for a k m :: nat |
|
259 |
using that by (simp add: cong_mult_rcancel_nat ac_simps) |
|
| 31719 | 260 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
261 |
lemma coprime_cong_mult_int: |
| 66380 | 262 |
"[a = b] (mod m) \<Longrightarrow> [a = b] (mod n) \<Longrightarrow> coprime m n \<Longrightarrow> [a = b] (mod m * n)" |
263 |
for a b :: int |
|
| 66954 | 264 |
by (simp add: cong_altdef_int divides_mult) |
| 31719 | 265 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
266 |
lemma coprime_cong_mult_nat: |
| 66380 | 267 |
"[a = b] (mod m) \<Longrightarrow> [a = b] (mod n) \<Longrightarrow> coprime m n \<Longrightarrow> [a = b] (mod m * n)" |
268 |
for a b :: nat |
|
| 66954 | 269 |
by (simp add: cong_altdef_nat' divides_mult) |
| 31719 | 270 |
|
| 66380 | 271 |
lemma cong_less_imp_eq_nat: "0 \<le> a \<Longrightarrow> a < m \<Longrightarrow> 0 \<le> b \<Longrightarrow> b < m \<Longrightarrow> [a = b] (mod m) \<Longrightarrow> a = b" |
272 |
for a b :: nat |
|
| 66888 | 273 |
by (auto simp add: cong_def) |
| 31719 | 274 |
|
| 66380 | 275 |
lemma cong_less_imp_eq_int: "0 \<le> a \<Longrightarrow> a < m \<Longrightarrow> 0 \<le> b \<Longrightarrow> b < m \<Longrightarrow> [a = b] (mod m) \<Longrightarrow> a = b" |
276 |
for a b :: int |
|
| 66888 | 277 |
by (auto simp add: cong_def) |
| 31719 | 278 |
|
| 66380 | 279 |
lemma cong_less_unique_nat: "0 < m \<Longrightarrow> (\<exists>!b. 0 \<le> b \<and> b < m \<and> [a = b] (mod m))" |
280 |
for a m :: nat |
|
| 66888 | 281 |
by (auto simp: cong_def) (metis mod_less_divisor mod_mod_trivial) |
| 31719 | 282 |
|
| 66380 | 283 |
lemma cong_less_unique_int: "0 < m \<Longrightarrow> (\<exists>!b. 0 \<le> b \<and> b < m \<and> [a = b] (mod m))" |
284 |
for a m :: int |
|
| 66888 | 285 |
by (auto simp: cong_def) (metis mod_mod_trivial pos_mod_conj) |
| 31719 | 286 |
|
| 66380 | 287 |
lemma cong_iff_lin_int: "[a = b] (mod m) \<longleftrightarrow> (\<exists>k. b = a + m * k)" |
288 |
for a b :: int |
|
| 66837 | 289 |
by (auto simp add: cong_altdef_int algebra_simps elim!: dvdE) |
290 |
(simp add: distrib_right [symmetric] add_eq_0_iff) |
|
| 31719 | 291 |
|
| 66380 | 292 |
lemma cong_iff_lin_nat: "([a = b] (mod m)) \<longleftrightarrow> (\<exists>k1 k2. b + k1 * m = a + k2 * m)" |
293 |
(is "?lhs = ?rhs") |
|
294 |
for a b :: nat |
|
295 |
proof |
|
296 |
assume ?lhs |
|
| 55371 | 297 |
show ?rhs |
298 |
proof (cases "b \<le> a") |
|
299 |
case True |
|
| 66380 | 300 |
with \<open>?lhs\<close> show ?rhs |
|
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57418
diff
changeset
|
301 |
by (metis cong_altdef_nat dvd_def le_add_diff_inverse add_0_right mult_0 mult.commute) |
| 55371 | 302 |
next |
303 |
case False |
|
| 66380 | 304 |
with \<open>?lhs\<close> show ?rhs |
| 66888 | 305 |
by (metis cong_def mult.commute nat_le_linear nat_mod_eq_lemma) |
| 55371 | 306 |
qed |
307 |
next |
|
308 |
assume ?rhs |
|
309 |
then show ?lhs |
|
| 66888 | 310 |
by (metis cong_def mult.commute nat_mod_eq_iff) |
| 55371 | 311 |
qed |
| 31719 | 312 |
|
| 66954 | 313 |
lemma cong_gcd_eq_nat: "[a = b] (mod m) \<Longrightarrow> gcd a m = gcd b m" |
314 |
for a b :: nat |
|
315 |
by (auto simp add: cong_def) (metis gcd_red_nat) |
|
316 |
||
| 66380 | 317 |
lemma cong_gcd_eq_int: "[a = b] (mod m) \<Longrightarrow> gcd a m = gcd b m" |
318 |
for a b :: int |
|
| 66888 | 319 |
by (auto simp add: cong_def) (metis gcd_red_int) |
| 31719 | 320 |
|
| 66380 | 321 |
lemma cong_imp_coprime_nat: "[a = b] (mod m) \<Longrightarrow> coprime a m \<Longrightarrow> coprime b m" |
322 |
for a b :: nat |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
323 |
by (auto simp add: cong_gcd_eq_nat) |
| 31719 | 324 |
|
| 66380 | 325 |
lemma cong_imp_coprime_int: "[a = b] (mod m) \<Longrightarrow> coprime a m \<Longrightarrow> coprime b m" |
326 |
for a b :: int |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
327 |
by (auto simp add: cong_gcd_eq_int) |
| 31719 | 328 |
|
| 66380 | 329 |
lemma cong_cong_mod_nat: "[a = b] (mod m) \<longleftrightarrow> [a mod m = b mod m] (mod m)" |
330 |
for a b :: nat |
|
| 66888 | 331 |
by simp |
| 31719 | 332 |
|
| 66380 | 333 |
lemma cong_cong_mod_int: "[a = b] (mod m) \<longleftrightarrow> [a mod m = b mod m] (mod m)" |
334 |
for a b :: int |
|
| 66888 | 335 |
by simp |
| 31719 | 336 |
|
| 66380 | 337 |
lemma cong_minus_int [iff]: "[a = b] (mod - m) \<longleftrightarrow> [a = b] (mod m)" |
338 |
for a b :: int |
|
| 55371 | 339 |
by (metis cong_iff_lin_int minus_equation_iff mult_minus_left mult_minus_right) |
| 31719 | 340 |
|
341 |
(* |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
342 |
lemma mod_dvd_mod_int: |
| 31719 | 343 |
"0 < (m::int) \<Longrightarrow> m dvd b \<Longrightarrow> (a mod b mod m) = (a mod m)" |
344 |
apply (unfold dvd_def, auto) |
|
345 |
apply (rule mod_mod_cancel) |
|
346 |
apply auto |
|
| 44872 | 347 |
done |
| 31719 | 348 |
|
349 |
lemma mod_dvd_mod: |
|
350 |
assumes "0 < (m::nat)" and "m dvd b" |
|
351 |
shows "(a mod b mod m) = (a mod m)" |
|
352 |
||
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
353 |
apply (rule mod_dvd_mod_int [transferred]) |
| 41541 | 354 |
using assms apply auto |
355 |
done |
|
| 31719 | 356 |
*) |
357 |
||
| 66380 | 358 |
lemma cong_add_lcancel_nat: "[a + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)" |
359 |
for a x y :: nat |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
360 |
by (simp add: cong_iff_lin_nat) |
| 31719 | 361 |
|
| 66380 | 362 |
lemma cong_add_lcancel_int: "[a + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)" |
363 |
for a x y :: int |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
364 |
by (simp add: cong_iff_lin_int) |
| 31719 | 365 |
|
| 66380 | 366 |
lemma cong_add_rcancel_nat: "[x + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)" |
367 |
for a x y :: nat |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
368 |
by (simp add: cong_iff_lin_nat) |
| 31719 | 369 |
|
| 66380 | 370 |
lemma cong_add_rcancel_int: "[x + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)" |
371 |
for a x y :: int |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
372 |
by (simp add: cong_iff_lin_int) |
| 31719 | 373 |
|
| 66380 | 374 |
lemma cong_add_lcancel_0_nat: "[a + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)" |
375 |
for a x :: nat |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
376 |
by (simp add: cong_iff_lin_nat) |
| 31719 | 377 |
|
| 66380 | 378 |
lemma cong_add_lcancel_0_int: "[a + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)" |
379 |
for a x :: int |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
380 |
by (simp add: cong_iff_lin_int) |
| 31719 | 381 |
|
| 66380 | 382 |
lemma cong_add_rcancel_0_nat: "[x + a = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)" |
383 |
for a x :: nat |
|
384 |
by (simp add: cong_iff_lin_nat) |
|
385 |
||
386 |
lemma cong_add_rcancel_0_int: "[x + a = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)" |
|
387 |
for a x :: int |
|
388 |
by (simp add: cong_iff_lin_int) |
|
389 |
||
390 |
lemma cong_dvd_modulus_nat: "[x = y] (mod m) \<Longrightarrow> n dvd m \<Longrightarrow> [x = y] (mod n)" |
|
391 |
for x y :: nat |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
392 |
apply (auto simp add: cong_iff_lin_nat dvd_def) |
| 66380 | 393 |
apply (rule_tac x= "k1 * k" in exI) |
394 |
apply (rule_tac x= "k2 * k" in exI) |
|
| 36350 | 395 |
apply (simp add: field_simps) |
| 44872 | 396 |
done |
| 31719 | 397 |
|
| 66380 | 398 |
lemma cong_dvd_modulus_int: "[x = y] (mod m) \<Longrightarrow> n dvd m \<Longrightarrow> [x = y] (mod n)" |
399 |
for x y :: int |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
400 |
by (auto simp add: cong_altdef_int dvd_def) |
| 31719 | 401 |
|
| 66380 | 402 |
lemma cong_to_1_nat: |
403 |
fixes a :: nat |
|
404 |
assumes "[a = 1] (mod n)" |
|
405 |
shows "n dvd (a - 1)" |
|
406 |
proof (cases "a = 0") |
|
407 |
case True |
|
408 |
then show ?thesis by force |
|
409 |
next |
|
410 |
case False |
|
411 |
with assms show ?thesis by (metis cong_altdef_nat leI less_one) |
|
412 |
qed |
|
413 |
||
414 |
lemma cong_0_1_nat': "[0 = Suc 0] (mod n) \<longleftrightarrow> n = Suc 0" |
|
| 66888 | 415 |
by (auto simp: cong_def) |
| 66380 | 416 |
|
417 |
lemma cong_0_1_nat: "[0 = 1] (mod n) \<longleftrightarrow> n = 1" |
|
418 |
for n :: nat |
|
| 66888 | 419 |
by (auto simp: cong_def) |
| 66380 | 420 |
|
421 |
lemma cong_0_1_int: "[0 = 1] (mod n) \<longleftrightarrow> n = 1 \<or> n = - 1" |
|
422 |
for n :: int |
|
| 66888 | 423 |
by (auto simp: cong_def zmult_eq_1_iff) |
| 66380 | 424 |
|
425 |
lemma cong_to_1'_nat: "[a = 1] (mod n) \<longleftrightarrow> a = 0 \<and> n = 1 \<or> (\<exists>m. a = 1 + m * n)" |
|
426 |
for a :: nat |
|
427 |
by (metis add.right_neutral cong_0_1_nat cong_iff_lin_nat cong_to_1_nat |
|
428 |
dvd_div_mult_self leI le_add_diff_inverse less_one mult_eq_if) |
|
429 |
||
430 |
lemma cong_le_nat: "y \<le> x \<Longrightarrow> [x = y] (mod n) \<longleftrightarrow> (\<exists>q. x = q * n + y)" |
|
431 |
for x y :: nat |
|
| 66837 | 432 |
by (auto simp add: cong_altdef_nat le_imp_diff_is_add elim!: dvdE) |
| 66380 | 433 |
|
434 |
lemma cong_solve_nat: |
|
435 |
fixes a :: nat |
|
436 |
assumes "a \<noteq> 0" |
|
437 |
shows "\<exists>x. [a * x = gcd a n] (mod n)" |
|
438 |
proof (cases "n = 0") |
|
439 |
case True |
|
440 |
then show ?thesis by force |
|
441 |
next |
|
442 |
case False |
|
443 |
then show ?thesis |
|
444 |
using bezout_nat [of a n, OF \<open>a \<noteq> 0\<close>] |
|
| 66888 | 445 |
by auto (metis cong_add_rcancel_0_nat cong_mult_self_left) |
| 66380 | 446 |
qed |
447 |
||
448 |
lemma cong_solve_int: "a \<noteq> 0 \<Longrightarrow> \<exists>x. [a * x = gcd a n] (mod n)" |
|
449 |
for a :: int |
|
450 |
apply (cases "n = 0") |
|
451 |
apply (cases "a \<ge> 0") |
|
452 |
apply auto |
|
453 |
apply (rule_tac x = "-1" in exI) |
|
454 |
apply auto |
|
455 |
apply (insert bezout_int [of a n], auto) |
|
456 |
apply (metis cong_iff_lin_int mult.commute) |
|
| 44872 | 457 |
done |
| 31719 | 458 |
|
| 44872 | 459 |
lemma cong_solve_dvd_nat: |
| 66380 | 460 |
fixes a :: nat |
461 |
assumes a: "a \<noteq> 0" and b: "gcd a n dvd d" |
|
462 |
shows "\<exists>x. [a * x = d] (mod n)" |
|
| 31719 | 463 |
proof - |
| 44872 | 464 |
from cong_solve_nat [OF a] obtain x where "[a * x = gcd a n](mod n)" |
| 31719 | 465 |
by auto |
| 44872 | 466 |
then have "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)" |
| 66888 | 467 |
using cong_scalar_left by blast |
| 31719 | 468 |
also from b have "(d div gcd a n) * gcd a n = d" |
469 |
by (rule dvd_div_mult_self) |
|
470 |
also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)" |
|
471 |
by auto |
|
472 |
finally show ?thesis |
|
473 |
by auto |
|
474 |
qed |
|
475 |
||
| 44872 | 476 |
lemma cong_solve_dvd_int: |
| 31719 | 477 |
assumes a: "(a::int) \<noteq> 0" and b: "gcd a n dvd d" |
| 66380 | 478 |
shows "\<exists>x. [a * x = d] (mod n)" |
| 31719 | 479 |
proof - |
| 44872 | 480 |
from cong_solve_int [OF a] obtain x where "[a * x = gcd a n](mod n)" |
| 31719 | 481 |
by auto |
| 44872 | 482 |
then have "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)" |
| 66888 | 483 |
using cong_scalar_left by blast |
| 31719 | 484 |
also from b have "(d div gcd a n) * gcd a n = d" |
485 |
by (rule dvd_div_mult_self) |
|
486 |
also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)" |
|
487 |
by auto |
|
488 |
finally show ?thesis |
|
489 |
by auto |
|
490 |
qed |
|
491 |
||
| 66380 | 492 |
lemma cong_solve_coprime_nat: |
493 |
fixes a :: nat |
|
494 |
assumes "coprime a n" |
|
495 |
shows "\<exists>x. [a * x = 1] (mod n)" |
|
496 |
proof (cases "a = 0") |
|
497 |
case True |
|
| 66888 | 498 |
with assms show ?thesis |
499 |
by (simp add: cong_0_1_nat') |
|
| 66380 | 500 |
next |
501 |
case False |
|
| 66888 | 502 |
with assms show ?thesis |
503 |
by (metis cong_solve_nat) |
|
| 66380 | 504 |
qed |
| 31719 | 505 |
|
| 66380 | 506 |
lemma cong_solve_coprime_int: "coprime (a::int) n \<Longrightarrow> \<exists>x. [a * x = 1] (mod n)" |
| 44872 | 507 |
apply (cases "a = 0") |
| 66380 | 508 |
apply auto |
509 |
apply (cases "n \<ge> 0") |
|
510 |
apply auto |
|
| 55161 | 511 |
apply (metis cong_solve_int) |
512 |
done |
|
513 |
||
|
62349
7c23469b5118
cleansed junk-producing interpretations for gcd/lcm on nat altogether
haftmann
parents:
62348
diff
changeset
|
514 |
lemma coprime_iff_invertible_nat: |
| 66380 | 515 |
"m > 0 \<Longrightarrow> coprime a m = (\<exists>x. [a * x = Suc 0] (mod m))" |
|
62429
25271ff79171
Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents:
62353
diff
changeset
|
516 |
by (metis One_nat_def cong_gcd_eq_nat cong_solve_coprime_nat coprime_lmult gcd.commute gcd_Suc_0) |
| 66380 | 517 |
|
518 |
lemma coprime_iff_invertible_int: "m > 0 \<Longrightarrow> coprime a m \<longleftrightarrow> (\<exists>x. [a * x = 1] (mod m))" |
|
519 |
for m :: int |
|
| 55161 | 520 |
apply (auto intro: cong_solve_coprime_int) |
| 66888 | 521 |
using cong_gcd_eq_int coprime_mul_eq' apply fastforce |
| 44872 | 522 |
done |
| 31719 | 523 |
|
| 66380 | 524 |
lemma coprime_iff_invertible'_nat: |
525 |
"m > 0 \<Longrightarrow> coprime a m \<longleftrightarrow> (\<exists>x. 0 \<le> x \<and> x < m \<and> [a * x = Suc 0] (mod m))" |
|
| 55161 | 526 |
apply (subst coprime_iff_invertible_nat) |
| 66380 | 527 |
apply auto |
| 66888 | 528 |
apply (auto simp add: cong_def) |
| 55161 | 529 |
apply (metis mod_less_divisor mod_mult_right_eq) |
| 44872 | 530 |
done |
| 31719 | 531 |
|
| 66380 | 532 |
lemma coprime_iff_invertible'_int: |
533 |
"m > 0 \<Longrightarrow> coprime a m \<longleftrightarrow> (\<exists>x. 0 \<le> x \<and> x < m \<and> [a * x = 1] (mod m))" |
|
534 |
for m :: int |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
535 |
apply (subst coprime_iff_invertible_int) |
| 66888 | 536 |
apply (auto simp add: cong_def) |
| 55371 | 537 |
apply (metis mod_mult_right_eq pos_mod_conj) |
| 44872 | 538 |
done |
| 31719 | 539 |
|
| 66380 | 540 |
lemma cong_cong_lcm_nat: "[x = y] (mod a) \<Longrightarrow> [x = y] (mod b) \<Longrightarrow> [x = y] (mod lcm a b)" |
541 |
for x y :: nat |
|
| 44872 | 542 |
apply (cases "y \<le> x") |
| 66888 | 543 |
apply (simp add: cong_altdef_nat) |
544 |
apply (meson cong_altdef_nat cong_sym lcm_least_iff nat_le_linear) |
|
| 44872 | 545 |
done |
| 31719 | 546 |
|
| 66380 | 547 |
lemma cong_cong_lcm_int: "[x = y] (mod a) \<Longrightarrow> [x = y] (mod b) \<Longrightarrow> [x = y] (mod lcm a b)" |
548 |
for x y :: int |
|
549 |
by (auto simp add: cong_altdef_int lcm_least) |
|
| 31719 | 550 |
|
| 66888 | 551 |
lemma cong_cong_prod_coprime_nat: |
552 |
"[x = y] (mod (\<Prod>i\<in>A. m i))" if |
|
553 |
"(\<forall>i\<in>A. [(x::nat) = y] (mod m i))" |
|
554 |
and "(\<forall>i\<in>A. (\<forall>j\<in>A. i \<noteq> j \<longrightarrow> coprime (m i) (m j)))" |
|
555 |
using that apply (induct A rule: infinite_finite_induct) |
|
556 |
apply auto |
|
| 64272 | 557 |
apply (metis One_nat_def coprime_cong_mult_nat gcd.commute prod_coprime) |
| 44872 | 558 |
done |
| 31719 | 559 |
|
| 66888 | 560 |
lemma cong_cong_prod_coprime_int [rule_format]: |
561 |
"[x = y] (mod (\<Prod>i\<in>A. m i))" if |
|
562 |
"(\<forall>i\<in>A. [(x::int) = y] (mod m i))" |
|
563 |
"(\<forall>i\<in>A. (\<forall>j\<in>A. i \<noteq> j \<longrightarrow> coprime (m i) (m j)))" |
|
564 |
using that apply (induct A rule: infinite_finite_induct) |
|
| 31719 | 565 |
apply auto |
| 64272 | 566 |
apply (metis coprime_cong_mult_int gcd.commute prod_coprime) |
| 44872 | 567 |
done |
| 31719 | 568 |
|
| 44872 | 569 |
lemma binary_chinese_remainder_aux_nat: |
| 66380 | 570 |
fixes m1 m2 :: nat |
571 |
assumes a: "coprime m1 m2" |
|
572 |
shows "\<exists>b1 b2. [b1 = 1] (mod m1) \<and> [b1 = 0] (mod m2) \<and> [b2 = 0] (mod m1) \<and> [b2 = 1] (mod m2)" |
|
| 31719 | 573 |
proof - |
| 66380 | 574 |
from cong_solve_coprime_nat [OF a] obtain x1 where 1: "[m1 * x1 = 1] (mod m2)" |
| 31719 | 575 |
by auto |
| 44872 | 576 |
from a have b: "coprime m2 m1" |
| 62348 | 577 |
by (subst gcd.commute) |
| 66380 | 578 |
from cong_solve_coprime_nat [OF b] obtain x2 where 2: "[m2 * x2 = 1] (mod m1)" |
| 31719 | 579 |
by auto |
580 |
have "[m1 * x1 = 0] (mod m1)" |
|
| 66888 | 581 |
by (simp add: cong_mult_self_left) |
| 31719 | 582 |
moreover have "[m2 * x2 = 0] (mod m2)" |
| 66888 | 583 |
by (simp add: cong_mult_self_left) |
| 66380 | 584 |
ultimately show ?thesis |
585 |
using 1 2 by blast |
|
| 31719 | 586 |
qed |
587 |
||
| 44872 | 588 |
lemma binary_chinese_remainder_aux_int: |
| 66380 | 589 |
fixes m1 m2 :: int |
590 |
assumes a: "coprime m1 m2" |
|
591 |
shows "\<exists>b1 b2. [b1 = 1] (mod m1) \<and> [b1 = 0] (mod m2) \<and> [b2 = 0] (mod m1) \<and> [b2 = 1] (mod m2)" |
|
| 31719 | 592 |
proof - |
| 66380 | 593 |
from cong_solve_coprime_int [OF a] obtain x1 where 1: "[m1 * x1 = 1] (mod m2)" |
| 31719 | 594 |
by auto |
| 44872 | 595 |
from a have b: "coprime m2 m1" |
| 62348 | 596 |
by (subst gcd.commute) |
| 66380 | 597 |
from cong_solve_coprime_int [OF b] obtain x2 where 2: "[m2 * x2 = 1] (mod m1)" |
| 31719 | 598 |
by auto |
599 |
have "[m1 * x1 = 0] (mod m1)" |
|
| 66888 | 600 |
by (simp add: cong_mult_self_left) |
| 31719 | 601 |
moreover have "[m2 * x2 = 0] (mod m2)" |
| 66888 | 602 |
by (simp add: cong_mult_self_left) |
| 66380 | 603 |
ultimately show ?thesis |
604 |
using 1 2 by blast |
|
| 31719 | 605 |
qed |
606 |
||
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
607 |
lemma binary_chinese_remainder_nat: |
| 66380 | 608 |
fixes m1 m2 :: nat |
609 |
assumes a: "coprime m1 m2" |
|
610 |
shows "\<exists>x. [x = u1] (mod m1) \<and> [x = u2] (mod m2)" |
|
| 31719 | 611 |
proof - |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
612 |
from binary_chinese_remainder_aux_nat [OF a] obtain b1 b2 |
| 66380 | 613 |
where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)" |
614 |
and "[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)" |
|
| 31719 | 615 |
by blast |
616 |
let ?x = "u1 * b1 + u2 * b2" |
|
617 |
have "[?x = u1 * 1 + u2 * 0] (mod m1)" |
|
| 66888 | 618 |
using \<open>[b1 = 1] (mod m1)\<close> \<open>[b2 = 0] (mod m1)\<close> cong_add cong_scalar_left by blast |
| 44872 | 619 |
then have "[?x = u1] (mod m1)" by simp |
| 31719 | 620 |
have "[?x = u1 * 0 + u2 * 1] (mod m2)" |
| 66888 | 621 |
using \<open>[b1 = 0] (mod m2)\<close> \<open>[b2 = 1] (mod m2)\<close> cong_add cong_scalar_left by blast |
| 66380 | 622 |
then have "[?x = u2] (mod m2)" |
623 |
by simp |
|
624 |
with \<open>[?x = u1] (mod m1)\<close> show ?thesis |
|
625 |
by blast |
|
| 31719 | 626 |
qed |
627 |
||
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
628 |
lemma binary_chinese_remainder_int: |
| 66380 | 629 |
fixes m1 m2 :: int |
630 |
assumes a: "coprime m1 m2" |
|
631 |
shows "\<exists>x. [x = u1] (mod m1) \<and> [x = u2] (mod m2)" |
|
| 31719 | 632 |
proof - |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
633 |
from binary_chinese_remainder_aux_int [OF a] obtain b1 b2 |
| 66380 | 634 |
where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)" |
635 |
and "[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)" |
|
| 31719 | 636 |
by blast |
637 |
let ?x = "u1 * b1 + u2 * b2" |
|
638 |
have "[?x = u1 * 1 + u2 * 0] (mod m1)" |
|
| 66888 | 639 |
using \<open>[b1 = 1] (mod m1)\<close> \<open>[b2 = 0] (mod m1)\<close> cong_add cong_scalar_left by blast |
| 44872 | 640 |
then have "[?x = u1] (mod m1)" by simp |
| 31719 | 641 |
have "[?x = u1 * 0 + u2 * 1] (mod m2)" |
| 66888 | 642 |
using \<open>[b1 = 0] (mod m2)\<close> \<open>[b2 = 1] (mod m2)\<close> cong_add cong_scalar_left by blast |
| 44872 | 643 |
then have "[?x = u2] (mod m2)" by simp |
| 66380 | 644 |
with \<open>[?x = u1] (mod m1)\<close> show ?thesis |
645 |
by blast |
|
| 31719 | 646 |
qed |
647 |
||
| 66380 | 648 |
lemma cong_modulus_mult_nat: "[x = y] (mod m * n) \<Longrightarrow> [x = y] (mod m)" |
649 |
for x y :: nat |
|
| 44872 | 650 |
apply (cases "y \<le> x") |
| 66380 | 651 |
apply (simp add: cong_altdef_nat) |
652 |
apply (erule dvd_mult_left) |
|
| 66888 | 653 |
apply (rule cong_sym) |
654 |
apply (subst (asm) cong_sym_eq) |
|
| 44872 | 655 |
apply (simp add: cong_altdef_nat) |
| 31719 | 656 |
apply (erule dvd_mult_left) |
| 44872 | 657 |
done |
| 31719 | 658 |
|
| 66380 | 659 |
lemma cong_modulus_mult_int: "[x = y] (mod m * n) \<Longrightarrow> [x = y] (mod m)" |
660 |
for x y :: int |
|
| 44872 | 661 |
apply (simp add: cong_altdef_int) |
| 31719 | 662 |
apply (erule dvd_mult_left) |
| 44872 | 663 |
done |
| 31719 | 664 |
|
| 66380 | 665 |
lemma cong_less_modulus_unique_nat: "[x = y] (mod m) \<Longrightarrow> x < m \<Longrightarrow> y < m \<Longrightarrow> x = y" |
666 |
for x y :: nat |
|
| 66888 | 667 |
by (simp add: cong_def) |
| 31719 | 668 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
669 |
lemma binary_chinese_remainder_unique_nat: |
| 66380 | 670 |
fixes m1 m2 :: nat |
671 |
assumes a: "coprime m1 m2" |
|
| 44872 | 672 |
and nz: "m1 \<noteq> 0" "m2 \<noteq> 0" |
| 63901 | 673 |
shows "\<exists>!x. x < m1 * m2 \<and> [x = u1] (mod m1) \<and> [x = u2] (mod m2)" |
| 31719 | 674 |
proof - |
| 66380 | 675 |
from binary_chinese_remainder_nat [OF a] obtain y |
676 |
where "[y = u1] (mod m1)" and "[y = u2] (mod m2)" |
|
| 31719 | 677 |
by blast |
678 |
let ?x = "y mod (m1 * m2)" |
|
679 |
from nz have less: "?x < m1 * m2" |
|
| 44872 | 680 |
by auto |
| 66380 | 681 |
have 1: "[?x = u1] (mod m1)" |
| 66888 | 682 |
apply (rule cong_trans) |
| 66380 | 683 |
prefer 2 |
684 |
apply (rule \<open>[y = u1] (mod m1)\<close>) |
|
| 66888 | 685 |
apply (rule cong_modulus_mult_nat [of _ _ _ m2]) |
686 |
apply simp |
|
| 31719 | 687 |
done |
| 66380 | 688 |
have 2: "[?x = u2] (mod m2)" |
| 66888 | 689 |
apply (rule cong_trans) |
| 66380 | 690 |
prefer 2 |
691 |
apply (rule \<open>[y = u2] (mod m2)\<close>) |
|
|
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57418
diff
changeset
|
692 |
apply (subst mult.commute) |
| 66888 | 693 |
apply (rule cong_modulus_mult_nat [of _ _ _ m1]) |
694 |
apply simp |
|
| 31719 | 695 |
done |
| 66380 | 696 |
have "\<forall>z. z < m1 * m2 \<and> [z = u1] (mod m1) \<and> [z = u2] (mod m2) \<longrightarrow> z = ?x" |
| 44872 | 697 |
proof clarify |
| 31719 | 698 |
fix z |
699 |
assume "z < m1 * m2" |
|
700 |
assume "[z = u1] (mod m1)" and "[z = u2] (mod m2)" |
|
701 |
have "[?x = z] (mod m1)" |
|
| 66888 | 702 |
apply (rule cong_trans) |
| 66380 | 703 |
apply (rule \<open>[?x = u1] (mod m1)\<close>) |
| 66888 | 704 |
apply (rule cong_sym) |
| 60526 | 705 |
apply (rule \<open>[z = u1] (mod m1)\<close>) |
| 31719 | 706 |
done |
707 |
moreover have "[?x = z] (mod m2)" |
|
| 66888 | 708 |
apply (rule cong_trans) |
| 66380 | 709 |
apply (rule \<open>[?x = u2] (mod m2)\<close>) |
| 66888 | 710 |
apply (rule cong_sym) |
| 60526 | 711 |
apply (rule \<open>[z = u2] (mod m2)\<close>) |
| 31719 | 712 |
done |
713 |
ultimately have "[?x = z] (mod m1 * m2)" |
|
| 66888 | 714 |
using a by (auto intro: coprime_cong_mult_nat simp add: mod_mult_cong_left mod_mult_cong_right) |
| 60526 | 715 |
with \<open>z < m1 * m2\<close> \<open>?x < m1 * m2\<close> show "z = ?x" |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
716 |
apply (intro cong_less_modulus_unique_nat) |
| 66888 | 717 |
apply (auto, erule cong_sym) |
| 31719 | 718 |
done |
| 44872 | 719 |
qed |
| 66380 | 720 |
with less 1 2 show ?thesis by auto |
| 31719 | 721 |
qed |
722 |
||
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
723 |
lemma chinese_remainder_aux_nat: |
| 44872 | 724 |
fixes A :: "'a set" |
725 |
and m :: "'a \<Rightarrow> nat" |
|
726 |
assumes fin: "finite A" |
|
| 66380 | 727 |
and cop: "\<forall>i \<in> A. (\<forall>j \<in> A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))" |
728 |
shows "\<exists>b. (\<forall>i \<in> A. [b i = 1] (mod m i) \<and> [b i = 0] (mod (\<Prod>j \<in> A - {i}. m j)))"
|
|
| 31719 | 729 |
proof (rule finite_set_choice, rule fin, rule ballI) |
730 |
fix i |
|
| 66380 | 731 |
assume "i \<in> A" |
| 61954 | 732 |
with cop have "coprime (\<Prod>j \<in> A - {i}. m j) (m i)"
|
| 66380 | 733 |
by (intro prod_coprime) auto |
734 |
then have "\<exists>x. [(\<Prod>j \<in> A - {i}. m j) * x = 1] (mod m i)"
|
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
735 |
by (elim cong_solve_coprime_nat) |
| 61954 | 736 |
then obtain x where "[(\<Prod>j \<in> A - {i}. m j) * x = 1] (mod m i)"
|
| 31719 | 737 |
by auto |
| 66380 | 738 |
moreover have "[(\<Prod>j \<in> A - {i}. m j) * x = 0] (mod (\<Prod>j \<in> A - {i}. m j))"
|
| 66888 | 739 |
by (simp add: cong_0_iff) |
| 66380 | 740 |
ultimately show "\<exists>a. [a = 1] (mod m i) \<and> [a = 0] (mod prod m (A - {i}))"
|
| 31719 | 741 |
by blast |
742 |
qed |
|
743 |
||
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
744 |
lemma chinese_remainder_nat: |
| 44872 | 745 |
fixes A :: "'a set" |
746 |
and m :: "'a \<Rightarrow> nat" |
|
747 |
and u :: "'a \<Rightarrow> nat" |
|
748 |
assumes fin: "finite A" |
|
| 66380 | 749 |
and cop: "\<forall>i \<in> A. \<forall>j \<in> A. i \<noteq> j \<longrightarrow> coprime (m i) (m j)" |
750 |
shows "\<exists>x. \<forall>i \<in> A. [x = u i] (mod m i)" |
|
| 31719 | 751 |
proof - |
| 66380 | 752 |
from chinese_remainder_aux_nat [OF fin cop] |
753 |
obtain b where b: "\<forall>i \<in> A. [b i = 1] (mod m i) \<and> [b i = 0] (mod (\<Prod>j \<in> A - {i}. m j))"
|
|
| 31719 | 754 |
by blast |
| 61954 | 755 |
let ?x = "\<Sum>i\<in>A. (u i) * (b i)" |
| 66380 | 756 |
show ?thesis |
| 31719 | 757 |
proof (rule exI, clarify) |
758 |
fix i |
|
| 66380 | 759 |
assume a: "i \<in> A" |
| 44872 | 760 |
show "[?x = u i] (mod m i)" |
| 31719 | 761 |
proof - |
| 66380 | 762 |
from fin a have "?x = (\<Sum>j \<in> {i}. u j * b j) + (\<Sum>j \<in> A - {i}. u j * b j)"
|
763 |
by (subst sum.union_disjoint [symmetric]) (auto intro: sum.cong) |
|
| 61954 | 764 |
then have "[?x = u i * b i + (\<Sum>j \<in> A - {i}. u j * b j)] (mod m i)"
|
| 31719 | 765 |
by auto |
| 61954 | 766 |
also have "[u i * b i + (\<Sum>j \<in> A - {i}. u j * b j) =
|
767 |
u i * 1 + (\<Sum>j \<in> A - {i}. u j * 0)] (mod m i)"
|
|
| 66888 | 768 |
apply (rule cong_add) |
769 |
apply (rule cong_scalar_left) |
|
| 66380 | 770 |
using b a apply blast |
| 66888 | 771 |
apply (rule cong_sum) |
772 |
apply (rule cong_scalar_left) |
|
| 66380 | 773 |
using b apply auto |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
774 |
apply (rule cong_dvd_modulus_nat) |
| 66380 | 775 |
apply (drule (1) bspec) |
776 |
apply (erule conjE) |
|
777 |
apply assumption |
|
| 59010 | 778 |
apply rule |
| 31719 | 779 |
using fin a apply auto |
780 |
done |
|
781 |
finally show ?thesis |
|
782 |
by simp |
|
783 |
qed |
|
784 |
qed |
|
785 |
qed |
|
786 |
||
| 66888 | 787 |
lemma coprime_cong_prod_nat: |
788 |
"[x = y] (mod (\<Prod>i\<in>A. m i))" |
|
789 |
if "\<forall>i\<in>A. (\<forall>j\<in>A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))" |
|
790 |
and "\<forall>i\<in>A. [x = y] (mod m i)" for x y :: nat |
|
791 |
using that apply (induct A rule: infinite_finite_induct) |
|
| 31719 | 792 |
apply auto |
| 64272 | 793 |
apply (metis One_nat_def coprime_cong_mult_nat gcd.commute prod_coprime) |
| 44872 | 794 |
done |
| 31719 | 795 |
|
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
796 |
lemma chinese_remainder_unique_nat: |
| 44872 | 797 |
fixes A :: "'a set" |
798 |
and m :: "'a \<Rightarrow> nat" |
|
799 |
and u :: "'a \<Rightarrow> nat" |
|
800 |
assumes fin: "finite A" |
|
| 61954 | 801 |
and nz: "\<forall>i\<in>A. m i \<noteq> 0" |
| 66380 | 802 |
and cop: "\<forall>i\<in>A. \<forall>j\<in>A. i \<noteq> j \<longrightarrow> coprime (m i) (m j)" |
| 63901 | 803 |
shows "\<exists>!x. x < (\<Prod>i\<in>A. m i) \<and> (\<forall>i\<in>A. [x = u i] (mod m i))" |
| 31719 | 804 |
proof - |
| 44872 | 805 |
from chinese_remainder_nat [OF fin cop] |
| 66380 | 806 |
obtain y where one: "(\<forall>i\<in>A. [y = u i] (mod m i))" |
| 31719 | 807 |
by blast |
| 61954 | 808 |
let ?x = "y mod (\<Prod>i\<in>A. m i)" |
809 |
from fin nz have prodnz: "(\<Prod>i\<in>A. m i) \<noteq> 0" |
|
| 31719 | 810 |
by auto |
| 61954 | 811 |
then have less: "?x < (\<Prod>i\<in>A. m i)" |
| 31719 | 812 |
by auto |
| 66380 | 813 |
have cong: "\<forall>i\<in>A. [?x = u i] (mod m i)" |
| 31719 | 814 |
apply auto |
| 66888 | 815 |
apply (rule cong_trans) |
| 66380 | 816 |
prefer 2 |
| 31719 | 817 |
using one apply auto |
| 66888 | 818 |
apply (rule cong_dvd_modulus_nat [of _ _ "prod m A"]) |
819 |
apply simp |
|
820 |
using fin apply auto |
|
| 31719 | 821 |
done |
| 66380 | 822 |
have unique: "\<forall>z. z < (\<Prod>i\<in>A. m i) \<and> (\<forall>i\<in>A. [z = u i] (mod m i)) \<longrightarrow> z = ?x" |
823 |
proof clarify |
|
| 31719 | 824 |
fix z |
| 61954 | 825 |
assume zless: "z < (\<Prod>i\<in>A. m i)" |
| 66380 | 826 |
assume zcong: "(\<forall>i\<in>A. [z = u i] (mod m i))" |
827 |
have "\<forall>i\<in>A. [?x = z] (mod m i)" |
|
| 44872 | 828 |
apply clarify |
| 66888 | 829 |
apply (rule cong_trans) |
| 31719 | 830 |
using cong apply (erule bspec) |
| 66888 | 831 |
apply (rule cong_sym) |
| 31719 | 832 |
using zcong apply auto |
833 |
done |
|
| 61954 | 834 |
with fin cop have "[?x = z] (mod (\<Prod>i\<in>A. m i))" |
| 44872 | 835 |
apply (intro coprime_cong_prod_nat) |
| 66380 | 836 |
apply auto |
| 44872 | 837 |
done |
| 31719 | 838 |
with zless less show "z = ?x" |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
839 |
apply (intro cong_less_modulus_unique_nat) |
| 66380 | 840 |
apply auto |
| 66888 | 841 |
apply (erule cong_sym) |
| 31719 | 842 |
done |
| 44872 | 843 |
qed |
| 66380 | 844 |
from less cong unique show ?thesis |
845 |
by blast |
|
| 31719 | 846 |
qed |
847 |
||
848 |
end |