| author | haftmann | 
| Thu, 23 Nov 2017 17:03:21 +0000 | |
| changeset 67085 | f5d7f37b4143 | 
| parent 67051 | e7e54a0b9197 | 
| child 67086 | 59d07a95be0e | 
| 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_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  | 
||
157  | 
lemma cong_modulus_minus_iff: "[b = c] (mod - a) \<longleftrightarrow> [b = c] (mod a)"  | 
|
158  | 
using cong_diff_iff_cong_0 [of b c a] cong_diff_iff_cong_0 [of b c " -a"]  | 
|
159  | 
by (simp add: cong_0_iff)  | 
|
| 31719 | 160  | 
|
161  | 
end  | 
|
162  | 
||
| 66380 | 163  | 
|
| 66888 | 164  | 
subsection \<open>Congruences on @{typ nat} and @{typ int}\<close>
 | 
| 31719 | 165  | 
|
| 66888 | 166  | 
lemma cong_int_iff:  | 
167  | 
"[int m = int q] (mod int n) \<longleftrightarrow> [m = q] (mod n)"  | 
|
168  | 
by (simp add: cong_def of_nat_mod [symmetric])  | 
|
169  | 
||
170  | 
lemma cong_Suc_0 [simp, presburger]:  | 
|
171  | 
"[m = n] (mod Suc 0)"  | 
|
172  | 
using cong_1 [of m n] by simp  | 
|
| 31719 | 173  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
174  | 
lemma cong_diff_nat:  | 
| 66888 | 175  | 
"[a - c = b - d] (mod m)" if "[a = b] (mod m)" "[c = d] (mod m)"  | 
| 66954 | 176  | 
and "a \<ge> c" "b \<ge> d" for a b c d m :: nat  | 
177  | 
proof -  | 
|
178  | 
have "[c + (a - c) = d + (b - d)] (mod m)"  | 
|
179  | 
using that by simp  | 
|
180  | 
with \<open>[c = d] (mod m)\<close> have "[c + (a - c) = c + (b - d)] (mod m)"  | 
|
181  | 
using mod_add_cong by (auto simp add: cong_def) fastforce  | 
|
182  | 
then show ?thesis  | 
|
183  | 
by (simp add: cong_def nat_mod_eq_iff)  | 
|
184  | 
qed  | 
|
| 31719 | 185  | 
|
| 66888 | 186  | 
lemma cong_diff_iff_cong_0_nat:  | 
| 66954 | 187  | 
"[a - b = 0] (mod m) \<longleftrightarrow> [a = b] (mod m)" if "a \<ge> b" for a b :: nat  | 
188  | 
using that by (auto simp add: cong_def le_imp_diff_is_add dest: nat_mod_eq_lemma)  | 
|
| 31719 | 189  | 
|
| 66954 | 190  | 
lemma cong_diff_iff_cong_0_nat':  | 
191  | 
"[nat \<bar>int a - int b\<bar> = 0] (mod m) \<longleftrightarrow> [a = b] (mod m)"  | 
|
192  | 
proof (cases "b \<le> a")  | 
|
193  | 
case True  | 
|
194  | 
then show ?thesis  | 
|
195  | 
by (simp add: nat_diff_distrib' cong_diff_iff_cong_0_nat [of b a m])  | 
|
196  | 
next  | 
|
197  | 
case False  | 
|
198  | 
then have "a \<le> b"  | 
|
199  | 
by simp  | 
|
200  | 
then show ?thesis  | 
|
201  | 
by (simp add: nat_diff_distrib' cong_diff_iff_cong_0_nat [of a b m])  | 
|
202  | 
(auto simp add: cong_def)  | 
|
203  | 
qed  | 
|
204  | 
||
205  | 
lemma cong_altdef_nat:  | 
|
206  | 
"a \<ge> b \<Longrightarrow> [a = b] (mod m) \<longleftrightarrow> m dvd (a - b)"  | 
|
| 66380 | 207  | 
for a b :: nat  | 
| 66888 | 208  | 
by (simp add: cong_0_iff [symmetric] cong_diff_iff_cong_0_nat)  | 
| 31719 | 209  | 
|
| 66954 | 210  | 
lemma cong_altdef_nat':  | 
211  | 
"[a = b] (mod m) \<longleftrightarrow> m dvd nat \<bar>int a - int b\<bar>"  | 
|
| 67085 | 212  | 
using cong_diff_iff_cong_0_nat' [of a b m]  | 
213  | 
by (simp only: cong_0_iff [symmetric])  | 
|
| 66954 | 214  | 
|
215  | 
lemma cong_altdef_int:  | 
|
216  | 
"[a = b] (mod m) \<longleftrightarrow> m dvd (a - b)"  | 
|
| 66380 | 217  | 
for a b :: int  | 
| 66888 | 218  | 
by (simp add: cong_0_iff [symmetric] cong_diff_iff_cong_0)  | 
| 31719 | 219  | 
|
| 67085 | 220  | 
lemma cong_abs_int [simp]: "[x = y] (mod \<bar>m\<bar>) \<longleftrightarrow> [x = y] (mod m)"  | 
| 66380 | 221  | 
for x y :: int  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
222  | 
by (simp add: cong_altdef_int)  | 
| 31719 | 223  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
224  | 
lemma cong_square_int:  | 
| 66380 | 225  | 
"prime p \<Longrightarrow> 0 < a \<Longrightarrow> [a * a = 1] (mod p) \<Longrightarrow> [a = 1] (mod p) \<or> [a = - 1] (mod p)"  | 
226  | 
for a :: int  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
227  | 
apply (simp only: cong_altdef_int)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
228  | 
apply (subst prime_dvd_mult_eq_int [symmetric], assumption)  | 
| 36350 | 229  | 
apply (auto simp add: field_simps)  | 
| 44872 | 230  | 
done  | 
| 31719 | 231  | 
|
| 66954 | 232  | 
lemma cong_mult_rcancel_int:  | 
233  | 
"[a * k = b * k] (mod m) \<longleftrightarrow> [a = b] (mod m)"  | 
|
234  | 
if "coprime k m" for a k m :: int  | 
|
| 67051 | 235  | 
using that by (simp add: cong_altdef_int)  | 
236  | 
(metis coprime_commute coprime_dvd_mult_right_iff mult.commute right_diff_distrib')  | 
|
| 31719 | 237  | 
|
| 66954 | 238  | 
lemma cong_mult_rcancel_nat:  | 
239  | 
"[a * k = b * k] (mod m) \<longleftrightarrow> [a = b] (mod m)"  | 
|
240  | 
if "coprime k m" for a k m :: nat  | 
|
241  | 
proof -  | 
|
242  | 
have "[a * k = b * k] (mod m) \<longleftrightarrow> m dvd nat \<bar>int (a * k) - int (b * k)\<bar>"  | 
|
243  | 
by (simp add: cong_altdef_nat')  | 
|
244  | 
also have "\<dots> \<longleftrightarrow> m dvd nat \<bar>(int a - int b) * int k\<bar>"  | 
|
245  | 
by (simp add: algebra_simps)  | 
|
246  | 
also have "\<dots> \<longleftrightarrow> m dvd nat \<bar>int a - int b\<bar> * k"  | 
|
247  | 
by (simp add: abs_mult nat_times_as_int)  | 
|
248  | 
also have "\<dots> \<longleftrightarrow> m dvd nat \<bar>int a - int b\<bar>"  | 
|
| 67051 | 249  | 
by (rule coprime_dvd_mult_left_iff) (use \<open>coprime k m\<close> in \<open>simp add: ac_simps\<close>)  | 
| 66954 | 250  | 
also have "\<dots> \<longleftrightarrow> [a = b] (mod m)"  | 
251  | 
by (simp add: cong_altdef_nat')  | 
|
252  | 
finally show ?thesis .  | 
|
253  | 
qed  | 
|
| 31719 | 254  | 
|
| 66954 | 255  | 
lemma cong_mult_lcancel_int:  | 
256  | 
"[k * a = k * b] (mod m) = [a = b] (mod m)"  | 
|
257  | 
if "coprime k m" for a k m :: int  | 
|
258  | 
using that by (simp add: cong_mult_rcancel_int ac_simps)  | 
|
259  | 
||
260  | 
lemma cong_mult_lcancel_nat:  | 
|
261  | 
"[k * a = k * b] (mod m) = [a = b] (mod m)"  | 
|
262  | 
if "coprime k m" for a k m :: nat  | 
|
263  | 
using that by (simp add: cong_mult_rcancel_nat ac_simps)  | 
|
| 31719 | 264  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
265  | 
lemma coprime_cong_mult_int:  | 
| 66380 | 266  | 
"[a = b] (mod m) \<Longrightarrow> [a = b] (mod n) \<Longrightarrow> coprime m n \<Longrightarrow> [a = b] (mod m * n)"  | 
267  | 
for a b :: int  | 
|
| 66954 | 268  | 
by (simp add: cong_altdef_int divides_mult)  | 
| 31719 | 269  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
270  | 
lemma coprime_cong_mult_nat:  | 
| 66380 | 271  | 
"[a = b] (mod m) \<Longrightarrow> [a = b] (mod n) \<Longrightarrow> coprime m n \<Longrightarrow> [a = b] (mod m * n)"  | 
272  | 
for a b :: nat  | 
|
| 66954 | 273  | 
by (simp add: cong_altdef_nat' divides_mult)  | 
| 31719 | 274  | 
|
| 66380 | 275  | 
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"  | 
276  | 
for a b :: nat  | 
|
| 66888 | 277  | 
by (auto simp add: cong_def)  | 
| 31719 | 278  | 
|
| 66380 | 279  | 
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"  | 
280  | 
for a b :: int  | 
|
| 66888 | 281  | 
by (auto simp add: cong_def)  | 
| 31719 | 282  | 
|
| 66380 | 283  | 
lemma cong_less_unique_nat: "0 < m \<Longrightarrow> (\<exists>!b. 0 \<le> b \<and> b < m \<and> [a = b] (mod m))"  | 
284  | 
for a m :: nat  | 
|
| 66888 | 285  | 
by (auto simp: cong_def) (metis mod_less_divisor mod_mod_trivial)  | 
| 31719 | 286  | 
|
| 66380 | 287  | 
lemma cong_less_unique_int: "0 < m \<Longrightarrow> (\<exists>!b. 0 \<le> b \<and> b < m \<and> [a = b] (mod m))"  | 
288  | 
for a m :: int  | 
|
| 66888 | 289  | 
by (auto simp: cong_def) (metis mod_mod_trivial pos_mod_conj)  | 
| 31719 | 290  | 
|
| 66380 | 291  | 
lemma cong_iff_lin_int: "[a = b] (mod m) \<longleftrightarrow> (\<exists>k. b = a + m * k)"  | 
292  | 
for a b :: int  | 
|
| 66837 | 293  | 
by (auto simp add: cong_altdef_int algebra_simps elim!: dvdE)  | 
294  | 
(simp add: distrib_right [symmetric] add_eq_0_iff)  | 
|
| 31719 | 295  | 
|
| 66380 | 296  | 
lemma cong_iff_lin_nat: "([a = b] (mod m)) \<longleftrightarrow> (\<exists>k1 k2. b + k1 * m = a + k2 * m)"  | 
297  | 
(is "?lhs = ?rhs")  | 
|
298  | 
for a b :: nat  | 
|
299  | 
proof  | 
|
300  | 
assume ?lhs  | 
|
| 55371 | 301  | 
show ?rhs  | 
302  | 
proof (cases "b \<le> a")  | 
|
303  | 
case True  | 
|
| 66380 | 304  | 
with \<open>?lhs\<close> show ?rhs  | 
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
57418 
diff
changeset
 | 
305  | 
by (metis cong_altdef_nat dvd_def le_add_diff_inverse add_0_right mult_0 mult.commute)  | 
| 55371 | 306  | 
next  | 
307  | 
case False  | 
|
| 66380 | 308  | 
with \<open>?lhs\<close> show ?rhs  | 
| 66888 | 309  | 
by (metis cong_def mult.commute nat_le_linear nat_mod_eq_lemma)  | 
| 55371 | 310  | 
qed  | 
311  | 
next  | 
|
312  | 
assume ?rhs  | 
|
313  | 
then show ?lhs  | 
|
| 66888 | 314  | 
by (metis cong_def mult.commute nat_mod_eq_iff)  | 
| 55371 | 315  | 
qed  | 
| 31719 | 316  | 
|
| 66954 | 317  | 
lemma cong_gcd_eq_nat: "[a = b] (mod m) \<Longrightarrow> gcd a m = gcd b m"  | 
318  | 
for a b :: nat  | 
|
319  | 
by (auto simp add: cong_def) (metis gcd_red_nat)  | 
|
320  | 
||
| 66380 | 321  | 
lemma cong_gcd_eq_int: "[a = b] (mod m) \<Longrightarrow> gcd a m = gcd b m"  | 
322  | 
for a b :: int  | 
|
| 66888 | 323  | 
by (auto simp add: cong_def) (metis gcd_red_int)  | 
| 31719 | 324  | 
|
| 66380 | 325  | 
lemma cong_imp_coprime_nat: "[a = b] (mod m) \<Longrightarrow> coprime a m \<Longrightarrow> coprime b m"  | 
326  | 
for a b :: nat  | 
|
| 67085 | 327  | 
by (auto simp add: coprime_iff_gcd_eq_1 dest: cong_gcd_eq_nat)  | 
| 31719 | 328  | 
|
| 66380 | 329  | 
lemma cong_imp_coprime_int: "[a = b] (mod m) \<Longrightarrow> coprime a m \<Longrightarrow> coprime b m"  | 
330  | 
for a b :: int  | 
|
| 67085 | 331  | 
by (auto simp add: coprime_iff_gcd_eq_1 dest: cong_gcd_eq_int)  | 
| 31719 | 332  | 
|
| 66380 | 333  | 
lemma cong_cong_mod_nat: "[a = b] (mod m) \<longleftrightarrow> [a mod m = b mod m] (mod m)"  | 
334  | 
for a b :: nat  | 
|
| 66888 | 335  | 
by simp  | 
| 31719 | 336  | 
|
| 66380 | 337  | 
lemma cong_cong_mod_int: "[a = b] (mod m) \<longleftrightarrow> [a mod m = b mod m] (mod m)"  | 
338  | 
for a b :: int  | 
|
| 66888 | 339  | 
by simp  | 
| 31719 | 340  | 
|
| 66380 | 341  | 
lemma cong_minus_int [iff]: "[a = b] (mod - m) \<longleftrightarrow> [a = b] (mod m)"  | 
342  | 
for a b :: int  | 
|
| 55371 | 343  | 
by (metis cong_iff_lin_int minus_equation_iff mult_minus_left mult_minus_right)  | 
| 31719 | 344  | 
|
| 66380 | 345  | 
lemma cong_add_lcancel_nat: "[a + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)"  | 
346  | 
for a x y :: nat  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
347  | 
by (simp add: cong_iff_lin_nat)  | 
| 31719 | 348  | 
|
| 66380 | 349  | 
lemma cong_add_lcancel_int: "[a + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)"  | 
350  | 
for a x y :: int  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
351  | 
by (simp add: cong_iff_lin_int)  | 
| 31719 | 352  | 
|
| 66380 | 353  | 
lemma cong_add_rcancel_nat: "[x + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)"  | 
354  | 
for a x y :: nat  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
355  | 
by (simp add: cong_iff_lin_nat)  | 
| 31719 | 356  | 
|
| 66380 | 357  | 
lemma cong_add_rcancel_int: "[x + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)"  | 
358  | 
for a x y :: int  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
359  | 
by (simp add: cong_iff_lin_int)  | 
| 31719 | 360  | 
|
| 66380 | 361  | 
lemma cong_add_lcancel_0_nat: "[a + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"  | 
362  | 
for a x :: nat  | 
|
| 67085 | 363  | 
using cong_add_lcancel_nat [of a x 0 n] by simp  | 
| 31719 | 364  | 
|
| 66380 | 365  | 
lemma cong_add_lcancel_0_int: "[a + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"  | 
366  | 
for a x :: int  | 
|
| 67085 | 367  | 
using cong_add_lcancel_int [of a x 0 n] by simp  | 
| 31719 | 368  | 
|
| 66380 | 369  | 
lemma cong_add_rcancel_0_nat: "[x + a = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"  | 
370  | 
for a x :: nat  | 
|
| 67085 | 371  | 
using cong_add_rcancel_nat [of x a 0 n] by simp  | 
| 66380 | 372  | 
|
373  | 
lemma cong_add_rcancel_0_int: "[x + a = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"  | 
|
374  | 
for a x :: int  | 
|
| 67085 | 375  | 
using cong_add_rcancel_int [of x a 0 n] by simp  | 
| 66380 | 376  | 
|
377  | 
lemma cong_dvd_modulus_nat: "[x = y] (mod m) \<Longrightarrow> n dvd m \<Longrightarrow> [x = y] (mod n)"  | 
|
378  | 
for x y :: nat  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
379  | 
apply (auto simp add: cong_iff_lin_nat dvd_def)  | 
| 66380 | 380  | 
apply (rule_tac x= "k1 * k" in exI)  | 
381  | 
apply (rule_tac x= "k2 * k" in exI)  | 
|
| 36350 | 382  | 
apply (simp add: field_simps)  | 
| 44872 | 383  | 
done  | 
| 31719 | 384  | 
|
| 66380 | 385  | 
lemma cong_dvd_modulus_int: "[x = y] (mod m) \<Longrightarrow> n dvd m \<Longrightarrow> [x = y] (mod n)"  | 
386  | 
for x y :: int  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
387  | 
by (auto simp add: cong_altdef_int dvd_def)  | 
| 31719 | 388  | 
|
| 66380 | 389  | 
lemma cong_to_1_nat:  | 
390  | 
fixes a :: nat  | 
|
391  | 
assumes "[a = 1] (mod n)"  | 
|
392  | 
shows "n dvd (a - 1)"  | 
|
393  | 
proof (cases "a = 0")  | 
|
394  | 
case True  | 
|
395  | 
then show ?thesis by force  | 
|
396  | 
next  | 
|
397  | 
case False  | 
|
398  | 
with assms show ?thesis by (metis cong_altdef_nat leI less_one)  | 
|
399  | 
qed  | 
|
400  | 
||
401  | 
lemma cong_0_1_nat': "[0 = Suc 0] (mod n) \<longleftrightarrow> n = Suc 0"  | 
|
| 66888 | 402  | 
by (auto simp: cong_def)  | 
| 66380 | 403  | 
|
404  | 
lemma cong_0_1_nat: "[0 = 1] (mod n) \<longleftrightarrow> n = 1"  | 
|
405  | 
for n :: nat  | 
|
| 66888 | 406  | 
by (auto simp: cong_def)  | 
| 66380 | 407  | 
|
408  | 
lemma cong_0_1_int: "[0 = 1] (mod n) \<longleftrightarrow> n = 1 \<or> n = - 1"  | 
|
409  | 
for n :: int  | 
|
| 66888 | 410  | 
by (auto simp: cong_def zmult_eq_1_iff)  | 
| 66380 | 411  | 
|
412  | 
lemma cong_to_1'_nat: "[a = 1] (mod n) \<longleftrightarrow> a = 0 \<and> n = 1 \<or> (\<exists>m. a = 1 + m * n)"  | 
|
413  | 
for a :: nat  | 
|
414  | 
by (metis add.right_neutral cong_0_1_nat cong_iff_lin_nat cong_to_1_nat  | 
|
415  | 
dvd_div_mult_self leI le_add_diff_inverse less_one mult_eq_if)  | 
|
416  | 
||
417  | 
lemma cong_le_nat: "y \<le> x \<Longrightarrow> [x = y] (mod n) \<longleftrightarrow> (\<exists>q. x = q * n + y)"  | 
|
418  | 
for x y :: nat  | 
|
| 66837 | 419  | 
by (auto simp add: cong_altdef_nat le_imp_diff_is_add elim!: dvdE)  | 
| 66380 | 420  | 
|
421  | 
lemma cong_solve_nat:  | 
|
422  | 
fixes a :: nat  | 
|
423  | 
assumes "a \<noteq> 0"  | 
|
424  | 
shows "\<exists>x. [a * x = gcd a n] (mod n)"  | 
|
425  | 
proof (cases "n = 0")  | 
|
426  | 
case True  | 
|
427  | 
then show ?thesis by force  | 
|
428  | 
next  | 
|
429  | 
case False  | 
|
430  | 
then show ?thesis  | 
|
431  | 
using bezout_nat [of a n, OF \<open>a \<noteq> 0\<close>]  | 
|
| 66888 | 432  | 
by auto (metis cong_add_rcancel_0_nat cong_mult_self_left)  | 
| 66380 | 433  | 
qed  | 
434  | 
||
435  | 
lemma cong_solve_int: "a \<noteq> 0 \<Longrightarrow> \<exists>x. [a * x = gcd a n] (mod n)"  | 
|
436  | 
for a :: int  | 
|
437  | 
apply (cases "n = 0")  | 
|
438  | 
apply (cases "a \<ge> 0")  | 
|
439  | 
apply auto  | 
|
440  | 
apply (rule_tac x = "-1" in exI)  | 
|
441  | 
apply auto  | 
|
442  | 
apply (insert bezout_int [of a n], auto)  | 
|
443  | 
apply (metis cong_iff_lin_int mult.commute)  | 
|
| 44872 | 444  | 
done  | 
| 31719 | 445  | 
|
| 44872 | 446  | 
lemma cong_solve_dvd_nat:  | 
| 66380 | 447  | 
fixes a :: nat  | 
448  | 
assumes a: "a \<noteq> 0" and b: "gcd a n dvd d"  | 
|
449  | 
shows "\<exists>x. [a * x = d] (mod n)"  | 
|
| 31719 | 450  | 
proof -  | 
| 44872 | 451  | 
from cong_solve_nat [OF a] obtain x where "[a * x = gcd a n](mod n)"  | 
| 31719 | 452  | 
by auto  | 
| 44872 | 453  | 
then have "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)"  | 
| 66888 | 454  | 
using cong_scalar_left by blast  | 
| 31719 | 455  | 
also from b have "(d div gcd a n) * gcd a n = d"  | 
456  | 
by (rule dvd_div_mult_self)  | 
|
457  | 
also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)"  | 
|
458  | 
by auto  | 
|
459  | 
finally show ?thesis  | 
|
460  | 
by auto  | 
|
461  | 
qed  | 
|
462  | 
||
| 44872 | 463  | 
lemma cong_solve_dvd_int:  | 
| 31719 | 464  | 
assumes a: "(a::int) \<noteq> 0" and b: "gcd a n dvd d"  | 
| 66380 | 465  | 
shows "\<exists>x. [a * x = d] (mod n)"  | 
| 31719 | 466  | 
proof -  | 
| 44872 | 467  | 
from cong_solve_int [OF a] obtain x where "[a * x = gcd a n](mod n)"  | 
| 31719 | 468  | 
by auto  | 
| 44872 | 469  | 
then have "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)"  | 
| 66888 | 470  | 
using cong_scalar_left by blast  | 
| 31719 | 471  | 
also from b have "(d div gcd a n) * gcd a n = d"  | 
472  | 
by (rule dvd_div_mult_self)  | 
|
473  | 
also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)"  | 
|
474  | 
by auto  | 
|
475  | 
finally show ?thesis  | 
|
476  | 
by auto  | 
|
477  | 
qed  | 
|
478  | 
||
| 66380 | 479  | 
lemma cong_solve_coprime_nat:  | 
| 67051 | 480  | 
"\<exists>x. [a * x = Suc 0] (mod n)" if "coprime a n"  | 
481  | 
using that cong_solve_nat [of a n] by (cases "a = 0") simp_all  | 
|
| 31719 | 482  | 
|
| 67051 | 483  | 
lemma cong_solve_coprime_int:  | 
484  | 
"\<exists>x. [a * x = 1] (mod n)" if "coprime a n" for a n x :: int  | 
|
485  | 
using that cong_solve_int [of a n] by (cases "a = 0")  | 
|
486  | 
(auto simp add: zabs_def split: if_splits)  | 
|
| 55161 | 487  | 
|
| 
62349
 
7c23469b5118
cleansed junk-producing interpretations for gcd/lcm on nat altogether
 
haftmann 
parents: 
62348 
diff
changeset
 | 
488  | 
lemma coprime_iff_invertible_nat:  | 
| 67085 | 489  | 
"coprime a m \<longleftrightarrow> (\<exists>x. [a * x = Suc 0] (mod m))" (is "?P \<longleftrightarrow> ?Q")  | 
490  | 
proof  | 
|
491  | 
assume ?P then show ?Q  | 
|
492  | 
by (auto dest!: cong_solve_coprime_nat)  | 
|
493  | 
next  | 
|
494  | 
assume ?Q  | 
|
495  | 
then obtain b where "[a * b = Suc 0] (mod m)"  | 
|
496  | 
by blast  | 
|
497  | 
with coprime_mod_left_iff [of m "a * b"] show ?P  | 
|
498  | 
by (cases "m = 0 \<or> m = 1")  | 
|
499  | 
(unfold cong_def, auto simp add: cong_def)  | 
|
500  | 
qed  | 
|
| 66380 | 501  | 
|
| 67051 | 502  | 
lemma coprime_iff_invertible_int:  | 
| 67085 | 503  | 
"coprime a m \<longleftrightarrow> (\<exists>x. [a * x = 1] (mod m))" (is "?P \<longleftrightarrow> ?Q") for m :: int  | 
504  | 
proof  | 
|
505  | 
assume ?P then show ?Q  | 
|
506  | 
by (auto dest: cong_solve_coprime_int)  | 
|
507  | 
next  | 
|
508  | 
assume ?Q  | 
|
509  | 
then obtain b where "[a * b = 1] (mod m)"  | 
|
510  | 
by blast  | 
|
511  | 
with coprime_mod_left_iff [of m "a * b"] show ?P  | 
|
512  | 
by (cases "m = 0 \<or> m = 1")  | 
|
513  | 
(unfold cong_def, auto simp add: zmult_eq_1_iff)  | 
|
514  | 
qed  | 
|
| 31719 | 515  | 
|
| 66380 | 516  | 
lemma coprime_iff_invertible'_nat:  | 
517  | 
"m > 0 \<Longrightarrow> coprime a m \<longleftrightarrow> (\<exists>x. 0 \<le> x \<and> x < m \<and> [a * x = Suc 0] (mod m))"  | 
|
| 55161 | 518  | 
apply (subst coprime_iff_invertible_nat)  | 
| 66380 | 519  | 
apply auto  | 
| 66888 | 520  | 
apply (auto simp add: cong_def)  | 
| 55161 | 521  | 
apply (metis mod_less_divisor mod_mult_right_eq)  | 
| 44872 | 522  | 
done  | 
| 31719 | 523  | 
|
| 66380 | 524  | 
lemma coprime_iff_invertible'_int:  | 
525  | 
"m > 0 \<Longrightarrow> coprime a m \<longleftrightarrow> (\<exists>x. 0 \<le> x \<and> x < m \<and> [a * x = 1] (mod m))"  | 
|
526  | 
for m :: int  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
527  | 
apply (subst coprime_iff_invertible_int)  | 
| 66888 | 528  | 
apply (auto simp add: cong_def)  | 
| 55371 | 529  | 
apply (metis mod_mult_right_eq pos_mod_conj)  | 
| 44872 | 530  | 
done  | 
| 31719 | 531  | 
|
| 66380 | 532  | 
lemma cong_cong_lcm_nat: "[x = y] (mod a) \<Longrightarrow> [x = y] (mod b) \<Longrightarrow> [x = y] (mod lcm a b)"  | 
533  | 
for x y :: nat  | 
|
| 44872 | 534  | 
apply (cases "y \<le> x")  | 
| 66888 | 535  | 
apply (simp add: cong_altdef_nat)  | 
536  | 
apply (meson cong_altdef_nat cong_sym lcm_least_iff nat_le_linear)  | 
|
| 44872 | 537  | 
done  | 
| 31719 | 538  | 
|
| 66380 | 539  | 
lemma cong_cong_lcm_int: "[x = y] (mod a) \<Longrightarrow> [x = y] (mod b) \<Longrightarrow> [x = y] (mod lcm a b)"  | 
540  | 
for x y :: int  | 
|
541  | 
by (auto simp add: cong_altdef_int lcm_least)  | 
|
| 31719 | 542  | 
|
| 66888 | 543  | 
lemma cong_cong_prod_coprime_nat:  | 
544  | 
"[x = y] (mod (\<Prod>i\<in>A. m i))" if  | 
|
545  | 
"(\<forall>i\<in>A. [(x::nat) = y] (mod m i))"  | 
|
546  | 
and "(\<forall>i\<in>A. (\<forall>j\<in>A. i \<noteq> j \<longrightarrow> coprime (m i) (m j)))"  | 
|
547  | 
using that apply (induct A rule: infinite_finite_induct)  | 
|
548  | 
apply auto  | 
|
| 67051 | 549  | 
apply (metis coprime_cong_mult_nat prod_coprime_right)  | 
| 44872 | 550  | 
done  | 
| 31719 | 551  | 
|
| 67051 | 552  | 
lemma cong_cong_prod_coprime_int:  | 
| 66888 | 553  | 
"[x = y] (mod (\<Prod>i\<in>A. m i))" if  | 
554  | 
"(\<forall>i\<in>A. [(x::int) = y] (mod m i))"  | 
|
555  | 
"(\<forall>i\<in>A. (\<forall>j\<in>A. i \<noteq> j \<longrightarrow> coprime (m i) (m j)))"  | 
|
556  | 
using that apply (induct A rule: infinite_finite_induct)  | 
|
| 67051 | 557  | 
apply auto  | 
558  | 
apply (metis coprime_cong_mult_int prod_coprime_right)  | 
|
| 44872 | 559  | 
done  | 
| 31719 | 560  | 
|
| 44872 | 561  | 
lemma binary_chinese_remainder_aux_nat:  | 
| 66380 | 562  | 
fixes m1 m2 :: nat  | 
563  | 
assumes a: "coprime m1 m2"  | 
|
564  | 
shows "\<exists>b1 b2. [b1 = 1] (mod m1) \<and> [b1 = 0] (mod m2) \<and> [b2 = 0] (mod m1) \<and> [b2 = 1] (mod m2)"  | 
|
| 31719 | 565  | 
proof -  | 
| 66380 | 566  | 
from cong_solve_coprime_nat [OF a] obtain x1 where 1: "[m1 * x1 = 1] (mod m2)"  | 
| 31719 | 567  | 
by auto  | 
| 44872 | 568  | 
from a have b: "coprime m2 m1"  | 
| 67051 | 569  | 
by (simp add: ac_simps)  | 
| 66380 | 570  | 
from cong_solve_coprime_nat [OF b] obtain x2 where 2: "[m2 * x2 = 1] (mod m1)"  | 
| 31719 | 571  | 
by auto  | 
572  | 
have "[m1 * x1 = 0] (mod m1)"  | 
|
| 66888 | 573  | 
by (simp add: cong_mult_self_left)  | 
| 31719 | 574  | 
moreover have "[m2 * x2 = 0] (mod m2)"  | 
| 66888 | 575  | 
by (simp add: cong_mult_self_left)  | 
| 66380 | 576  | 
ultimately show ?thesis  | 
577  | 
using 1 2 by blast  | 
|
| 31719 | 578  | 
qed  | 
579  | 
||
| 44872 | 580  | 
lemma binary_chinese_remainder_aux_int:  | 
| 66380 | 581  | 
fixes m1 m2 :: int  | 
582  | 
assumes a: "coprime m1 m2"  | 
|
583  | 
shows "\<exists>b1 b2. [b1 = 1] (mod m1) \<and> [b1 = 0] (mod m2) \<and> [b2 = 0] (mod m1) \<and> [b2 = 1] (mod m2)"  | 
|
| 31719 | 584  | 
proof -  | 
| 66380 | 585  | 
from cong_solve_coprime_int [OF a] obtain x1 where 1: "[m1 * x1 = 1] (mod m2)"  | 
| 31719 | 586  | 
by auto  | 
| 44872 | 587  | 
from a have b: "coprime m2 m1"  | 
| 67051 | 588  | 
by (simp add: ac_simps)  | 
| 66380 | 589  | 
from cong_solve_coprime_int [OF b] obtain x2 where 2: "[m2 * x2 = 1] (mod m1)"  | 
| 31719 | 590  | 
by auto  | 
591  | 
have "[m1 * x1 = 0] (mod m1)"  | 
|
| 66888 | 592  | 
by (simp add: cong_mult_self_left)  | 
| 31719 | 593  | 
moreover have "[m2 * x2 = 0] (mod m2)"  | 
| 66888 | 594  | 
by (simp add: cong_mult_self_left)  | 
| 66380 | 595  | 
ultimately show ?thesis  | 
596  | 
using 1 2 by blast  | 
|
| 31719 | 597  | 
qed  | 
598  | 
||
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
599  | 
lemma binary_chinese_remainder_nat:  | 
| 66380 | 600  | 
fixes m1 m2 :: nat  | 
601  | 
assumes a: "coprime m1 m2"  | 
|
602  | 
shows "\<exists>x. [x = u1] (mod m1) \<and> [x = u2] (mod m2)"  | 
|
| 31719 | 603  | 
proof -  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
604  | 
from binary_chinese_remainder_aux_nat [OF a] obtain b1 b2  | 
| 66380 | 605  | 
where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)"  | 
606  | 
and "[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)"  | 
|
| 31719 | 607  | 
by blast  | 
608  | 
let ?x = "u1 * b1 + u2 * b2"  | 
|
609  | 
have "[?x = u1 * 1 + u2 * 0] (mod m1)"  | 
|
| 66888 | 610  | 
using \<open>[b1 = 1] (mod m1)\<close> \<open>[b2 = 0] (mod m1)\<close> cong_add cong_scalar_left by blast  | 
| 44872 | 611  | 
then have "[?x = u1] (mod m1)" by simp  | 
| 31719 | 612  | 
have "[?x = u1 * 0 + u2 * 1] (mod m2)"  | 
| 66888 | 613  | 
using \<open>[b1 = 0] (mod m2)\<close> \<open>[b2 = 1] (mod m2)\<close> cong_add cong_scalar_left by blast  | 
| 66380 | 614  | 
then have "[?x = u2] (mod m2)"  | 
615  | 
by simp  | 
|
616  | 
with \<open>[?x = u1] (mod m1)\<close> show ?thesis  | 
|
617  | 
by blast  | 
|
| 31719 | 618  | 
qed  | 
619  | 
||
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
620  | 
lemma binary_chinese_remainder_int:  | 
| 66380 | 621  | 
fixes m1 m2 :: int  | 
622  | 
assumes a: "coprime m1 m2"  | 
|
623  | 
shows "\<exists>x. [x = u1] (mod m1) \<and> [x = u2] (mod m2)"  | 
|
| 31719 | 624  | 
proof -  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
625  | 
from binary_chinese_remainder_aux_int [OF a] obtain b1 b2  | 
| 66380 | 626  | 
where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)"  | 
627  | 
and "[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)"  | 
|
| 31719 | 628  | 
by blast  | 
629  | 
let ?x = "u1 * b1 + u2 * b2"  | 
|
630  | 
have "[?x = u1 * 1 + u2 * 0] (mod m1)"  | 
|
| 66888 | 631  | 
using \<open>[b1 = 1] (mod m1)\<close> \<open>[b2 = 0] (mod m1)\<close> cong_add cong_scalar_left by blast  | 
| 44872 | 632  | 
then have "[?x = u1] (mod m1)" by simp  | 
| 31719 | 633  | 
have "[?x = u1 * 0 + u2 * 1] (mod m2)"  | 
| 66888 | 634  | 
using \<open>[b1 = 0] (mod m2)\<close> \<open>[b2 = 1] (mod m2)\<close> cong_add cong_scalar_left by blast  | 
| 44872 | 635  | 
then have "[?x = u2] (mod m2)" by simp  | 
| 66380 | 636  | 
with \<open>[?x = u1] (mod m1)\<close> show ?thesis  | 
637  | 
by blast  | 
|
| 31719 | 638  | 
qed  | 
639  | 
||
| 66380 | 640  | 
lemma cong_modulus_mult_nat: "[x = y] (mod m * n) \<Longrightarrow> [x = y] (mod m)"  | 
641  | 
for x y :: nat  | 
|
| 44872 | 642  | 
apply (cases "y \<le> x")  | 
| 67085 | 643  | 
apply (auto simp add: cong_altdef_nat elim: dvd_mult_left)  | 
644  | 
apply (metis cong_def mod_mult_cong_right)  | 
|
| 44872 | 645  | 
done  | 
| 31719 | 646  | 
|
| 66380 | 647  | 
lemma cong_modulus_mult_int: "[x = y] (mod m * n) \<Longrightarrow> [x = y] (mod m)"  | 
648  | 
for x y :: int  | 
|
| 44872 | 649  | 
apply (simp add: cong_altdef_int)  | 
| 31719 | 650  | 
apply (erule dvd_mult_left)  | 
| 44872 | 651  | 
done  | 
| 31719 | 652  | 
|
| 66380 | 653  | 
lemma cong_less_modulus_unique_nat: "[x = y] (mod m) \<Longrightarrow> x < m \<Longrightarrow> y < m \<Longrightarrow> x = y"  | 
654  | 
for x y :: nat  | 
|
| 66888 | 655  | 
by (simp add: cong_def)  | 
| 31719 | 656  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
657  | 
lemma binary_chinese_remainder_unique_nat:  | 
| 66380 | 658  | 
fixes m1 m2 :: nat  | 
659  | 
assumes a: "coprime m1 m2"  | 
|
| 44872 | 660  | 
and nz: "m1 \<noteq> 0" "m2 \<noteq> 0"  | 
| 63901 | 661  | 
shows "\<exists>!x. x < m1 * m2 \<and> [x = u1] (mod m1) \<and> [x = u2] (mod m2)"  | 
| 31719 | 662  | 
proof -  | 
| 66380 | 663  | 
from binary_chinese_remainder_nat [OF a] obtain y  | 
664  | 
where "[y = u1] (mod m1)" and "[y = u2] (mod m2)"  | 
|
| 31719 | 665  | 
by blast  | 
666  | 
let ?x = "y mod (m1 * m2)"  | 
|
667  | 
from nz have less: "?x < m1 * m2"  | 
|
| 44872 | 668  | 
by auto  | 
| 66380 | 669  | 
have 1: "[?x = u1] (mod m1)"  | 
| 66888 | 670  | 
apply (rule cong_trans)  | 
| 66380 | 671  | 
prefer 2  | 
672  | 
apply (rule \<open>[y = u1] (mod m1)\<close>)  | 
|
| 66888 | 673  | 
apply (rule cong_modulus_mult_nat [of _ _ _ m2])  | 
674  | 
apply simp  | 
|
| 31719 | 675  | 
done  | 
| 66380 | 676  | 
have 2: "[?x = u2] (mod m2)"  | 
| 66888 | 677  | 
apply (rule cong_trans)  | 
| 66380 | 678  | 
prefer 2  | 
679  | 
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
 | 
680  | 
apply (subst mult.commute)  | 
| 66888 | 681  | 
apply (rule cong_modulus_mult_nat [of _ _ _ m1])  | 
682  | 
apply simp  | 
|
| 31719 | 683  | 
done  | 
| 66380 | 684  | 
have "\<forall>z. z < m1 * m2 \<and> [z = u1] (mod m1) \<and> [z = u2] (mod m2) \<longrightarrow> z = ?x"  | 
| 44872 | 685  | 
proof clarify  | 
| 31719 | 686  | 
fix z  | 
687  | 
assume "z < m1 * m2"  | 
|
688  | 
assume "[z = u1] (mod m1)" and "[z = u2] (mod m2)"  | 
|
689  | 
have "[?x = z] (mod m1)"  | 
|
| 66888 | 690  | 
apply (rule cong_trans)  | 
| 66380 | 691  | 
apply (rule \<open>[?x = u1] (mod m1)\<close>)  | 
| 66888 | 692  | 
apply (rule cong_sym)  | 
| 60526 | 693  | 
apply (rule \<open>[z = u1] (mod m1)\<close>)  | 
| 31719 | 694  | 
done  | 
695  | 
moreover have "[?x = z] (mod m2)"  | 
|
| 66888 | 696  | 
apply (rule cong_trans)  | 
| 66380 | 697  | 
apply (rule \<open>[?x = u2] (mod m2)\<close>)  | 
| 66888 | 698  | 
apply (rule cong_sym)  | 
| 60526 | 699  | 
apply (rule \<open>[z = u2] (mod m2)\<close>)  | 
| 31719 | 700  | 
done  | 
701  | 
ultimately have "[?x = z] (mod m1 * m2)"  | 
|
| 66888 | 702  | 
using a by (auto intro: coprime_cong_mult_nat simp add: mod_mult_cong_left mod_mult_cong_right)  | 
| 60526 | 703  | 
with \<open>z < m1 * m2\<close> \<open>?x < m1 * m2\<close> show "z = ?x"  | 
| 67085 | 704  | 
by (auto simp add: cong_def)  | 
| 44872 | 705  | 
qed  | 
| 66380 | 706  | 
with less 1 2 show ?thesis by auto  | 
| 31719 | 707  | 
qed  | 
708  | 
||
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
709  | 
lemma chinese_remainder_aux_nat:  | 
| 44872 | 710  | 
fixes A :: "'a set"  | 
711  | 
and m :: "'a \<Rightarrow> nat"  | 
|
712  | 
assumes fin: "finite A"  | 
|
| 66380 | 713  | 
and cop: "\<forall>i \<in> A. (\<forall>j \<in> A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))"  | 
714  | 
  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 | 715  | 
proof (rule finite_set_choice, rule fin, rule ballI)  | 
716  | 
fix i  | 
|
| 66380 | 717  | 
assume "i \<in> A"  | 
| 61954 | 718  | 
  with cop have "coprime (\<Prod>j \<in> A - {i}. m j) (m i)"
 | 
| 67051 | 719  | 
by (intro prod_coprime_left) auto  | 
720  | 
  then have "\<exists>x. [(\<Prod>j \<in> A - {i}. m j) * x = Suc 0] (mod m i)"
 | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
721  | 
by (elim cong_solve_coprime_nat)  | 
| 61954 | 722  | 
  then obtain x where "[(\<Prod>j \<in> A - {i}. m j) * x = 1] (mod m i)"
 | 
| 31719 | 723  | 
by auto  | 
| 66380 | 724  | 
  moreover have "[(\<Prod>j \<in> A - {i}. m j) * x = 0] (mod (\<Prod>j \<in> A - {i}. m j))"
 | 
| 66888 | 725  | 
by (simp add: cong_0_iff)  | 
| 66380 | 726  | 
  ultimately show "\<exists>a. [a = 1] (mod m i) \<and> [a = 0] (mod prod m (A - {i}))"
 | 
| 31719 | 727  | 
by blast  | 
728  | 
qed  | 
|
729  | 
||
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
730  | 
lemma chinese_remainder_nat:  | 
| 44872 | 731  | 
fixes A :: "'a set"  | 
732  | 
and m :: "'a \<Rightarrow> nat"  | 
|
733  | 
and u :: "'a \<Rightarrow> nat"  | 
|
734  | 
assumes fin: "finite A"  | 
|
| 66380 | 735  | 
and cop: "\<forall>i \<in> A. \<forall>j \<in> A. i \<noteq> j \<longrightarrow> coprime (m i) (m j)"  | 
736  | 
shows "\<exists>x. \<forall>i \<in> A. [x = u i] (mod m i)"  | 
|
| 31719 | 737  | 
proof -  | 
| 66380 | 738  | 
from chinese_remainder_aux_nat [OF fin cop]  | 
739  | 
  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 | 740  | 
by blast  | 
| 61954 | 741  | 
let ?x = "\<Sum>i\<in>A. (u i) * (b i)"  | 
| 66380 | 742  | 
show ?thesis  | 
| 31719 | 743  | 
proof (rule exI, clarify)  | 
744  | 
fix i  | 
|
| 66380 | 745  | 
assume a: "i \<in> A"  | 
| 44872 | 746  | 
show "[?x = u i] (mod m i)"  | 
| 31719 | 747  | 
proof -  | 
| 66380 | 748  | 
      from fin a have "?x = (\<Sum>j \<in> {i}. u j * b j) + (\<Sum>j \<in> A - {i}. u j * b j)"
 | 
749  | 
by (subst sum.union_disjoint [symmetric]) (auto intro: sum.cong)  | 
|
| 61954 | 750  | 
      then have "[?x = u i * b i + (\<Sum>j \<in> A - {i}. u j * b j)] (mod m i)"
 | 
| 31719 | 751  | 
by auto  | 
| 61954 | 752  | 
      also have "[u i * b i + (\<Sum>j \<in> A - {i}. u j * b j) =
 | 
753  | 
                  u i * 1 + (\<Sum>j \<in> A - {i}. u j * 0)] (mod m i)"
 | 
|
| 66888 | 754  | 
apply (rule cong_add)  | 
755  | 
apply (rule cong_scalar_left)  | 
|
| 66380 | 756  | 
using b a apply blast  | 
| 66888 | 757  | 
apply (rule cong_sum)  | 
758  | 
apply (rule cong_scalar_left)  | 
|
| 67085 | 759  | 
using b apply (auto simp add: mod_eq_0_iff_dvd cong_def)  | 
760  | 
        apply (rule dvd_trans [of _ "prod m (A - {x})" "b x" for x])
 | 
|
761  | 
using a fin apply auto  | 
|
| 31719 | 762  | 
done  | 
763  | 
finally show ?thesis  | 
|
764  | 
by simp  | 
|
765  | 
qed  | 
|
766  | 
qed  | 
|
767  | 
qed  | 
|
768  | 
||
| 66888 | 769  | 
lemma coprime_cong_prod_nat:  | 
770  | 
"[x = y] (mod (\<Prod>i\<in>A. m i))"  | 
|
771  | 
if "\<forall>i\<in>A. (\<forall>j\<in>A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))"  | 
|
772  | 
and "\<forall>i\<in>A. [x = y] (mod m i)" for x y :: nat  | 
|
773  | 
using that apply (induct A rule: infinite_finite_induct)  | 
|
| 67051 | 774  | 
apply auto  | 
775  | 
apply (metis coprime_cong_mult_nat prod_coprime_right)  | 
|
| 44872 | 776  | 
done  | 
| 31719 | 777  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
778  | 
lemma chinese_remainder_unique_nat:  | 
| 44872 | 779  | 
fixes A :: "'a set"  | 
780  | 
and m :: "'a \<Rightarrow> nat"  | 
|
781  | 
and u :: "'a \<Rightarrow> nat"  | 
|
782  | 
assumes fin: "finite A"  | 
|
| 61954 | 783  | 
and nz: "\<forall>i\<in>A. m i \<noteq> 0"  | 
| 66380 | 784  | 
and cop: "\<forall>i\<in>A. \<forall>j\<in>A. i \<noteq> j \<longrightarrow> coprime (m i) (m j)"  | 
| 63901 | 785  | 
shows "\<exists>!x. x < (\<Prod>i\<in>A. m i) \<and> (\<forall>i\<in>A. [x = u i] (mod m i))"  | 
| 31719 | 786  | 
proof -  | 
| 44872 | 787  | 
from chinese_remainder_nat [OF fin cop]  | 
| 66380 | 788  | 
obtain y where one: "(\<forall>i\<in>A. [y = u i] (mod m i))"  | 
| 31719 | 789  | 
by blast  | 
| 61954 | 790  | 
let ?x = "y mod (\<Prod>i\<in>A. m i)"  | 
791  | 
from fin nz have prodnz: "(\<Prod>i\<in>A. m i) \<noteq> 0"  | 
|
| 31719 | 792  | 
by auto  | 
| 61954 | 793  | 
then have less: "?x < (\<Prod>i\<in>A. m i)"  | 
| 31719 | 794  | 
by auto  | 
| 66380 | 795  | 
have cong: "\<forall>i\<in>A. [?x = u i] (mod m i)"  | 
| 67085 | 796  | 
using fin one  | 
797  | 
by (auto simp add: cong_def dvd_prod_eqI mod_mod_cancel)  | 
|
| 66380 | 798  | 
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"  | 
799  | 
proof clarify  | 
|
| 31719 | 800  | 
fix z  | 
| 61954 | 801  | 
assume zless: "z < (\<Prod>i\<in>A. m i)"  | 
| 66380 | 802  | 
assume zcong: "(\<forall>i\<in>A. [z = u i] (mod m i))"  | 
803  | 
have "\<forall>i\<in>A. [?x = z] (mod m i)"  | 
|
| 67085 | 804  | 
using cong zcong by (auto simp add: cong_def)  | 
| 61954 | 805  | 
with fin cop have "[?x = z] (mod (\<Prod>i\<in>A. m i))"  | 
| 67085 | 806  | 
by (intro coprime_cong_prod_nat) auto  | 
| 31719 | 807  | 
with zless less show "z = ?x"  | 
| 67085 | 808  | 
by (auto simp add: cong_def)  | 
| 44872 | 809  | 
qed  | 
| 66380 | 810  | 
from less cong unique show ?thesis  | 
811  | 
by blast  | 
|
| 31719 | 812  | 
qed  | 
813  | 
||
814  | 
end  |