author | haftmann |
Sat, 01 Oct 2022 07:56:53 +0000 | |
changeset 76231 | 8a48e18f081e |
parent 76224 | 64e8d4afcf10 |
child 80084 | 173548e4d5d0 |
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 |
||
71546 | 40 |
definition cong :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" (\<open>(1[_ = _] '(' mod _'))\<close>) |
66888 | 41 |
where "cong b c a \<longleftrightarrow> b mod a = c mod a" |
42 |
||
71546 | 43 |
abbreviation notcong :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" (\<open>(1[_ \<noteq> _] '(' mod _'))\<close>) |
66888 | 44 |
where "notcong b c a \<equiv> \<not> cong b c a" |
45 |
||
46 |
lemma cong_refl [simp]: |
|
47 |
"[b = b] (mod a)" |
|
48 |
by (simp add: cong_def) |
|
49 |
||
50 |
lemma cong_sym: |
|
51 |
"[b = c] (mod a) \<Longrightarrow> [c = b] (mod a)" |
|
52 |
by (simp add: cong_def) |
|
53 |
||
54 |
lemma cong_sym_eq: |
|
55 |
"[b = c] (mod a) \<longleftrightarrow> [c = b] (mod a)" |
|
56 |
by (auto simp add: cong_def) |
|
57 |
||
58 |
lemma cong_trans [trans]: |
|
59 |
"[b = c] (mod a) \<Longrightarrow> [c = d] (mod a) \<Longrightarrow> [b = d] (mod a)" |
|
60 |
by (simp add: cong_def) |
|
61 |
||
67085 | 62 |
lemma cong_mult_self_right: |
63 |
"[b * a = 0] (mod a)" |
|
64 |
by (simp add: cong_def) |
|
65 |
||
66 |
lemma cong_mult_self_left: |
|
67 |
"[a * b = 0] (mod a)" |
|
68 |
by (simp add: cong_def) |
|
69 |
||
70 |
lemma cong_mod_left [simp]: |
|
71 |
"[b mod a = c] (mod a) \<longleftrightarrow> [b = c] (mod a)" |
|
72 |
by (simp add: cong_def) |
|
73 |
||
74 |
lemma cong_mod_right [simp]: |
|
75 |
"[b = c mod a] (mod a) \<longleftrightarrow> [b = c] (mod a)" |
|
76 |
by (simp add: cong_def) |
|
77 |
||
78 |
lemma cong_0 [simp, presburger]: |
|
79 |
"[b = c] (mod 0) \<longleftrightarrow> b = c" |
|
80 |
by (simp add: cong_def) |
|
81 |
||
82 |
lemma cong_1 [simp, presburger]: |
|
83 |
"[b = c] (mod 1)" |
|
84 |
by (simp add: cong_def) |
|
85 |
||
86 |
lemma cong_dvd_iff: |
|
87 |
"a dvd b \<longleftrightarrow> a dvd c" if "[b = c] (mod a)" |
|
88 |
using that by (auto simp: cong_def dvd_eq_mod_eq_0) |
|
89 |
||
90 |
lemma cong_0_iff: "[b = 0] (mod a) \<longleftrightarrow> a dvd b" |
|
91 |
by (simp add: cong_def dvd_eq_mod_eq_0) |
|
92 |
||
66888 | 93 |
lemma cong_add: |
94 |
"[b = c] (mod a) \<Longrightarrow> [d = e] (mod a) \<Longrightarrow> [b + d = c + e] (mod a)" |
|
95 |
by (auto simp add: cong_def intro: mod_add_cong) |
|
96 |
||
97 |
lemma cong_mult: |
|
98 |
"[b = c] (mod a) \<Longrightarrow> [d = e] (mod a) \<Longrightarrow> [b * d = c * e] (mod a)" |
|
99 |
by (auto simp add: cong_def intro: mod_mult_cong) |
|
100 |
||
67085 | 101 |
lemma cong_scalar_right: |
102 |
"[b = c] (mod a) \<Longrightarrow> [b * d = c * d] (mod a)" |
|
103 |
by (simp add: cong_mult) |
|
104 |
||
105 |
lemma cong_scalar_left: |
|
106 |
"[b = c] (mod a) \<Longrightarrow> [d * b = d * c] (mod a)" |
|
107 |
by (simp add: cong_mult) |
|
108 |
||
66888 | 109 |
lemma cong_pow: |
110 |
"[b = c] (mod a) \<Longrightarrow> [b ^ n = c ^ n] (mod a)" |
|
111 |
by (simp add: cong_def power_mod [symmetric, of b n a] power_mod [symmetric, of c n a]) |
|
112 |
||
113 |
lemma cong_sum: |
|
114 |
"[sum f A = sum g A] (mod a)" if "\<And>x. x \<in> A \<Longrightarrow> [f x = g x] (mod a)" |
|
115 |
using that by (induct A rule: infinite_finite_induct) (auto intro: cong_add) |
|
116 |
||
117 |
lemma cong_prod: |
|
118 |
"[prod f A = prod g A] (mod a)" if "(\<And>x. x \<in> A \<Longrightarrow> [f x = g x] (mod a))" |
|
119 |
using that by (induct A rule: infinite_finite_induct) (auto intro: cong_mult) |
|
120 |
||
121 |
lemma mod_mult_cong_right: |
|
122 |
"[c mod (a * b) = d] (mod a) \<longleftrightarrow> [c = d] (mod a)" |
|
123 |
by (simp add: cong_def mod_mod_cancel mod_add_left_eq) |
|
124 |
||
125 |
lemma mod_mult_cong_left: |
|
126 |
"[c mod (b * a) = d] (mod a) \<longleftrightarrow> [c = d] (mod a)" |
|
127 |
using mod_mult_cong_right [of c a b d] by (simp add: ac_simps) |
|
128 |
||
129 |
end |
|
130 |
||
131 |
context unique_euclidean_ring |
|
132 |
begin |
|
133 |
||
134 |
lemma cong_diff: |
|
135 |
"[b = c] (mod a) \<Longrightarrow> [d = e] (mod a) \<Longrightarrow> [b - d = c - e] (mod a)" |
|
136 |
by (auto simp add: cong_def intro: mod_diff_cong) |
|
137 |
||
138 |
lemma cong_diff_iff_cong_0: |
|
139 |
"[b - c = 0] (mod a) \<longleftrightarrow> [b = c] (mod a)" (is "?P \<longleftrightarrow> ?Q") |
|
140 |
proof |
|
141 |
assume ?P |
|
142 |
then have "[b - c + c = 0 + c] (mod a)" |
|
143 |
by (rule cong_add) simp |
|
144 |
then show ?Q |
|
145 |
by simp |
|
146 |
next |
|
147 |
assume ?Q |
|
148 |
with cong_diff [of b c a c c] show ?P |
|
149 |
by simp |
|
150 |
qed |
|
151 |
||
152 |
lemma cong_minus_minus_iff: |
|
153 |
"[- b = - c] (mod a) \<longleftrightarrow> [b = c] (mod a)" |
|
154 |
using cong_diff_iff_cong_0 [of b c a] cong_diff_iff_cong_0 [of "- b" "- c" a] |
|
155 |
by (simp add: cong_0_iff dvd_diff_commute) |
|
156 |
||
67115 | 157 |
lemma cong_modulus_minus_iff [iff]: |
67087 | 158 |
"[b = c] (mod - a) \<longleftrightarrow> [b = c] (mod a)" |
66888 | 159 |
using cong_diff_iff_cong_0 [of b c a] cong_diff_iff_cong_0 [of b c " -a"] |
160 |
by (simp add: cong_0_iff) |
|
31719 | 161 |
|
67087 | 162 |
lemma cong_iff_dvd_diff: |
163 |
"[a = b] (mod m) \<longleftrightarrow> m dvd (a - b)" |
|
164 |
by (simp add: cong_0_iff [symmetric] cong_diff_iff_cong_0) |
|
165 |
||
166 |
lemma cong_iff_lin: |
|
167 |
"[a = b] (mod m) \<longleftrightarrow> (\<exists>k. b = a + m * k)" (is "?P \<longleftrightarrow> ?Q") |
|
168 |
proof - |
|
169 |
have "?P \<longleftrightarrow> m dvd b - a" |
|
170 |
by (simp add: cong_iff_dvd_diff dvd_diff_commute) |
|
171 |
also have "\<dots> \<longleftrightarrow> ?Q" |
|
172 |
by (auto simp add: algebra_simps elim!: dvdE) |
|
173 |
finally show ?thesis |
|
174 |
by simp |
|
175 |
qed |
|
176 |
||
67115 | 177 |
lemma cong_add_lcancel: |
178 |
"[a + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)" |
|
179 |
by (simp add: cong_iff_lin algebra_simps) |
|
180 |
||
181 |
lemma cong_add_rcancel: |
|
182 |
"[x + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)" |
|
183 |
by (simp add: cong_iff_lin algebra_simps) |
|
184 |
||
185 |
lemma cong_add_lcancel_0: |
|
186 |
"[a + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)" |
|
187 |
using cong_add_lcancel [of a x 0 n] by simp |
|
188 |
||
189 |
lemma cong_add_rcancel_0: |
|
190 |
"[x + a = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)" |
|
191 |
using cong_add_rcancel [of x a 0 n] by simp |
|
192 |
||
193 |
lemma cong_dvd_modulus: |
|
194 |
"[x = y] (mod n)" if "[x = y] (mod m)" and "n dvd m" |
|
195 |
using that by (auto intro: dvd_trans simp add: cong_iff_dvd_diff) |
|
196 |
||
197 |
lemma cong_modulus_mult: |
|
198 |
"[x = y] (mod m)" if "[x = y] (mod m * n)" |
|
199 |
using that by (simp add: cong_iff_dvd_diff) (rule dvd_mult_left) |
|
200 |
||
31719 | 201 |
end |
202 |
||
67115 | 203 |
lemma cong_abs [simp]: |
204 |
"[x = y] (mod \<bar>m\<bar>) \<longleftrightarrow> [x = y] (mod m)" |
|
205 |
for x y :: "'a :: {unique_euclidean_ring, linordered_idom}" |
|
206 |
by (simp add: cong_iff_dvd_diff) |
|
207 |
||
208 |
lemma cong_square: |
|
209 |
"prime p \<Longrightarrow> 0 < a \<Longrightarrow> [a * a = 1] (mod p) \<Longrightarrow> [a = 1] (mod p) \<or> [a = - 1] (mod p)" |
|
210 |
for a p :: "'a :: {normalization_semidom, linordered_idom, unique_euclidean_ring}" |
|
211 |
by (auto simp add: cong_iff_dvd_diff square_diff_one_factored dest: prime_dvd_multD) |
|
212 |
||
213 |
lemma cong_mult_rcancel: |
|
214 |
"[a * k = b * k] (mod m) \<longleftrightarrow> [a = b] (mod m)" |
|
215 |
if "coprime k m" for a k m :: "'a::{unique_euclidean_ring, ring_gcd}" |
|
216 |
using that by (auto simp add: cong_iff_dvd_diff left_diff_distrib [symmetric] ac_simps coprime_dvd_mult_right_iff) |
|
217 |
||
218 |
lemma cong_mult_lcancel: |
|
219 |
"[k * a = k * b] (mod m) = [a = b] (mod m)" |
|
220 |
if "coprime k m" for a k m :: "'a::{unique_euclidean_ring, ring_gcd}" |
|
221 |
using that cong_mult_rcancel [of k m a b] by (simp add: ac_simps) |
|
222 |
||
223 |
lemma coprime_cong_mult: |
|
224 |
"[a = b] (mod m) \<Longrightarrow> [a = b] (mod n) \<Longrightarrow> coprime m n \<Longrightarrow> [a = b] (mod m * n)" |
|
225 |
for a b :: "'a :: {unique_euclidean_ring, semiring_gcd}" |
|
226 |
by (simp add: cong_iff_dvd_diff divides_mult) |
|
227 |
||
228 |
lemma cong_gcd_eq: |
|
229 |
"gcd a m = gcd b m" if "[a = b] (mod m)" |
|
230 |
for a b :: "'a :: {unique_euclidean_semiring, euclidean_semiring_gcd}" |
|
231 |
proof (cases "m = 0") |
|
232 |
case True |
|
233 |
with that show ?thesis |
|
234 |
by simp |
|
235 |
next |
|
236 |
case False |
|
237 |
moreover have "gcd (a mod m) m = gcd (b mod m) m" |
|
238 |
using that by (simp add: cong_def) |
|
239 |
ultimately show ?thesis |
|
240 |
by simp |
|
241 |
qed |
|
242 |
||
243 |
lemma cong_imp_coprime: |
|
244 |
"[a = b] (mod m) \<Longrightarrow> coprime a m \<Longrightarrow> coprime b m" |
|
245 |
for a b :: "'a :: {unique_euclidean_semiring, euclidean_semiring_gcd}" |
|
246 |
by (auto simp add: coprime_iff_gcd_eq_1 dest: cong_gcd_eq) |
|
247 |
||
248 |
lemma cong_cong_prod_coprime: |
|
249 |
"[x = y] (mod (\<Prod>i\<in>A. m i))" if |
|
250 |
"(\<forall>i\<in>A. [x = y] (mod m i))" |
|
251 |
"(\<forall>i\<in>A. (\<forall>j\<in>A. i \<noteq> j \<longrightarrow> coprime (m i) (m j)))" |
|
252 |
for x y :: "'a :: {unique_euclidean_ring, semiring_gcd}" |
|
253 |
using that by (induct A rule: infinite_finite_induct) |
|
254 |
(auto intro!: coprime_cong_mult prod_coprime_right) |
|
255 |
||
66380 | 256 |
|
69597 | 257 |
subsection \<open>Congruences on \<^typ>\<open>nat\<close> and \<^typ>\<open>int\<close>\<close> |
31719 | 258 |
|
66888 | 259 |
lemma cong_int_iff: |
260 |
"[int m = int q] (mod int n) \<longleftrightarrow> [m = q] (mod n)" |
|
261 |
by (simp add: cong_def of_nat_mod [symmetric]) |
|
262 |
||
263 |
lemma cong_Suc_0 [simp, presburger]: |
|
264 |
"[m = n] (mod Suc 0)" |
|
265 |
using cong_1 [of m n] by simp |
|
31719 | 266 |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
267 |
lemma cong_diff_nat: |
66888 | 268 |
"[a - c = b - d] (mod m)" if "[a = b] (mod m)" "[c = d] (mod m)" |
66954 | 269 |
and "a \<ge> c" "b \<ge> d" for a b c d m :: nat |
270 |
proof - |
|
271 |
have "[c + (a - c) = d + (b - d)] (mod m)" |
|
272 |
using that by simp |
|
273 |
with \<open>[c = d] (mod m)\<close> have "[c + (a - c) = c + (b - d)] (mod m)" |
|
274 |
using mod_add_cong by (auto simp add: cong_def) fastforce |
|
275 |
then show ?thesis |
|
276 |
by (simp add: cong_def nat_mod_eq_iff) |
|
277 |
qed |
|
31719 | 278 |
|
66888 | 279 |
lemma cong_diff_iff_cong_0_nat: |
66954 | 280 |
"[a - b = 0] (mod m) \<longleftrightarrow> [a = b] (mod m)" if "a \<ge> b" for a b :: nat |
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
71546
diff
changeset
|
281 |
using that by (simp add: cong_0_iff) (simp add: cong_def mod_eq_dvd_iff_nat) |
31719 | 282 |
|
66954 | 283 |
lemma cong_diff_iff_cong_0_nat': |
284 |
"[nat \<bar>int a - int b\<bar> = 0] (mod m) \<longleftrightarrow> [a = b] (mod m)" |
|
285 |
proof (cases "b \<le> a") |
|
286 |
case True |
|
287 |
then show ?thesis |
|
288 |
by (simp add: nat_diff_distrib' cong_diff_iff_cong_0_nat [of b a m]) |
|
289 |
next |
|
290 |
case False |
|
291 |
then have "a \<le> b" |
|
292 |
by simp |
|
293 |
then show ?thesis |
|
294 |
by (simp add: nat_diff_distrib' cong_diff_iff_cong_0_nat [of a b m]) |
|
295 |
(auto simp add: cong_def) |
|
296 |
qed |
|
297 |
||
298 |
lemma cong_altdef_nat: |
|
299 |
"a \<ge> b \<Longrightarrow> [a = b] (mod m) \<longleftrightarrow> m dvd (a - b)" |
|
66380 | 300 |
for a b :: nat |
66888 | 301 |
by (simp add: cong_0_iff [symmetric] cong_diff_iff_cong_0_nat) |
31719 | 302 |
|
66954 | 303 |
lemma cong_altdef_nat': |
304 |
"[a = b] (mod m) \<longleftrightarrow> m dvd nat \<bar>int a - int b\<bar>" |
|
67085 | 305 |
using cong_diff_iff_cong_0_nat' [of a b m] |
306 |
by (simp only: cong_0_iff [symmetric]) |
|
66954 | 307 |
|
308 |
lemma cong_mult_rcancel_nat: |
|
309 |
"[a * k = b * k] (mod m) \<longleftrightarrow> [a = b] (mod m)" |
|
310 |
if "coprime k m" for a k m :: nat |
|
311 |
proof - |
|
312 |
have "[a * k = b * k] (mod m) \<longleftrightarrow> m dvd nat \<bar>int (a * k) - int (b * k)\<bar>" |
|
313 |
by (simp add: cong_altdef_nat') |
|
314 |
also have "\<dots> \<longleftrightarrow> m dvd nat \<bar>(int a - int b) * int k\<bar>" |
|
315 |
by (simp add: algebra_simps) |
|
316 |
also have "\<dots> \<longleftrightarrow> m dvd nat \<bar>int a - int b\<bar> * k" |
|
317 |
by (simp add: abs_mult nat_times_as_int) |
|
318 |
also have "\<dots> \<longleftrightarrow> m dvd nat \<bar>int a - int b\<bar>" |
|
67051 | 319 |
by (rule coprime_dvd_mult_left_iff) (use \<open>coprime k m\<close> in \<open>simp add: ac_simps\<close>) |
66954 | 320 |
also have "\<dots> \<longleftrightarrow> [a = b] (mod m)" |
321 |
by (simp add: cong_altdef_nat') |
|
322 |
finally show ?thesis . |
|
323 |
qed |
|
31719 | 324 |
|
66954 | 325 |
lemma cong_mult_lcancel_nat: |
326 |
"[k * a = k * b] (mod m) = [a = b] (mod m)" |
|
327 |
if "coprime k m" for a k m :: nat |
|
328 |
using that by (simp add: cong_mult_rcancel_nat ac_simps) |
|
31719 | 329 |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
330 |
lemma coprime_cong_mult_nat: |
66380 | 331 |
"[a = b] (mod m) \<Longrightarrow> [a = b] (mod n) \<Longrightarrow> coprime m n \<Longrightarrow> [a = b] (mod m * n)" |
332 |
for a b :: nat |
|
66954 | 333 |
by (simp add: cong_altdef_nat' divides_mult) |
31719 | 334 |
|
66380 | 335 |
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" |
336 |
for a b :: nat |
|
66888 | 337 |
by (auto simp add: cong_def) |
31719 | 338 |
|
66380 | 339 |
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" |
340 |
for a b :: int |
|
66888 | 341 |
by (auto simp add: cong_def) |
31719 | 342 |
|
66380 | 343 |
lemma cong_less_unique_nat: "0 < m \<Longrightarrow> (\<exists>!b. 0 \<le> b \<and> b < m \<and> [a = b] (mod m))" |
344 |
for a m :: nat |
|
76231 | 345 |
by (auto simp: cong_def) (metis mod_mod_trivial mod_less_divisor) |
31719 | 346 |
|
66380 | 347 |
lemma cong_less_unique_int: "0 < m \<Longrightarrow> (\<exists>!b. 0 \<le> b \<and> b < m \<and> [a = b] (mod m))" |
348 |
for a m :: int |
|
76231 | 349 |
by (auto simp add: cong_def) (metis mod_mod_trivial pos_mod_bound pos_mod_sign) |
31719 | 350 |
|
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
71546
diff
changeset
|
351 |
lemma cong_iff_lin_nat: "[a = b] (mod m) \<longleftrightarrow> (\<exists>k1 k2. b + k1 * m = a + k2 * m)" |
66380 | 352 |
for a b :: nat |
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
71546
diff
changeset
|
353 |
apply (auto simp add: cong_def nat_mod_eq_iff) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
71546
diff
changeset
|
354 |
apply (metis mult.commute) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
71546
diff
changeset
|
355 |
apply (metis mult.commute) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
71546
diff
changeset
|
356 |
done |
31719 | 357 |
|
66380 | 358 |
lemma cong_cong_mod_nat: "[a = b] (mod m) \<longleftrightarrow> [a mod m = b mod m] (mod m)" |
359 |
for a b :: nat |
|
66888 | 360 |
by simp |
31719 | 361 |
|
66380 | 362 |
lemma cong_cong_mod_int: "[a = b] (mod m) \<longleftrightarrow> [a mod m = b mod m] (mod m)" |
363 |
for a b :: int |
|
66888 | 364 |
by simp |
31719 | 365 |
|
66380 | 366 |
lemma cong_add_lcancel_nat: "[a + x = a + y] (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_nat: "[x + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)" |
371 |
for a x y :: nat |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
372 |
by (simp add: cong_iff_lin_nat) |
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 |
|
67085 | 376 |
using cong_add_lcancel_nat [of a x 0 n] by simp |
31719 | 377 |
|
66380 | 378 |
lemma cong_add_rcancel_0_nat: "[x + a = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)" |
379 |
for a x :: nat |
|
67085 | 380 |
using cong_add_rcancel_nat [of x a 0 n] by simp |
66380 | 381 |
|
382 |
lemma cong_dvd_modulus_nat: "[x = y] (mod m) \<Longrightarrow> n dvd m \<Longrightarrow> [x = y] (mod n)" |
|
383 |
for x y :: nat |
|
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
71546
diff
changeset
|
384 |
by (auto simp add: cong_altdef_nat') |
31719 | 385 |
|
66380 | 386 |
lemma cong_to_1_nat: |
387 |
fixes a :: nat |
|
388 |
assumes "[a = 1] (mod n)" |
|
389 |
shows "n dvd (a - 1)" |
|
390 |
proof (cases "a = 0") |
|
391 |
case True |
|
392 |
then show ?thesis by force |
|
393 |
next |
|
394 |
case False |
|
395 |
with assms show ?thesis by (metis cong_altdef_nat leI less_one) |
|
396 |
qed |
|
397 |
||
398 |
lemma cong_0_1_nat': "[0 = Suc 0] (mod n) \<longleftrightarrow> n = Suc 0" |
|
66888 | 399 |
by (auto simp: cong_def) |
66380 | 400 |
|
401 |
lemma cong_0_1_nat: "[0 = 1] (mod n) \<longleftrightarrow> n = 1" |
|
402 |
for n :: nat |
|
66888 | 403 |
by (auto simp: cong_def) |
66380 | 404 |
|
405 |
lemma cong_0_1_int: "[0 = 1] (mod n) \<longleftrightarrow> n = 1 \<or> n = - 1" |
|
406 |
for n :: int |
|
66888 | 407 |
by (auto simp: cong_def zmult_eq_1_iff) |
66380 | 408 |
|
409 |
lemma cong_to_1'_nat: "[a = 1] (mod n) \<longleftrightarrow> a = 0 \<and> n = 1 \<or> (\<exists>m. a = 1 + m * n)" |
|
410 |
for a :: nat |
|
411 |
by (metis add.right_neutral cong_0_1_nat cong_iff_lin_nat cong_to_1_nat |
|
412 |
dvd_div_mult_self leI le_add_diff_inverse less_one mult_eq_if) |
|
413 |
||
414 |
lemma cong_le_nat: "y \<le> x \<Longrightarrow> [x = y] (mod n) \<longleftrightarrow> (\<exists>q. x = q * n + y)" |
|
415 |
for x y :: nat |
|
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
71546
diff
changeset
|
416 |
by (auto simp add: cong_altdef_nat le_imp_diff_is_add) |
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
417 |
|
66380 | 418 |
lemma cong_solve_nat: |
419 |
fixes a :: nat |
|
420 |
shows "\<exists>x. [a * x = gcd a n] (mod n)" |
|
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
421 |
proof (cases "a = 0 \<or> n = 0") |
66380 | 422 |
case True |
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
423 |
then show ?thesis |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
424 |
by (force simp add: cong_0_iff cong_sym) |
66380 | 425 |
next |
426 |
case False |
|
427 |
then show ?thesis |
|
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
428 |
using bezout_nat [of a n] |
66888 | 429 |
by auto (metis cong_add_rcancel_0_nat cong_mult_self_left) |
66380 | 430 |
qed |
431 |
||
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
432 |
lemma cong_solve_int: |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
433 |
fixes a :: int |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
434 |
shows "\<exists>x. [a * x = gcd a n] (mod n)" |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
435 |
by (metis bezout_int cong_iff_lin mult.commute) |
31719 | 436 |
|
44872 | 437 |
lemma cong_solve_dvd_nat: |
66380 | 438 |
fixes a :: nat |
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
439 |
assumes "gcd a n dvd d" |
66380 | 440 |
shows "\<exists>x. [a * x = d] (mod n)" |
31719 | 441 |
proof - |
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
442 |
from cong_solve_nat [of a] obtain x where "[a * x = gcd a n](mod n)" |
31719 | 443 |
by auto |
44872 | 444 |
then have "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)" |
66888 | 445 |
using cong_scalar_left by blast |
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
446 |
also from assms have "(d div gcd a n) * gcd a n = d" |
31719 | 447 |
by (rule dvd_div_mult_self) |
448 |
also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)" |
|
449 |
by auto |
|
450 |
finally show ?thesis |
|
451 |
by auto |
|
452 |
qed |
|
453 |
||
44872 | 454 |
lemma cong_solve_dvd_int: |
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
455 |
fixes a::int |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
456 |
assumes b: "gcd a n dvd d" |
66380 | 457 |
shows "\<exists>x. [a * x = d] (mod n)" |
31719 | 458 |
proof - |
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
459 |
from cong_solve_int [of a] obtain x where "[a * x = gcd a n](mod n)" |
31719 | 460 |
by auto |
44872 | 461 |
then have "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)" |
66888 | 462 |
using cong_scalar_left by blast |
31719 | 463 |
also from b have "(d div gcd a n) * gcd a n = d" |
464 |
by (rule dvd_div_mult_self) |
|
465 |
also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)" |
|
466 |
by auto |
|
467 |
finally show ?thesis |
|
468 |
by auto |
|
469 |
qed |
|
470 |
||
66380 | 471 |
lemma cong_solve_coprime_nat: |
67051 | 472 |
"\<exists>x. [a * x = Suc 0] (mod n)" if "coprime a n" |
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
473 |
using that cong_solve_nat [of a n] by auto |
31719 | 474 |
|
67051 | 475 |
lemma cong_solve_coprime_int: |
476 |
"\<exists>x. [a * x = 1] (mod n)" if "coprime a n" for a n x :: int |
|
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
477 |
using that cong_solve_int [of a n] by (auto simp add: zabs_def split: if_splits) |
55161 | 478 |
|
62349
7c23469b5118
cleansed junk-producing interpretations for gcd/lcm on nat altogether
haftmann
parents:
62348
diff
changeset
|
479 |
lemma coprime_iff_invertible_nat: |
67085 | 480 |
"coprime a m \<longleftrightarrow> (\<exists>x. [a * x = Suc 0] (mod m))" (is "?P \<longleftrightarrow> ?Q") |
481 |
proof |
|
482 |
assume ?P then show ?Q |
|
483 |
by (auto dest!: cong_solve_coprime_nat) |
|
484 |
next |
|
485 |
assume ?Q |
|
486 |
then obtain b where "[a * b = Suc 0] (mod m)" |
|
487 |
by blast |
|
488 |
with coprime_mod_left_iff [of m "a * b"] show ?P |
|
489 |
by (cases "m = 0 \<or> m = 1") |
|
490 |
(unfold cong_def, auto simp add: cong_def) |
|
491 |
qed |
|
66380 | 492 |
|
67051 | 493 |
lemma coprime_iff_invertible_int: |
67085 | 494 |
"coprime a m \<longleftrightarrow> (\<exists>x. [a * x = 1] (mod m))" (is "?P \<longleftrightarrow> ?Q") for m :: int |
495 |
proof |
|
496 |
assume ?P then show ?Q |
|
497 |
by (auto dest: cong_solve_coprime_int) |
|
498 |
next |
|
499 |
assume ?Q |
|
500 |
then obtain b where "[a * b = 1] (mod m)" |
|
501 |
by blast |
|
502 |
with coprime_mod_left_iff [of m "a * b"] show ?P |
|
503 |
by (cases "m = 0 \<or> m = 1") |
|
504 |
(unfold cong_def, auto simp add: zmult_eq_1_iff) |
|
505 |
qed |
|
31719 | 506 |
|
66380 | 507 |
lemma coprime_iff_invertible'_nat: |
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
508 |
assumes "m > 0" |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
509 |
shows "coprime a m \<longleftrightarrow> (\<exists>x. 0 \<le> x \<and> x < m \<and> [a * x = Suc 0] (mod m))" |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
510 |
proof - |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
511 |
have "\<And>b. \<lbrakk>0 < m; [a * b = Suc 0] (mod m)\<rbrakk> \<Longrightarrow> \<exists>b'<m. [a * b' = Suc 0] (mod m)" |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
512 |
by (metis cong_def mod_less_divisor [OF assms] mod_mult_right_eq) |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
513 |
then show ?thesis |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
514 |
using assms coprime_iff_invertible_nat by auto |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
515 |
qed |
31719 | 516 |
|
66380 | 517 |
lemma coprime_iff_invertible'_int: |
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
518 |
fixes m :: int |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
519 |
assumes "m > 0" |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
520 |
shows "coprime a m \<longleftrightarrow> (\<exists>x. 0 \<le> x \<and> x < m \<and> [a * x = 1] (mod m))" |
76231 | 521 |
using assms by (simp add: coprime_iff_invertible_int) |
522 |
(metis assms cong_mod_left mod_mult_right_eq pos_mod_bound pos_mod_sign) |
|
31719 | 523 |
|
66380 | 524 |
lemma cong_cong_lcm_nat: "[x = y] (mod a) \<Longrightarrow> [x = y] (mod b) \<Longrightarrow> [x = y] (mod lcm a b)" |
525 |
for x y :: nat |
|
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
526 |
by (meson cong_altdef_nat' lcm_least) |
31719 | 527 |
|
66380 | 528 |
lemma cong_cong_lcm_int: "[x = y] (mod a) \<Longrightarrow> [x = y] (mod b) \<Longrightarrow> [x = y] (mod lcm a b)" |
529 |
for x y :: int |
|
67115 | 530 |
by (auto simp add: cong_iff_dvd_diff lcm_least) |
31719 | 531 |
|
66888 | 532 |
lemma cong_cong_prod_coprime_nat: |
533 |
"[x = y] (mod (\<Prod>i\<in>A. m i))" if |
|
67115 | 534 |
"(\<forall>i\<in>A. [x = y] (mod m i))" |
66888 | 535 |
"(\<forall>i\<in>A. (\<forall>j\<in>A. i \<noteq> j \<longrightarrow> coprime (m i) (m j)))" |
67115 | 536 |
for x y :: nat |
537 |
using that by (induct A rule: infinite_finite_induct) |
|
538 |
(auto intro!: coprime_cong_mult_nat prod_coprime_right) |
|
31719 | 539 |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
540 |
lemma binary_chinese_remainder_nat: |
66380 | 541 |
fixes m1 m2 :: nat |
542 |
assumes a: "coprime m1 m2" |
|
543 |
shows "\<exists>x. [x = u1] (mod m1) \<and> [x = u2] (mod m2)" |
|
31719 | 544 |
proof - |
67086 | 545 |
have "\<exists>b1 b2. [b1 = 1] (mod m1) \<and> [b1 = 0] (mod m2) \<and> [b2 = 0] (mod m1) \<and> [b2 = 1] (mod m2)" |
546 |
proof - |
|
547 |
from cong_solve_coprime_nat [OF a] obtain x1 where 1: "[m1 * x1 = 1] (mod m2)" |
|
548 |
by auto |
|
549 |
from a have b: "coprime m2 m1" |
|
550 |
by (simp add: ac_simps) |
|
551 |
from cong_solve_coprime_nat [OF b] obtain x2 where 2: "[m2 * x2 = 1] (mod m1)" |
|
552 |
by auto |
|
553 |
have "[m1 * x1 = 0] (mod m1)" |
|
554 |
by (simp add: cong_mult_self_left) |
|
555 |
moreover have "[m2 * x2 = 0] (mod m2)" |
|
556 |
by (simp add: cong_mult_self_left) |
|
557 |
ultimately show ?thesis |
|
558 |
using 1 2 by blast |
|
559 |
qed |
|
560 |
then obtain b1 b2 |
|
66380 | 561 |
where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)" |
562 |
and "[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)" |
|
31719 | 563 |
by blast |
564 |
let ?x = "u1 * b1 + u2 * b2" |
|
565 |
have "[?x = u1 * 1 + u2 * 0] (mod m1)" |
|
66888 | 566 |
using \<open>[b1 = 1] (mod m1)\<close> \<open>[b2 = 0] (mod m1)\<close> cong_add cong_scalar_left by blast |
44872 | 567 |
then have "[?x = u1] (mod m1)" by simp |
31719 | 568 |
have "[?x = u1 * 0 + u2 * 1] (mod m2)" |
66888 | 569 |
using \<open>[b1 = 0] (mod m2)\<close> \<open>[b2 = 1] (mod m2)\<close> cong_add cong_scalar_left by blast |
66380 | 570 |
then have "[?x = u2] (mod m2)" |
571 |
by simp |
|
572 |
with \<open>[?x = u1] (mod m1)\<close> show ?thesis |
|
573 |
by blast |
|
31719 | 574 |
qed |
575 |
||
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
576 |
lemma binary_chinese_remainder_int: |
66380 | 577 |
fixes m1 m2 :: int |
578 |
assumes a: "coprime m1 m2" |
|
579 |
shows "\<exists>x. [x = u1] (mod m1) \<and> [x = u2] (mod m2)" |
|
31719 | 580 |
proof - |
67086 | 581 |
have "\<exists>b1 b2. [b1 = 1] (mod m1) \<and> [b1 = 0] (mod m2) \<and> [b2 = 0] (mod m1) \<and> [b2 = 1] (mod m2)" |
582 |
proof - |
|
583 |
from cong_solve_coprime_int [OF a] obtain x1 where 1: "[m1 * x1 = 1] (mod m2)" |
|
584 |
by auto |
|
585 |
from a have b: "coprime m2 m1" |
|
586 |
by (simp add: ac_simps) |
|
587 |
from cong_solve_coprime_int [OF b] obtain x2 where 2: "[m2 * x2 = 1] (mod m1)" |
|
588 |
by auto |
|
589 |
have "[m1 * x1 = 0] (mod m1)" |
|
590 |
by (simp add: cong_mult_self_left) |
|
591 |
moreover have "[m2 * x2 = 0] (mod m2)" |
|
592 |
by (simp add: cong_mult_self_left) |
|
593 |
ultimately show ?thesis |
|
594 |
using 1 2 by blast |
|
595 |
qed |
|
596 |
then obtain b1 b2 |
|
66380 | 597 |
where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)" |
598 |
and "[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)" |
|
31719 | 599 |
by blast |
600 |
let ?x = "u1 * b1 + u2 * b2" |
|
601 |
have "[?x = u1 * 1 + u2 * 0] (mod m1)" |
|
66888 | 602 |
using \<open>[b1 = 1] (mod m1)\<close> \<open>[b2 = 0] (mod m1)\<close> cong_add cong_scalar_left by blast |
44872 | 603 |
then have "[?x = u1] (mod m1)" by simp |
31719 | 604 |
have "[?x = u1 * 0 + u2 * 1] (mod m2)" |
66888 | 605 |
using \<open>[b1 = 0] (mod m2)\<close> \<open>[b2 = 1] (mod m2)\<close> cong_add cong_scalar_left by blast |
44872 | 606 |
then have "[?x = u2] (mod m2)" by simp |
66380 | 607 |
with \<open>[?x = u1] (mod m1)\<close> show ?thesis |
608 |
by blast |
|
31719 | 609 |
qed |
610 |
||
66380 | 611 |
lemma cong_modulus_mult_nat: "[x = y] (mod m * n) \<Longrightarrow> [x = y] (mod m)" |
612 |
for x y :: nat |
|
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
613 |
by (metis cong_def mod_mult_cong_right) |
31719 | 614 |
|
66380 | 615 |
lemma cong_less_modulus_unique_nat: "[x = y] (mod m) \<Longrightarrow> x < m \<Longrightarrow> y < m \<Longrightarrow> x = y" |
616 |
for x y :: nat |
|
66888 | 617 |
by (simp add: cong_def) |
31719 | 618 |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
619 |
lemma binary_chinese_remainder_unique_nat: |
66380 | 620 |
fixes m1 m2 :: nat |
621 |
assumes a: "coprime m1 m2" |
|
44872 | 622 |
and nz: "m1 \<noteq> 0" "m2 \<noteq> 0" |
63901 | 623 |
shows "\<exists>!x. x < m1 * m2 \<and> [x = u1] (mod m1) \<and> [x = u2] (mod m2)" |
31719 | 624 |
proof - |
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
625 |
obtain y where y1: "[y = u1] (mod m1)" and y2: "[y = u2] (mod m2)" |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
626 |
using binary_chinese_remainder_nat [OF a] by blast |
31719 | 627 |
let ?x = "y mod (m1 * m2)" |
628 |
from nz have less: "?x < m1 * m2" |
|
44872 | 629 |
by auto |
66380 | 630 |
have 1: "[?x = u1] (mod m1)" |
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
631 |
using y1 mod_mult_cong_right by blast |
66380 | 632 |
have 2: "[?x = u2] (mod m2)" |
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
633 |
using y2 mod_mult_cong_left by blast |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
634 |
have "z = ?x" if "z < m1 * m2" "[z = u1] (mod m1)" "[z = u2] (mod m2)" for z |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
635 |
proof - |
31719 | 636 |
have "[?x = z] (mod m1)" |
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
637 |
by (metis "1" cong_def that(2)) |
31719 | 638 |
moreover have "[?x = z] (mod m2)" |
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
639 |
by (metis "2" cong_def that(3)) |
31719 | 640 |
ultimately have "[?x = z] (mod m1 * m2)" |
66888 | 641 |
using a by (auto intro: coprime_cong_mult_nat simp add: mod_mult_cong_left mod_mult_cong_right) |
60526 | 642 |
with \<open>z < m1 * m2\<close> \<open>?x < m1 * m2\<close> show "z = ?x" |
67085 | 643 |
by (auto simp add: cong_def) |
44872 | 644 |
qed |
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
645 |
with less 1 2 show ?thesis |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
646 |
by blast |
31719 | 647 |
qed |
648 |
||
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
649 |
lemma chinese_remainder_nat: |
44872 | 650 |
fixes A :: "'a set" |
651 |
and m :: "'a \<Rightarrow> nat" |
|
652 |
and u :: "'a \<Rightarrow> nat" |
|
653 |
assumes fin: "finite A" |
|
66380 | 654 |
and cop: "\<forall>i \<in> A. \<forall>j \<in> A. i \<noteq> j \<longrightarrow> coprime (m i) (m j)" |
655 |
shows "\<exists>x. \<forall>i \<in> A. [x = u i] (mod m i)" |
|
31719 | 656 |
proof - |
67086 | 657 |
have "\<exists>b. (\<forall>i \<in> A. [b i = 1] (mod m i) \<and> [b i = 0] (mod (\<Prod>j \<in> A - {i}. m j)))" |
658 |
proof (rule finite_set_choice, rule fin, rule ballI) |
|
659 |
fix i |
|
660 |
assume "i \<in> A" |
|
661 |
with cop have "coprime (\<Prod>j \<in> A - {i}. m j) (m i)" |
|
662 |
by (intro prod_coprime_left) auto |
|
663 |
then have "\<exists>x. [(\<Prod>j \<in> A - {i}. m j) * x = Suc 0] (mod m i)" |
|
664 |
by (elim cong_solve_coprime_nat) |
|
665 |
then obtain x where "[(\<Prod>j \<in> A - {i}. m j) * x = 1] (mod m i)" |
|
666 |
by auto |
|
667 |
moreover have "[(\<Prod>j \<in> A - {i}. m j) * x = 0] (mod (\<Prod>j \<in> A - {i}. m j))" |
|
668 |
by (simp add: cong_0_iff) |
|
669 |
ultimately show "\<exists>a. [a = 1] (mod m i) \<and> [a = 0] (mod prod m (A - {i}))" |
|
670 |
by blast |
|
671 |
qed |
|
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
672 |
then obtain b where b: "\<And>i. i \<in> A \<Longrightarrow> [b i = 1] (mod m i) \<and> [b i = 0] (mod (\<Prod>j \<in> A - {i}. m j))" |
31719 | 673 |
by blast |
61954 | 674 |
let ?x = "\<Sum>i\<in>A. (u i) * (b i)" |
66380 | 675 |
show ?thesis |
31719 | 676 |
proof (rule exI, clarify) |
677 |
fix i |
|
66380 | 678 |
assume a: "i \<in> A" |
44872 | 679 |
show "[?x = u i] (mod m i)" |
31719 | 680 |
proof - |
66380 | 681 |
from fin a have "?x = (\<Sum>j \<in> {i}. u j * b j) + (\<Sum>j \<in> A - {i}. u j * b j)" |
682 |
by (subst sum.union_disjoint [symmetric]) (auto intro: sum.cong) |
|
61954 | 683 |
then have "[?x = u i * b i + (\<Sum>j \<in> A - {i}. u j * b j)] (mod m i)" |
31719 | 684 |
by auto |
61954 | 685 |
also have "[u i * b i + (\<Sum>j \<in> A - {i}. u j * b j) = |
686 |
u i * 1 + (\<Sum>j \<in> A - {i}. u j * 0)] (mod m i)" |
|
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
687 |
proof (intro cong_add cong_scalar_left cong_sum) |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
688 |
show "[b i = 1] (mod m i)" |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
689 |
using a b by blast |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
690 |
show "[b x = 0] (mod m i)" if "x \<in> A - {i}" for x |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
691 |
proof - |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
692 |
have "x \<in> A" "x \<noteq> i" |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
693 |
using that by auto |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
694 |
then show ?thesis |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
695 |
using a b [OF \<open>x \<in> A\<close>] cong_dvd_modulus_nat fin by blast |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
696 |
qed |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
697 |
qed |
31719 | 698 |
finally show ?thesis |
699 |
by simp |
|
700 |
qed |
|
701 |
qed |
|
702 |
qed |
|
703 |
||
68707
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
704 |
lemma coprime_cong_prod_nat: "[x = y] (mod (\<Prod>i\<in>A. m i))" |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
705 |
if "\<And>i j. \<lbrakk>i \<in> A; j \<in> A; i \<noteq> j\<rbrakk> \<Longrightarrow> coprime (m i) (m j)" |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
706 |
and "\<And>i. i \<in> A \<Longrightarrow> [x = y] (mod m i)" for x y :: nat |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
707 |
using that |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
708 |
proof (induct A rule: infinite_finite_induct) |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
709 |
case (insert x A) |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
710 |
then show ?case |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
711 |
by simp (metis coprime_cong_mult_nat prod_coprime_right) |
5b269897df9d
de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents:
67115
diff
changeset
|
712 |
qed auto |
31719 | 713 |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
714 |
lemma chinese_remainder_unique_nat: |
44872 | 715 |
fixes A :: "'a set" |
716 |
and m :: "'a \<Rightarrow> nat" |
|
717 |
and u :: "'a \<Rightarrow> nat" |
|
718 |
assumes fin: "finite A" |
|
61954 | 719 |
and nz: "\<forall>i\<in>A. m i \<noteq> 0" |
66380 | 720 |
and cop: "\<forall>i\<in>A. \<forall>j\<in>A. i \<noteq> j \<longrightarrow> coprime (m i) (m j)" |
63901 | 721 |
shows "\<exists>!x. x < (\<Prod>i\<in>A. m i) \<and> (\<forall>i\<in>A. [x = u i] (mod m i))" |
31719 | 722 |
proof - |
44872 | 723 |
from chinese_remainder_nat [OF fin cop] |
66380 | 724 |
obtain y where one: "(\<forall>i\<in>A. [y = u i] (mod m i))" |
31719 | 725 |
by blast |
61954 | 726 |
let ?x = "y mod (\<Prod>i\<in>A. m i)" |
727 |
from fin nz have prodnz: "(\<Prod>i\<in>A. m i) \<noteq> 0" |
|
31719 | 728 |
by auto |
61954 | 729 |
then have less: "?x < (\<Prod>i\<in>A. m i)" |
31719 | 730 |
by auto |
66380 | 731 |
have cong: "\<forall>i\<in>A. [?x = u i] (mod m i)" |
67085 | 732 |
using fin one |
733 |
by (auto simp add: cong_def dvd_prod_eqI mod_mod_cancel) |
|
66380 | 734 |
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" |
735 |
proof clarify |
|
31719 | 736 |
fix z |
61954 | 737 |
assume zless: "z < (\<Prod>i\<in>A. m i)" |
66380 | 738 |
assume zcong: "(\<forall>i\<in>A. [z = u i] (mod m i))" |
739 |
have "\<forall>i\<in>A. [?x = z] (mod m i)" |
|
67085 | 740 |
using cong zcong by (auto simp add: cong_def) |
61954 | 741 |
with fin cop have "[?x = z] (mod (\<Prod>i\<in>A. m i))" |
67085 | 742 |
by (intro coprime_cong_prod_nat) auto |
31719 | 743 |
with zless less show "z = ?x" |
67085 | 744 |
by (auto simp add: cong_def) |
44872 | 745 |
qed |
66380 | 746 |
from less cong unique show ?thesis |
747 |
by blast |
|
31719 | 748 |
qed |
749 |
||
750 |
end |