| author | blanchet | 
| Mon, 20 Jan 2014 18:25:44 +0100 | |
| changeset 55077 | 4cf280104b85 | 
| parent 54489 | 03ff4d1e6784 | 
| child 55130 | 70db8d380d62 | 
| permissions | -rw-r--r-- | 
| 41959 | 1  | 
(* Title: HOL/Number_Theory/Cong.thy  | 
| 31719 | 2  | 
Authors: Christophe Tabacznyj, Lawrence C. Paulson, Amine Chaieb,  | 
3  | 
Thomas M. Rasmussen, Jeremy Avigad  | 
|
4  | 
||
5  | 
Defines congruence (notation: [x = y] (mod z)) for natural numbers and  | 
|
6  | 
integers.  | 
|
7  | 
||
8  | 
This file combines and revises a number of prior developments.  | 
|
9  | 
||
10  | 
The original theories "GCD" and "Primes" were by Christophe Tabacznyj  | 
|
11  | 
and Lawrence C. Paulson, based on \cite{davenport92}. They introduced
 | 
|
12  | 
gcd, lcm, and prime for the natural numbers.  | 
|
13  | 
||
14  | 
The original theory "IntPrimes" was by Thomas M. Rasmussen, and  | 
|
15  | 
extended gcd, lcm, primes to the integers. Amine Chaieb provided  | 
|
16  | 
another extension of the notions to the integers, and added a number  | 
|
| 44872 | 17  | 
of results to "Primes" and "GCD".  | 
| 31719 | 18  | 
|
19  | 
The original theory, "IntPrimes", by Thomas M. Rasmussen, defined and  | 
|
20  | 
developed the congruence relations on the integers. The notion was  | 
|
| 33718 | 21  | 
extended to the natural numbers by Chaieb. Jeremy Avigad combined  | 
| 31719 | 22  | 
these, revised and tidied them, made the development uniform for the  | 
23  | 
natural numbers and the integers, and added a number of new theorems.  | 
|
24  | 
*)  | 
|
25  | 
||
26  | 
header {* Congruence *}
 | 
|
27  | 
||
28  | 
theory Cong  | 
|
| 37293 | 29  | 
imports Primes  | 
| 31719 | 30  | 
begin  | 
31  | 
||
| 44872 | 32  | 
subsection {* Turn off @{text One_nat_def} *}
 | 
| 31719 | 33  | 
|
| 44872 | 34  | 
lemma induct'_nat [case_names zero plus1, induct type: nat]:  | 
35  | 
"P (0::nat) \<Longrightarrow> (\<And>n. P n \<Longrightarrow> P (n + 1)) \<Longrightarrow> P n"  | 
|
36  | 
by (induct n) (simp_all add: One_nat_def)  | 
|
| 31719 | 37  | 
|
| 44872 | 38  | 
lemma cases_nat [case_names zero plus1, cases type: nat]:  | 
39  | 
"P (0::nat) \<Longrightarrow> (\<And>n. P (n + 1)) \<Longrightarrow> P n"  | 
|
40  | 
by (rule induct'_nat)  | 
|
| 31719 | 41  | 
|
42  | 
lemma power_plus_one [simp]: "(x::'a::power)^(n + 1) = x * x^n"  | 
|
| 44872 | 43  | 
by (simp add: One_nat_def)  | 
| 31719 | 44  | 
|
| 44872 | 45  | 
lemma power_eq_one_eq_nat [simp]: "((x::nat)^m = 1) = (m = 0 | x = 1)"  | 
46  | 
by (induct m) auto  | 
|
| 31719 | 47  | 
|
48  | 
lemma card_insert_if' [simp]: "finite A \<Longrightarrow>  | 
|
| 44872 | 49  | 
card (insert x A) = (if x \<in> A then (card A) else (card A) + 1)"  | 
50  | 
by (auto simp add: insert_absorb)  | 
|
| 31719 | 51  | 
|
52  | 
lemma nat_1' [simp]: "nat 1 = 1"  | 
|
| 44872 | 53  | 
by simp  | 
| 31719 | 54  | 
|
| 31792 | 55  | 
(* For those annoying moments where Suc reappears, use Suc_eq_plus1 *)  | 
| 31719 | 56  | 
|
57  | 
declare nat_1 [simp del]  | 
|
| 44872 | 58  | 
declare add_2_eq_Suc [simp del]  | 
| 31719 | 59  | 
declare add_2_eq_Suc' [simp del]  | 
60  | 
||
61  | 
||
62  | 
declare mod_pos_pos_trivial [simp]  | 
|
63  | 
||
64  | 
||
65  | 
subsection {* Main definitions *}
 | 
|
66  | 
||
67  | 
class cong =  | 
|
| 44872 | 68  | 
  fixes cong :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" ("(1[_ = _] '(mod _'))")
 | 
| 31719 | 69  | 
begin  | 
70  | 
||
| 44872 | 71  | 
abbreviation notcong :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"  ("(1[_ \<noteq> _] '(mod _'))")
 | 
72  | 
where "notcong x y m \<equiv> \<not> cong x y m"  | 
|
| 31719 | 73  | 
|
74  | 
end  | 
|
75  | 
||
76  | 
(* definitions for the natural numbers *)  | 
|
77  | 
||
78  | 
instantiation nat :: cong  | 
|
| 44872 | 79  | 
begin  | 
| 31719 | 80  | 
|
| 44872 | 81  | 
definition cong_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool"  | 
82  | 
where "cong_nat x y m = ((x mod m) = (y mod m))"  | 
|
| 31719 | 83  | 
|
| 44872 | 84  | 
instance ..  | 
| 31719 | 85  | 
|
86  | 
end  | 
|
87  | 
||
88  | 
||
89  | 
(* definitions for the integers *)  | 
|
90  | 
||
91  | 
instantiation int :: cong  | 
|
| 44872 | 92  | 
begin  | 
| 31719 | 93  | 
|
| 44872 | 94  | 
definition cong_int :: "int \<Rightarrow> int \<Rightarrow> int \<Rightarrow> bool"  | 
95  | 
where "cong_int x y m = ((x mod m) = (y mod m))"  | 
|
| 31719 | 96  | 
|
| 44872 | 97  | 
instance ..  | 
| 31719 | 98  | 
|
99  | 
end  | 
|
100  | 
||
101  | 
||
102  | 
subsection {* Set up Transfer *}
 | 
|
103  | 
||
104  | 
||
105  | 
lemma transfer_nat_int_cong:  | 
|
| 44872 | 106  | 
"(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> m >= 0 \<Longrightarrow>  | 
| 31719 | 107  | 
([(nat x) = (nat y)] (mod (nat m))) = ([x = y] (mod m))"  | 
| 44872 | 108  | 
unfolding cong_int_def cong_nat_def  | 
| 31719 | 109  | 
apply (auto simp add: nat_mod_distrib [symmetric])  | 
110  | 
apply (subst (asm) eq_nat_nat_iff)  | 
|
| 44872 | 111  | 
apply (cases "m = 0", force, rule pos_mod_sign, force)+  | 
| 31719 | 112  | 
apply assumption  | 
| 44872 | 113  | 
done  | 
| 31719 | 114  | 
|
| 44872 | 115  | 
declare transfer_morphism_nat_int[transfer add return:  | 
| 31719 | 116  | 
transfer_nat_int_cong]  | 
117  | 
||
118  | 
lemma transfer_int_nat_cong:  | 
|
119  | 
"[(int x) = (int y)] (mod (int m)) = [x = y] (mod m)"  | 
|
120  | 
apply (auto simp add: cong_int_def cong_nat_def)  | 
|
121  | 
apply (auto simp add: zmod_int [symmetric])  | 
|
| 44872 | 122  | 
done  | 
| 31719 | 123  | 
|
| 44872 | 124  | 
declare transfer_morphism_int_nat[transfer add return:  | 
| 31719 | 125  | 
transfer_int_nat_cong]  | 
126  | 
||
127  | 
||
128  | 
subsection {* Congruence *}
 | 
|
129  | 
||
130  | 
(* was zcong_0, etc. *)  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
131  | 
lemma cong_0_nat [simp, presburger]: "([(a::nat) = b] (mod 0)) = (a = b)"  | 
| 44872 | 132  | 
unfolding cong_nat_def by auto  | 
| 31719 | 133  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
134  | 
lemma cong_0_int [simp, presburger]: "([(a::int) = b] (mod 0)) = (a = b)"  | 
| 44872 | 135  | 
unfolding cong_int_def by auto  | 
| 31719 | 136  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
137  | 
lemma cong_1_nat [simp, presburger]: "[(a::nat) = b] (mod 1)"  | 
| 44872 | 138  | 
unfolding cong_nat_def by auto  | 
| 31719 | 139  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
140  | 
lemma cong_Suc_0_nat [simp, presburger]: "[(a::nat) = b] (mod Suc 0)"  | 
| 44872 | 141  | 
unfolding cong_nat_def by (auto simp add: One_nat_def)  | 
| 31719 | 142  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
143  | 
lemma cong_1_int [simp, presburger]: "[(a::int) = b] (mod 1)"  | 
| 44872 | 144  | 
unfolding cong_int_def by auto  | 
| 31719 | 145  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
146  | 
lemma cong_refl_nat [simp]: "[(k::nat) = k] (mod m)"  | 
| 44872 | 147  | 
unfolding cong_nat_def by auto  | 
| 31719 | 148  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
149  | 
lemma cong_refl_int [simp]: "[(k::int) = k] (mod m)"  | 
| 44872 | 150  | 
unfolding cong_int_def by auto  | 
| 31719 | 151  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
152  | 
lemma cong_sym_nat: "[(a::nat) = b] (mod m) \<Longrightarrow> [b = a] (mod m)"  | 
| 44872 | 153  | 
unfolding cong_nat_def by auto  | 
| 31719 | 154  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
155  | 
lemma cong_sym_int: "[(a::int) = b] (mod m) \<Longrightarrow> [b = a] (mod m)"  | 
| 44872 | 156  | 
unfolding cong_int_def by auto  | 
| 31719 | 157  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
158  | 
lemma cong_sym_eq_nat: "[(a::nat) = b] (mod m) = [b = a] (mod m)"  | 
| 44872 | 159  | 
unfolding cong_nat_def by auto  | 
| 31719 | 160  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
161  | 
lemma cong_sym_eq_int: "[(a::int) = b] (mod m) = [b = a] (mod m)"  | 
| 44872 | 162  | 
unfolding cong_int_def by auto  | 
| 31719 | 163  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
164  | 
lemma cong_trans_nat [trans]:  | 
| 31719 | 165  | 
"[(a::nat) = b] (mod m) \<Longrightarrow> [b = c] (mod m) \<Longrightarrow> [a = c] (mod m)"  | 
| 44872 | 166  | 
unfolding cong_nat_def by auto  | 
| 31719 | 167  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
168  | 
lemma cong_trans_int [trans]:  | 
| 31719 | 169  | 
"[(a::int) = b] (mod m) \<Longrightarrow> [b = c] (mod m) \<Longrightarrow> [a = c] (mod m)"  | 
| 44872 | 170  | 
unfolding cong_int_def by auto  | 
| 31719 | 171  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
172  | 
lemma cong_add_nat:  | 
| 31719 | 173  | 
"[(a::nat) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a + c = b + d] (mod m)"  | 
174  | 
apply (unfold cong_nat_def)  | 
|
175  | 
apply (subst (1 2) mod_add_eq)  | 
|
176  | 
apply simp  | 
|
| 44872 | 177  | 
done  | 
| 31719 | 178  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
179  | 
lemma cong_add_int:  | 
| 31719 | 180  | 
"[(a::int) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a + c = b + d] (mod m)"  | 
181  | 
apply (unfold cong_int_def)  | 
|
182  | 
apply (subst (1 2) mod_add_left_eq)  | 
|
183  | 
apply (subst (1 2) mod_add_right_eq)  | 
|
184  | 
apply simp  | 
|
| 44872 | 185  | 
done  | 
| 31719 | 186  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
187  | 
lemma cong_diff_int:  | 
| 31719 | 188  | 
"[(a::int) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a - c = b - d] (mod m)"  | 
189  | 
apply (unfold cong_int_def)  | 
|
190  | 
apply (subst (1 2) mod_diff_eq)  | 
|
191  | 
apply simp  | 
|
| 44872 | 192  | 
done  | 
| 31719 | 193  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
194  | 
lemma cong_diff_aux_int:  | 
| 44872 | 195  | 
"(a::int) >= c \<Longrightarrow> b >= d \<Longrightarrow> [(a::int) = b] (mod m) \<Longrightarrow>  | 
| 31719 | 196  | 
[c = d] (mod m) \<Longrightarrow> [tsub a c = tsub b d] (mod m)"  | 
197  | 
apply (subst (1 2) tsub_eq)  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
198  | 
apply (auto intro: cong_diff_int)  | 
| 44872 | 199  | 
done  | 
| 31719 | 200  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
201  | 
lemma cong_diff_nat:  | 
| 31719 | 202  | 
assumes "(a::nat) >= c" and "b >= d" and "[a = b] (mod m)" and  | 
203  | 
"[c = d] (mod m)"  | 
|
204  | 
shows "[a - c = b - d] (mod m)"  | 
|
| 41541 | 205  | 
using assms by (rule cong_diff_aux_int [transferred]);  | 
| 31719 | 206  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
207  | 
lemma cong_mult_nat:  | 
| 31719 | 208  | 
"[(a::nat) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a * c = b * d] (mod m)"  | 
209  | 
apply (unfold cong_nat_def)  | 
|
210  | 
apply (subst (1 2) mod_mult_eq)  | 
|
211  | 
apply simp  | 
|
| 44872 | 212  | 
done  | 
| 31719 | 213  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
214  | 
lemma cong_mult_int:  | 
| 31719 | 215  | 
"[(a::int) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a * c = b * d] (mod m)"  | 
216  | 
apply (unfold cong_int_def)  | 
|
| 47163 | 217  | 
apply (subst (1 2) mod_mult_right_eq)  | 
| 31719 | 218  | 
apply (subst (1 2) mult_commute)  | 
| 47163 | 219  | 
apply (subst (1 2) mod_mult_right_eq)  | 
| 31719 | 220  | 
apply simp  | 
| 41541 | 221  | 
done  | 
| 31719 | 222  | 
|
| 44872 | 223  | 
lemma cong_exp_nat: "[(x::nat) = y] (mod n) \<Longrightarrow> [x^k = y^k] (mod n)"  | 
224  | 
by (induct k) (auto simp add: cong_mult_nat)  | 
|
225  | 
||
226  | 
lemma cong_exp_int: "[(x::int) = y] (mod n) \<Longrightarrow> [x^k = y^k] (mod n)"  | 
|
227  | 
by (induct k) (auto simp add: cong_mult_int)  | 
|
228  | 
||
229  | 
lemma cong_setsum_nat [rule_format]:  | 
|
230  | 
"(ALL x: A. [((f x)::nat) = g x] (mod m)) \<longrightarrow>  | 
|
| 31719 | 231  | 
[(SUM x:A. f x) = (SUM x:A. g x)] (mod m)"  | 
| 44872 | 232  | 
apply (cases "finite A")  | 
| 31719 | 233  | 
apply (induct set: finite)  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
234  | 
apply (auto intro: cong_add_nat)  | 
| 44872 | 235  | 
done  | 
| 31719 | 236  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
237  | 
lemma cong_setsum_int [rule_format]:  | 
| 44872 | 238  | 
"(ALL x: A. [((f x)::int) = g x] (mod m)) \<longrightarrow>  | 
| 31719 | 239  | 
[(SUM x:A. f x) = (SUM x:A. g x)] (mod m)"  | 
| 44872 | 240  | 
apply (cases "finite A")  | 
| 31719 | 241  | 
apply (induct set: finite)  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
242  | 
apply (auto intro: cong_add_int)  | 
| 44872 | 243  | 
done  | 
| 31719 | 244  | 
|
| 44872 | 245  | 
lemma cong_setprod_nat [rule_format]:  | 
246  | 
"(ALL x: A. [((f x)::nat) = g x] (mod m)) \<longrightarrow>  | 
|
| 31719 | 247  | 
[(PROD x:A. f x) = (PROD x:A. g x)] (mod m)"  | 
| 44872 | 248  | 
apply (cases "finite A")  | 
| 31719 | 249  | 
apply (induct set: finite)  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
250  | 
apply (auto intro: cong_mult_nat)  | 
| 44872 | 251  | 
done  | 
| 31719 | 252  | 
|
| 44872 | 253  | 
lemma cong_setprod_int [rule_format]:  | 
254  | 
"(ALL x: A. [((f x)::int) = g x] (mod m)) \<longrightarrow>  | 
|
| 31719 | 255  | 
[(PROD x:A. f x) = (PROD x:A. g x)] (mod m)"  | 
| 44872 | 256  | 
apply (cases "finite A")  | 
| 31719 | 257  | 
apply (induct set: finite)  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
258  | 
apply (auto intro: cong_mult_int)  | 
| 44872 | 259  | 
done  | 
| 31719 | 260  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
261  | 
lemma cong_scalar_nat: "[(a::nat)= b] (mod m) \<Longrightarrow> [a * k = b * k] (mod m)"  | 
| 44872 | 262  | 
by (rule cong_mult_nat) simp_all  | 
| 31719 | 263  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
264  | 
lemma cong_scalar_int: "[(a::int)= b] (mod m) \<Longrightarrow> [a * k = b * k] (mod m)"  | 
| 44872 | 265  | 
by (rule cong_mult_int) simp_all  | 
| 31719 | 266  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
267  | 
lemma cong_scalar2_nat: "[(a::nat)= b] (mod m) \<Longrightarrow> [k * a = k * b] (mod m)"  | 
| 44872 | 268  | 
by (rule cong_mult_nat) simp_all  | 
| 31719 | 269  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
270  | 
lemma cong_scalar2_int: "[(a::int)= b] (mod m) \<Longrightarrow> [k * a = k * b] (mod m)"  | 
| 44872 | 271  | 
by (rule cong_mult_int) simp_all  | 
| 31719 | 272  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
273  | 
lemma cong_mult_self_nat: "[(a::nat) * m = 0] (mod m)"  | 
| 44872 | 274  | 
unfolding cong_nat_def by auto  | 
| 31719 | 275  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
276  | 
lemma cong_mult_self_int: "[(a::int) * m = 0] (mod m)"  | 
| 44872 | 277  | 
unfolding cong_int_def by auto  | 
| 31719 | 278  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
279  | 
lemma cong_eq_diff_cong_0_int: "[(a::int) = b] (mod m) = [a - b = 0] (mod m)"  | 
| 31719 | 280  | 
apply (rule iffI)  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
281  | 
apply (erule cong_diff_int [of a b m b b, simplified])  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
282  | 
apply (erule cong_add_int [of "a - b" 0 m b b, simplified])  | 
| 44872 | 283  | 
done  | 
| 31719 | 284  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
285  | 
lemma cong_eq_diff_cong_0_aux_int: "a >= b \<Longrightarrow>  | 
| 31719 | 286  | 
[(a::int) = b] (mod m) = [tsub a b = 0] (mod m)"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
287  | 
by (subst tsub_eq, assumption, rule cong_eq_diff_cong_0_int)  | 
| 31719 | 288  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
289  | 
lemma cong_eq_diff_cong_0_nat:  | 
| 31719 | 290  | 
assumes "(a::nat) >= b"  | 
291  | 
shows "[a = b] (mod m) = [a - b = 0] (mod m)"  | 
|
| 41541 | 292  | 
using assms by (rule cong_eq_diff_cong_0_aux_int [transferred])  | 
| 31719 | 293  | 
|
| 44872 | 294  | 
lemma cong_diff_cong_0'_nat:  | 
295  | 
"[(x::nat) = y] (mod n) \<longleftrightarrow>  | 
|
| 31719 | 296  | 
(if x <= y then [y - x = 0] (mod n) else [x - y = 0] (mod n))"  | 
| 44872 | 297  | 
apply (cases "y <= x")  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
298  | 
apply (frule cong_eq_diff_cong_0_nat [where m = n])  | 
| 31719 | 299  | 
apply auto [1]  | 
300  | 
apply (subgoal_tac "x <= y")  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
301  | 
apply (frule cong_eq_diff_cong_0_nat [where m = n])  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
302  | 
apply (subst cong_sym_eq_nat)  | 
| 31719 | 303  | 
apply auto  | 
| 44872 | 304  | 
done  | 
| 31719 | 305  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
306  | 
lemma cong_altdef_nat: "(a::nat) >= b \<Longrightarrow> [a = b] (mod m) = (m dvd (a - b))"  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
307  | 
apply (subst cong_eq_diff_cong_0_nat, assumption)  | 
| 31719 | 308  | 
apply (unfold cong_nat_def)  | 
309  | 
apply (simp add: dvd_eq_mod_eq_0 [symmetric])  | 
|
| 44872 | 310  | 
done  | 
| 31719 | 311  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
312  | 
lemma cong_altdef_int: "[(a::int) = b] (mod m) = (m dvd (a - b))"  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
313  | 
apply (subst cong_eq_diff_cong_0_int)  | 
| 31719 | 314  | 
apply (unfold cong_int_def)  | 
315  | 
apply (simp add: dvd_eq_mod_eq_0 [symmetric])  | 
|
| 44872 | 316  | 
done  | 
| 31719 | 317  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
318  | 
lemma cong_abs_int: "[(x::int) = y] (mod abs m) = [x = y] (mod m)"  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
319  | 
by (simp add: cong_altdef_int)  | 
| 31719 | 320  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
321  | 
lemma cong_square_int:  | 
| 31719 | 322  | 
"\<lbrakk> prime (p::int); 0 < a; [a * a = 1] (mod p) \<rbrakk>  | 
323  | 
\<Longrightarrow> [a = 1] (mod p) \<or> [a = - 1] (mod p)"  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
324  | 
apply (simp only: cong_altdef_int)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
325  | 
apply (subst prime_dvd_mult_eq_int [symmetric], assumption)  | 
| 36350 | 326  | 
apply (auto simp add: field_simps)  | 
| 44872 | 327  | 
done  | 
| 31719 | 328  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
329  | 
lemma cong_mult_rcancel_int:  | 
| 44872 | 330  | 
"coprime k (m::int) \<Longrightarrow> [a * k = b * k] (mod m) = [a = b] (mod m)"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
331  | 
apply (subst (1 2) cong_altdef_int)  | 
| 31719 | 332  | 
apply (subst left_diff_distrib [symmetric])  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
333  | 
apply (rule coprime_dvd_mult_iff_int)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
334  | 
apply (subst gcd_commute_int, assumption)  | 
| 44872 | 335  | 
done  | 
| 31719 | 336  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
337  | 
lemma cong_mult_rcancel_nat:  | 
| 31719 | 338  | 
assumes "coprime k (m::nat)"  | 
339  | 
shows "[a * k = b * k] (mod m) = [a = b] (mod m)"  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
340  | 
apply (rule cong_mult_rcancel_int [transferred])  | 
| 41541 | 341  | 
using assms apply auto  | 
| 44872 | 342  | 
done  | 
| 31719 | 343  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
344  | 
lemma cong_mult_lcancel_nat:  | 
| 44872 | 345  | 
"coprime k (m::nat) \<Longrightarrow> [k * a = k * b ] (mod m) = [a = b] (mod m)"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
346  | 
by (simp add: mult_commute cong_mult_rcancel_nat)  | 
| 31719 | 347  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
348  | 
lemma cong_mult_lcancel_int:  | 
| 44872 | 349  | 
"coprime k (m::int) \<Longrightarrow> [k * a = k * b] (mod m) = [a = b] (mod m)"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
350  | 
by (simp add: mult_commute cong_mult_rcancel_int)  | 
| 31719 | 351  | 
|
352  | 
(* was zcong_zgcd_zmult_zmod *)  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
353  | 
lemma coprime_cong_mult_int:  | 
| 31719 | 354  | 
"[(a::int) = b] (mod m) \<Longrightarrow> [a = b] (mod n) \<Longrightarrow> coprime m n  | 
355  | 
\<Longrightarrow> [a = b] (mod m * n)"  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
356  | 
apply (simp only: cong_altdef_int)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
357  | 
apply (erule (2) divides_mult_int)  | 
| 41541 | 358  | 
done  | 
| 31719 | 359  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
360  | 
lemma coprime_cong_mult_nat:  | 
| 31719 | 361  | 
assumes "[(a::nat) = b] (mod m)" and "[a = b] (mod n)" and "coprime m n"  | 
362  | 
shows "[a = b] (mod m * n)"  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
363  | 
apply (rule coprime_cong_mult_int [transferred])  | 
| 41541 | 364  | 
using assms apply auto  | 
365  | 
done  | 
|
| 31719 | 366  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
367  | 
lemma cong_less_imp_eq_nat: "0 \<le> (a::nat) \<Longrightarrow>  | 
| 31719 | 368  | 
a < m \<Longrightarrow> 0 \<le> b \<Longrightarrow> b < m \<Longrightarrow> [a = b] (mod m) \<Longrightarrow> a = b"  | 
| 41541 | 369  | 
by (auto simp add: cong_nat_def)  | 
| 31719 | 370  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
371  | 
lemma cong_less_imp_eq_int: "0 \<le> (a::int) \<Longrightarrow>  | 
| 31719 | 372  | 
a < m \<Longrightarrow> 0 \<le> b \<Longrightarrow> b < m \<Longrightarrow> [a = b] (mod m) \<Longrightarrow> a = b"  | 
| 41541 | 373  | 
by (auto simp add: cong_int_def)  | 
| 31719 | 374  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
375  | 
lemma cong_less_unique_nat:  | 
| 31719 | 376  | 
"0 < (m::nat) \<Longrightarrow> (\<exists>!b. 0 \<le> b \<and> b < m \<and> [a = b] (mod m))"  | 
377  | 
apply auto  | 
|
378  | 
apply (rule_tac x = "a mod m" in exI)  | 
|
379  | 
apply (unfold cong_nat_def, auto)  | 
|
| 44872 | 380  | 
done  | 
| 31719 | 381  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
382  | 
lemma cong_less_unique_int:  | 
| 31719 | 383  | 
"0 < (m::int) \<Longrightarrow> (\<exists>!b. 0 \<le> b \<and> b < m \<and> [a = b] (mod m))"  | 
384  | 
apply auto  | 
|
385  | 
apply (rule_tac x = "a mod m" in exI)  | 
|
| 41541 | 386  | 
apply (unfold cong_int_def, auto)  | 
387  | 
done  | 
|
| 31719 | 388  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
389  | 
lemma cong_iff_lin_int: "([(a::int) = b] (mod m)) = (\<exists>k. b = a + m * k)"  | 
| 36350 | 390  | 
apply (auto simp add: cong_altdef_int dvd_def field_simps)  | 
| 31719 | 391  | 
apply (rule_tac [!] x = "-k" in exI, auto)  | 
| 44872 | 392  | 
done  | 
| 31719 | 393  | 
|
| 44872 | 394  | 
lemma cong_iff_lin_nat: "([(a::nat) = b] (mod m)) =  | 
| 31719 | 395  | 
(\<exists>k1 k2. b + k1 * m = a + k2 * m)"  | 
396  | 
apply (rule iffI)  | 
|
| 44872 | 397  | 
apply (cases "b <= a")  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
398  | 
apply (subst (asm) cong_altdef_nat, assumption)  | 
| 31719 | 399  | 
apply (unfold dvd_def, auto)  | 
400  | 
apply (rule_tac x = k in exI)  | 
|
401  | 
apply (rule_tac x = 0 in exI)  | 
|
| 36350 | 402  | 
apply (auto simp add: field_simps)  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
403  | 
apply (subst (asm) cong_sym_eq_nat)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
404  | 
apply (subst (asm) cong_altdef_nat)  | 
| 31719 | 405  | 
apply force  | 
406  | 
apply (unfold dvd_def, auto)  | 
|
407  | 
apply (rule_tac x = 0 in exI)  | 
|
408  | 
apply (rule_tac x = k in exI)  | 
|
| 36350 | 409  | 
apply (auto simp add: field_simps)  | 
| 31719 | 410  | 
apply (unfold cong_nat_def)  | 
411  | 
apply (subgoal_tac "a mod m = (a + k2 * m) mod m")  | 
|
412  | 
apply (erule ssubst)back  | 
|
413  | 
apply (erule subst)  | 
|
414  | 
apply auto  | 
|
| 44872 | 415  | 
done  | 
| 31719 | 416  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
417  | 
lemma cong_gcd_eq_int: "[(a::int) = b] (mod m) \<Longrightarrow> gcd a m = gcd b m"  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
418  | 
apply (subst (asm) cong_iff_lin_int, auto)  | 
| 44872 | 419  | 
apply (subst add_commute)  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
420  | 
apply (subst (2) gcd_commute_int)  | 
| 31719 | 421  | 
apply (subst mult_commute)  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
422  | 
apply (subst gcd_add_mult_int)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
423  | 
apply (rule gcd_commute_int)  | 
| 41541 | 424  | 
done  | 
| 31719 | 425  | 
|
| 44872 | 426  | 
lemma cong_gcd_eq_nat:  | 
| 31719 | 427  | 
assumes "[(a::nat) = b] (mod m)"  | 
428  | 
shows "gcd a m = gcd b m"  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
429  | 
apply (rule cong_gcd_eq_int [transferred])  | 
| 41541 | 430  | 
using assms apply auto  | 
431  | 
done  | 
|
| 31719 | 432  | 
|
| 44872 | 433  | 
lemma cong_imp_coprime_nat: "[(a::nat) = b] (mod m) \<Longrightarrow> coprime a m \<Longrightarrow> coprime b m"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
434  | 
by (auto simp add: cong_gcd_eq_nat)  | 
| 31719 | 435  | 
|
| 44872 | 436  | 
lemma cong_imp_coprime_int: "[(a::int) = b] (mod m) \<Longrightarrow> coprime a m \<Longrightarrow> coprime b m"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
437  | 
by (auto simp add: cong_gcd_eq_int)  | 
| 31719 | 438  | 
|
| 44872 | 439  | 
lemma cong_cong_mod_nat: "[(a::nat) = b] (mod m) = [a mod m = b mod m] (mod m)"  | 
| 31719 | 440  | 
by (auto simp add: cong_nat_def)  | 
441  | 
||
| 44872 | 442  | 
lemma cong_cong_mod_int: "[(a::int) = b] (mod m) = [a mod m = b mod m] (mod m)"  | 
| 31719 | 443  | 
by (auto simp add: cong_int_def)  | 
444  | 
||
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
445  | 
lemma cong_minus_int [iff]: "[(a::int) = b] (mod -m) = [a = b] (mod m)"  | 
| 44872 | 446  | 
apply (subst (1 2) cong_altdef_int)  | 
447  | 
apply auto  | 
|
448  | 
done  | 
|
| 31719 | 449  | 
|
| 41541 | 450  | 
lemma cong_zero_nat: "[(a::nat) = b] (mod 0) = (a = b)"  | 
451  | 
by auto  | 
|
| 31719 | 452  | 
|
| 41541 | 453  | 
lemma cong_zero_int: "[(a::int) = b] (mod 0) = (a = b)"  | 
454  | 
by auto  | 
|
| 31719 | 455  | 
|
456  | 
(*  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
457  | 
lemma mod_dvd_mod_int:  | 
| 31719 | 458  | 
"0 < (m::int) \<Longrightarrow> m dvd b \<Longrightarrow> (a mod b mod m) = (a mod m)"  | 
459  | 
apply (unfold dvd_def, auto)  | 
|
460  | 
apply (rule mod_mod_cancel)  | 
|
461  | 
apply auto  | 
|
| 44872 | 462  | 
done  | 
| 31719 | 463  | 
|
464  | 
lemma mod_dvd_mod:  | 
|
465  | 
assumes "0 < (m::nat)" and "m dvd b"  | 
|
466  | 
shows "(a mod b mod m) = (a mod m)"  | 
|
467  | 
||
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
468  | 
apply (rule mod_dvd_mod_int [transferred])  | 
| 41541 | 469  | 
using assms apply auto  | 
470  | 
done  | 
|
| 31719 | 471  | 
*)  | 
472  | 
||
| 44872 | 473  | 
lemma cong_add_lcancel_nat:  | 
474  | 
"[(a::nat) + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)"  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
475  | 
by (simp add: cong_iff_lin_nat)  | 
| 31719 | 476  | 
|
| 44872 | 477  | 
lemma cong_add_lcancel_int:  | 
478  | 
"[(a::int) + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)"  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
479  | 
by (simp add: cong_iff_lin_int)  | 
| 31719 | 480  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
481  | 
lemma cong_add_rcancel_nat: "[(x::nat) + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)"  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
482  | 
by (simp add: cong_iff_lin_nat)  | 
| 31719 | 483  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
484  | 
lemma cong_add_rcancel_int: "[(x::int) + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)"  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
485  | 
by (simp add: cong_iff_lin_int)  | 
| 31719 | 486  | 
|
| 44872 | 487  | 
lemma cong_add_lcancel_0_nat: "[(a::nat) + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
488  | 
by (simp add: cong_iff_lin_nat)  | 
| 31719 | 489  | 
|
| 44872 | 490  | 
lemma cong_add_lcancel_0_int: "[(a::int) + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
491  | 
by (simp add: cong_iff_lin_int)  | 
| 31719 | 492  | 
|
| 44872 | 493  | 
lemma cong_add_rcancel_0_nat: "[x + (a::nat) = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
494  | 
by (simp add: cong_iff_lin_nat)  | 
| 31719 | 495  | 
|
| 44872 | 496  | 
lemma cong_add_rcancel_0_int: "[x + (a::int) = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
497  | 
by (simp add: cong_iff_lin_int)  | 
| 31719 | 498  | 
|
| 44872 | 499  | 
lemma cong_dvd_modulus_nat: "[(x::nat) = y] (mod m) \<Longrightarrow> n dvd m \<Longrightarrow>  | 
| 31719 | 500  | 
[x = y] (mod n)"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
501  | 
apply (auto simp add: cong_iff_lin_nat dvd_def)  | 
| 31719 | 502  | 
apply (rule_tac x="k1 * k" in exI)  | 
503  | 
apply (rule_tac x="k2 * k" in exI)  | 
|
| 36350 | 504  | 
apply (simp add: field_simps)  | 
| 44872 | 505  | 
done  | 
| 31719 | 506  | 
|
| 44872 | 507  | 
lemma cong_dvd_modulus_int: "[(x::int) = y] (mod m) \<Longrightarrow> n dvd m \<Longrightarrow> [x = y] (mod n)"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
508  | 
by (auto simp add: cong_altdef_int dvd_def)  | 
| 31719 | 509  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
510  | 
lemma cong_dvd_eq_nat: "[(x::nat) = y] (mod n) \<Longrightarrow> n dvd x \<longleftrightarrow> n dvd y"  | 
| 44872 | 511  | 
unfolding cong_nat_def by (auto simp add: dvd_eq_mod_eq_0)  | 
| 31719 | 512  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
513  | 
lemma cong_dvd_eq_int: "[(x::int) = y] (mod n) \<Longrightarrow> n dvd x \<longleftrightarrow> n dvd y"  | 
| 44872 | 514  | 
unfolding cong_int_def by (auto simp add: dvd_eq_mod_eq_0)  | 
| 31719 | 515  | 
|
| 44872 | 516  | 
lemma cong_mod_nat: "(n::nat) ~= 0 \<Longrightarrow> [a mod n = a] (mod n)"  | 
| 31719 | 517  | 
by (simp add: cong_nat_def)  | 
518  | 
||
| 44872 | 519  | 
lemma cong_mod_int: "(n::int) ~= 0 \<Longrightarrow> [a mod n = a] (mod n)"  | 
| 31719 | 520  | 
by (simp add: cong_int_def)  | 
521  | 
||
| 44872 | 522  | 
lemma mod_mult_cong_nat: "(a::nat) ~= 0 \<Longrightarrow> b ~= 0  | 
| 31719 | 523  | 
\<Longrightarrow> [x mod (a * b) = y] (mod a) \<longleftrightarrow> [x = y] (mod a)"  | 
524  | 
by (simp add: cong_nat_def mod_mult2_eq mod_add_left_eq)  | 
|
525  | 
||
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
526  | 
lemma neg_cong_int: "([(a::int) = b] (mod m)) = ([-a = -b] (mod m))"  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
527  | 
apply (simp add: cong_altdef_int)  | 
| 31719 | 528  | 
apply (subst dvd_minus_iff [symmetric])  | 
| 36350 | 529  | 
apply (simp add: field_simps)  | 
| 44872 | 530  | 
done  | 
| 31719 | 531  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
532  | 
lemma cong_modulus_neg_int: "([(a::int) = b] (mod m)) = ([a = b] (mod -m))"  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
533  | 
by (auto simp add: cong_altdef_int)  | 
| 31719 | 534  | 
|
| 44872 | 535  | 
lemma mod_mult_cong_int: "(a::int) ~= 0 \<Longrightarrow> b ~= 0  | 
| 31719 | 536  | 
\<Longrightarrow> [x mod (a * b) = y] (mod a) \<longleftrightarrow> [x = y] (mod a)"  | 
| 44872 | 537  | 
apply (cases "b > 0")  | 
| 31719 | 538  | 
apply (simp add: cong_int_def mod_mod_cancel mod_add_left_eq)  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
539  | 
apply (subst (1 2) cong_modulus_neg_int)  | 
| 31719 | 540  | 
apply (unfold cong_int_def)  | 
541  | 
apply (subgoal_tac "a * b = (-a * -b)")  | 
|
542  | 
apply (erule ssubst)  | 
|
543  | 
apply (subst zmod_zmult2_eq)  | 
|
| 
54230
 
b1d955791529
more simplification rules on unary and binary minus
 
haftmann 
parents: 
47163 
diff
changeset
 | 
544  | 
apply (auto simp add: mod_add_left_eq mod_minus_right div_minus_right)  | 
| 
 
b1d955791529
more simplification rules on unary and binary minus
 
haftmann 
parents: 
47163 
diff
changeset
 | 
545  | 
apply (metis mod_diff_left_eq mod_diff_right_eq mod_mult_self1_is_0 semiring_numeral_div_class.diff_zero)+  | 
| 44872 | 546  | 
done  | 
| 31719 | 547  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
548  | 
lemma cong_to_1_nat: "([(a::nat) = 1] (mod n)) \<Longrightarrow> (n dvd (a - 1))"  | 
| 44872 | 549  | 
apply (cases "a = 0")  | 
| 31719 | 550  | 
apply force  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
551  | 
apply (subst (asm) cong_altdef_nat)  | 
| 31719 | 552  | 
apply auto  | 
| 44872 | 553  | 
done  | 
| 31719 | 554  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
555  | 
lemma cong_0_1_nat: "[(0::nat) = 1] (mod n) = (n = 1)"  | 
| 44872 | 556  | 
unfolding cong_nat_def by auto  | 
| 31719 | 557  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
558  | 
lemma cong_0_1_int: "[(0::int) = 1] (mod n) = ((n = 1) | (n = -1))"  | 
| 44872 | 559  | 
unfolding cong_int_def by (auto simp add: zmult_eq_1_iff)  | 
| 31719 | 560  | 
|
| 44872 | 561  | 
lemma cong_to_1'_nat: "[(a::nat) = 1] (mod n) \<longleftrightarrow>  | 
| 31719 | 562  | 
a = 0 \<and> n = 1 \<or> (\<exists>m. a = 1 + m * n)"  | 
| 44872 | 563  | 
apply (cases "n = 1")  | 
| 31719 | 564  | 
apply auto [1]  | 
565  | 
apply (drule_tac x = "a - 1" in spec)  | 
|
566  | 
apply force  | 
|
| 44872 | 567  | 
apply (cases "a = 0")  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
568  | 
apply (auto simp add: cong_0_1_nat) [1]  | 
| 31719 | 569  | 
apply (rule iffI)  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
570  | 
apply (drule cong_to_1_nat)  | 
| 31719 | 571  | 
apply (unfold dvd_def)  | 
572  | 
apply auto [1]  | 
|
573  | 
apply (rule_tac x = k in exI)  | 
|
| 36350 | 574  | 
apply (auto simp add: field_simps) [1]  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
575  | 
apply (subst cong_altdef_nat)  | 
| 31719 | 576  | 
apply (auto simp add: dvd_def)  | 
| 44872 | 577  | 
done  | 
| 31719 | 578  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
579  | 
lemma cong_le_nat: "(y::nat) <= x \<Longrightarrow> [x = y] (mod n) \<longleftrightarrow> (\<exists>q. x = q * n + y)"  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
580  | 
apply (subst cong_altdef_nat)  | 
| 31719 | 581  | 
apply assumption  | 
| 36350 | 582  | 
apply (unfold dvd_def, auto simp add: field_simps)  | 
| 31719 | 583  | 
apply (rule_tac x = k in exI)  | 
584  | 
apply auto  | 
|
| 44872 | 585  | 
done  | 
| 31719 | 586  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
587  | 
lemma cong_solve_nat: "(a::nat) \<noteq> 0 \<Longrightarrow> EX x. [a * x = gcd a n] (mod n)"  | 
| 44872 | 588  | 
apply (cases "n = 0")  | 
| 31719 | 589  | 
apply force  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
590  | 
apply (frule bezout_nat [of a n], auto)  | 
| 31719 | 591  | 
apply (rule exI, erule ssubst)  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
592  | 
apply (rule cong_trans_nat)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
593  | 
apply (rule cong_add_nat)  | 
| 31719 | 594  | 
apply (subst mult_commute)  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
595  | 
apply (rule cong_mult_self_nat)  | 
| 31719 | 596  | 
prefer 2  | 
597  | 
apply simp  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
598  | 
apply (rule cong_refl_nat)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
599  | 
apply (rule cong_refl_nat)  | 
| 44872 | 600  | 
done  | 
| 31719 | 601  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
602  | 
lemma cong_solve_int: "(a::int) \<noteq> 0 \<Longrightarrow> EX x. [a * x = gcd a n] (mod n)"  | 
| 44872 | 603  | 
apply (cases "n = 0")  | 
604  | 
apply (cases "a \<ge> 0")  | 
|
| 31719 | 605  | 
apply auto  | 
606  | 
apply (rule_tac x = "-1" in exI)  | 
|
607  | 
apply auto  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
608  | 
apply (insert bezout_int [of a n], auto)  | 
| 31719 | 609  | 
apply (rule exI)  | 
610  | 
apply (erule subst)  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
611  | 
apply (rule cong_trans_int)  | 
| 31719 | 612  | 
prefer 2  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
613  | 
apply (rule cong_add_int)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
614  | 
apply (rule cong_refl_int)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
615  | 
apply (rule cong_sym_int)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
616  | 
apply (rule cong_mult_self_int)  | 
| 31719 | 617  | 
apply simp  | 
618  | 
apply (subst mult_commute)  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
619  | 
apply (rule cong_refl_int)  | 
| 44872 | 620  | 
done  | 
621  | 
||
622  | 
lemma cong_solve_dvd_nat:  | 
|
| 31719 | 623  | 
assumes a: "(a::nat) \<noteq> 0" and b: "gcd a n dvd d"  | 
624  | 
shows "EX x. [a * x = d] (mod n)"  | 
|
625  | 
proof -  | 
|
| 44872 | 626  | 
from cong_solve_nat [OF a] obtain x where "[a * x = gcd a n](mod n)"  | 
| 31719 | 627  | 
by auto  | 
| 44872 | 628  | 
then have "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
629  | 
by (elim cong_scalar2_nat)  | 
| 31719 | 630  | 
also from b have "(d div gcd a n) * gcd a n = d"  | 
631  | 
by (rule dvd_div_mult_self)  | 
|
632  | 
also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)"  | 
|
633  | 
by auto  | 
|
634  | 
finally show ?thesis  | 
|
635  | 
by auto  | 
|
636  | 
qed  | 
|
637  | 
||
| 44872 | 638  | 
lemma cong_solve_dvd_int:  | 
| 31719 | 639  | 
assumes a: "(a::int) \<noteq> 0" and b: "gcd a n dvd d"  | 
640  | 
shows "EX x. [a * x = d] (mod n)"  | 
|
641  | 
proof -  | 
|
| 44872 | 642  | 
from cong_solve_int [OF a] obtain x where "[a * x = gcd a n](mod n)"  | 
| 31719 | 643  | 
by auto  | 
| 44872 | 644  | 
then have "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
645  | 
by (elim cong_scalar2_int)  | 
| 31719 | 646  | 
also from b have "(d div gcd a n) * gcd a n = d"  | 
647  | 
by (rule dvd_div_mult_self)  | 
|
648  | 
also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)"  | 
|
649  | 
by auto  | 
|
650  | 
finally show ?thesis  | 
|
651  | 
by auto  | 
|
652  | 
qed  | 
|
653  | 
||
| 44872 | 654  | 
lemma cong_solve_coprime_nat: "coprime (a::nat) n \<Longrightarrow> EX x. [a * x = 1] (mod n)"  | 
655  | 
apply (cases "a = 0")  | 
|
| 31719 | 656  | 
apply force  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
657  | 
apply (frule cong_solve_nat [of a n])  | 
| 31719 | 658  | 
apply auto  | 
| 44872 | 659  | 
done  | 
| 31719 | 660  | 
|
| 44872 | 661  | 
lemma cong_solve_coprime_int: "coprime (a::int) n \<Longrightarrow> EX x. [a * x = 1] (mod n)"  | 
662  | 
apply (cases "a = 0")  | 
|
| 31719 | 663  | 
apply auto  | 
| 44872 | 664  | 
apply (cases "n \<ge> 0")  | 
| 31719 | 665  | 
apply auto  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
666  | 
apply (frule cong_solve_int [of a n])  | 
| 31719 | 667  | 
apply auto  | 
| 44872 | 668  | 
done  | 
| 31719 | 669  | 
|
| 44872 | 670  | 
lemma coprime_iff_invertible_nat: "m > (1::nat) \<Longrightarrow> coprime a m = (EX x. [a * x = 1] (mod m))"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
671  | 
apply (auto intro: cong_solve_coprime_nat)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
672  | 
apply (unfold cong_nat_def, auto intro: invertible_coprime_nat)  | 
| 44872 | 673  | 
done  | 
| 31719 | 674  | 
|
| 44872 | 675  | 
lemma coprime_iff_invertible_int: "m > (1::int) \<Longrightarrow> coprime a m = (EX x. [a * x = 1] (mod m))"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
676  | 
apply (auto intro: cong_solve_coprime_int)  | 
| 31719 | 677  | 
apply (unfold cong_int_def)  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
678  | 
apply (auto intro: invertible_coprime_int)  | 
| 44872 | 679  | 
done  | 
| 31719 | 680  | 
|
| 44872 | 681  | 
lemma coprime_iff_invertible'_int: "m > (1::int) \<Longrightarrow> coprime a m =  | 
| 31719 | 682  | 
(EX x. 0 <= x & x < m & [a * x = 1] (mod m))"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
683  | 
apply (subst coprime_iff_invertible_int)  | 
| 31719 | 684  | 
apply auto  | 
685  | 
apply (auto simp add: cong_int_def)  | 
|
686  | 
apply (rule_tac x = "x mod m" in exI)  | 
|
687  | 
apply (auto simp add: mod_mult_right_eq [symmetric])  | 
|
| 44872 | 688  | 
done  | 
| 31719 | 689  | 
|
690  | 
||
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
691  | 
lemma cong_cong_lcm_nat: "[(x::nat) = y] (mod a) \<Longrightarrow>  | 
| 31719 | 692  | 
[x = y] (mod b) \<Longrightarrow> [x = y] (mod lcm a b)"  | 
| 44872 | 693  | 
apply (cases "y \<le> x")  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
694  | 
apply (auto simp add: cong_altdef_nat lcm_least_nat) [1]  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
695  | 
apply (rule cong_sym_nat)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
696  | 
apply (subst (asm) (1 2) cong_sym_eq_nat)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
697  | 
apply (auto simp add: cong_altdef_nat lcm_least_nat)  | 
| 44872 | 698  | 
done  | 
| 31719 | 699  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
700  | 
lemma cong_cong_lcm_int: "[(x::int) = y] (mod a) \<Longrightarrow>  | 
| 31719 | 701  | 
[x = y] (mod b) \<Longrightarrow> [x = y] (mod lcm a b)"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
702  | 
by (auto simp add: cong_altdef_int lcm_least_int) [1]  | 
| 31719 | 703  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
704  | 
lemma cong_cong_coprime_nat: "coprime a b \<Longrightarrow> [(x::nat) = y] (mod a) \<Longrightarrow>  | 
| 31719 | 705  | 
[x = y] (mod b) \<Longrightarrow> [x = y] (mod a * b)"  | 
| 44872 | 706  | 
apply (frule (1) cong_cong_lcm_nat)  | 
707  | 
back  | 
|
| 31719 | 708  | 
apply (simp add: lcm_nat_def)  | 
| 44872 | 709  | 
done  | 
| 31719 | 710  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
711  | 
lemma cong_cong_coprime_int: "coprime a b \<Longrightarrow> [(x::int) = y] (mod a) \<Longrightarrow>  | 
| 31719 | 712  | 
[x = y] (mod b) \<Longrightarrow> [x = y] (mod a * b)"  | 
| 44872 | 713  | 
apply (frule (1) cong_cong_lcm_int)  | 
714  | 
back  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
715  | 
apply (simp add: lcm_altdef_int cong_abs_int abs_mult [symmetric])  | 
| 44872 | 716  | 
done  | 
| 31719 | 717  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
718  | 
lemma cong_cong_setprod_coprime_nat [rule_format]: "finite A \<Longrightarrow>  | 
| 31719 | 719  | 
(ALL i:A. (ALL j:A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))) \<longrightarrow>  | 
720  | 
(ALL i:A. [(x::nat) = y] (mod m i)) \<longrightarrow>  | 
|
721  | 
[x = y] (mod (PROD i:A. m i))"  | 
|
722  | 
apply (induct set: finite)  | 
|
723  | 
apply auto  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
724  | 
apply (rule cong_cong_coprime_nat)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
725  | 
apply (subst gcd_commute_nat)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
726  | 
apply (rule setprod_coprime_nat)  | 
| 31719 | 727  | 
apply auto  | 
| 44872 | 728  | 
done  | 
| 31719 | 729  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
730  | 
lemma cong_cong_setprod_coprime_int [rule_format]: "finite A \<Longrightarrow>  | 
| 31719 | 731  | 
(ALL i:A. (ALL j:A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))) \<longrightarrow>  | 
732  | 
(ALL i:A. [(x::int) = y] (mod m i)) \<longrightarrow>  | 
|
733  | 
[x = y] (mod (PROD i:A. m i))"  | 
|
734  | 
apply (induct set: finite)  | 
|
735  | 
apply auto  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
736  | 
apply (rule cong_cong_coprime_int)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
737  | 
apply (subst gcd_commute_int)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
738  | 
apply (rule setprod_coprime_int)  | 
| 31719 | 739  | 
apply auto  | 
| 44872 | 740  | 
done  | 
| 31719 | 741  | 
|
| 44872 | 742  | 
lemma binary_chinese_remainder_aux_nat:  | 
| 31719 | 743  | 
assumes a: "coprime (m1::nat) m2"  | 
744  | 
shows "EX b1 b2. [b1 = 1] (mod m1) \<and> [b1 = 0] (mod m2) \<and>  | 
|
745  | 
[b2 = 0] (mod m1) \<and> [b2 = 1] (mod m2)"  | 
|
746  | 
proof -  | 
|
| 44872 | 747  | 
from cong_solve_coprime_nat [OF a] obtain x1 where one: "[m1 * x1 = 1] (mod m2)"  | 
| 31719 | 748  | 
by auto  | 
| 44872 | 749  | 
from a have b: "coprime m2 m1"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
750  | 
by (subst gcd_commute_nat)  | 
| 44872 | 751  | 
from cong_solve_coprime_nat [OF b] obtain x2 where two: "[m2 * x2 = 1] (mod m1)"  | 
| 31719 | 752  | 
by auto  | 
753  | 
have "[m1 * x1 = 0] (mod m1)"  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
754  | 
by (subst mult_commute, rule cong_mult_self_nat)  | 
| 31719 | 755  | 
moreover have "[m2 * x2 = 0] (mod m2)"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
756  | 
by (subst mult_commute, rule cong_mult_self_nat)  | 
| 31719 | 757  | 
moreover note one two  | 
758  | 
ultimately show ?thesis by blast  | 
|
759  | 
qed  | 
|
760  | 
||
| 44872 | 761  | 
lemma binary_chinese_remainder_aux_int:  | 
| 31719 | 762  | 
assumes a: "coprime (m1::int) m2"  | 
763  | 
shows "EX b1 b2. [b1 = 1] (mod m1) \<and> [b1 = 0] (mod m2) \<and>  | 
|
764  | 
[b2 = 0] (mod m1) \<and> [b2 = 1] (mod m2)"  | 
|
765  | 
proof -  | 
|
| 44872 | 766  | 
from cong_solve_coprime_int [OF a] obtain x1 where one: "[m1 * x1 = 1] (mod m2)"  | 
| 31719 | 767  | 
by auto  | 
| 44872 | 768  | 
from a have b: "coprime m2 m1"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
769  | 
by (subst gcd_commute_int)  | 
| 44872 | 770  | 
from cong_solve_coprime_int [OF b] obtain x2 where two: "[m2 * x2 = 1] (mod m1)"  | 
| 31719 | 771  | 
by auto  | 
772  | 
have "[m1 * x1 = 0] (mod m1)"  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
773  | 
by (subst mult_commute, rule cong_mult_self_int)  | 
| 31719 | 774  | 
moreover have "[m2 * x2 = 0] (mod m2)"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
775  | 
by (subst mult_commute, rule cong_mult_self_int)  | 
| 31719 | 776  | 
moreover note one two  | 
777  | 
ultimately show ?thesis by blast  | 
|
778  | 
qed  | 
|
779  | 
||
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
780  | 
lemma binary_chinese_remainder_nat:  | 
| 31719 | 781  | 
assumes a: "coprime (m1::nat) m2"  | 
782  | 
shows "EX x. [x = u1] (mod m1) \<and> [x = u2] (mod m2)"  | 
|
783  | 
proof -  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
784  | 
from binary_chinese_remainder_aux_nat [OF a] obtain b1 b2  | 
| 44872 | 785  | 
where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)" and  | 
786  | 
"[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)"  | 
|
| 31719 | 787  | 
by blast  | 
788  | 
let ?x = "u1 * b1 + u2 * b2"  | 
|
789  | 
have "[?x = u1 * 1 + u2 * 0] (mod m1)"  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
790  | 
apply (rule cong_add_nat)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
791  | 
apply (rule cong_scalar2_nat)  | 
| 31719 | 792  | 
apply (rule `[b1 = 1] (mod m1)`)  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
793  | 
apply (rule cong_scalar2_nat)  | 
| 31719 | 794  | 
apply (rule `[b2 = 0] (mod m1)`)  | 
795  | 
done  | 
|
| 44872 | 796  | 
then have "[?x = u1] (mod m1)" by simp  | 
| 31719 | 797  | 
have "[?x = u1 * 0 + u2 * 1] (mod m2)"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
798  | 
apply (rule cong_add_nat)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
799  | 
apply (rule cong_scalar2_nat)  | 
| 31719 | 800  | 
apply (rule `[b1 = 0] (mod m2)`)  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
801  | 
apply (rule cong_scalar2_nat)  | 
| 31719 | 802  | 
apply (rule `[b2 = 1] (mod m2)`)  | 
803  | 
done  | 
|
| 44872 | 804  | 
then have "[?x = u2] (mod m2)" by simp  | 
| 31719 | 805  | 
with `[?x = u1] (mod m1)` show ?thesis by blast  | 
806  | 
qed  | 
|
807  | 
||
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
808  | 
lemma binary_chinese_remainder_int:  | 
| 31719 | 809  | 
assumes a: "coprime (m1::int) m2"  | 
810  | 
shows "EX x. [x = u1] (mod m1) \<and> [x = u2] (mod m2)"  | 
|
811  | 
proof -  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
812  | 
from binary_chinese_remainder_aux_int [OF a] obtain b1 b2  | 
| 31719 | 813  | 
where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)" and  | 
814  | 
"[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)"  | 
|
815  | 
by blast  | 
|
816  | 
let ?x = "u1 * b1 + u2 * b2"  | 
|
817  | 
have "[?x = u1 * 1 + u2 * 0] (mod m1)"  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
818  | 
apply (rule cong_add_int)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
819  | 
apply (rule cong_scalar2_int)  | 
| 31719 | 820  | 
apply (rule `[b1 = 1] (mod m1)`)  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
821  | 
apply (rule cong_scalar2_int)  | 
| 31719 | 822  | 
apply (rule `[b2 = 0] (mod m1)`)  | 
823  | 
done  | 
|
| 44872 | 824  | 
then have "[?x = u1] (mod m1)" by simp  | 
| 31719 | 825  | 
have "[?x = u1 * 0 + u2 * 1] (mod m2)"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
826  | 
apply (rule cong_add_int)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
827  | 
apply (rule cong_scalar2_int)  | 
| 31719 | 828  | 
apply (rule `[b1 = 0] (mod m2)`)  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
829  | 
apply (rule cong_scalar2_int)  | 
| 31719 | 830  | 
apply (rule `[b2 = 1] (mod m2)`)  | 
831  | 
done  | 
|
| 44872 | 832  | 
then have "[?x = u2] (mod m2)" by simp  | 
| 31719 | 833  | 
with `[?x = u1] (mod m1)` show ?thesis by blast  | 
834  | 
qed  | 
|
835  | 
||
| 44872 | 836  | 
lemma cong_modulus_mult_nat: "[(x::nat) = y] (mod m * n) \<Longrightarrow>  | 
| 31719 | 837  | 
[x = y] (mod m)"  | 
| 44872 | 838  | 
apply (cases "y \<le> x")  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
839  | 
apply (simp add: cong_altdef_nat)  | 
| 31719 | 840  | 
apply (erule dvd_mult_left)  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
841  | 
apply (rule cong_sym_nat)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
842  | 
apply (subst (asm) cong_sym_eq_nat)  | 
| 44872 | 843  | 
apply (simp add: cong_altdef_nat)  | 
| 31719 | 844  | 
apply (erule dvd_mult_left)  | 
| 44872 | 845  | 
done  | 
| 31719 | 846  | 
|
| 44872 | 847  | 
lemma cong_modulus_mult_int: "[(x::int) = y] (mod m * n) \<Longrightarrow>  | 
| 31719 | 848  | 
[x = y] (mod m)"  | 
| 44872 | 849  | 
apply (simp add: cong_altdef_int)  | 
| 31719 | 850  | 
apply (erule dvd_mult_left)  | 
| 44872 | 851  | 
done  | 
| 31719 | 852  | 
|
| 44872 | 853  | 
lemma cong_less_modulus_unique_nat:  | 
| 31719 | 854  | 
"[(x::nat) = y] (mod m) \<Longrightarrow> x < m \<Longrightarrow> y < m \<Longrightarrow> x = y"  | 
855  | 
by (simp add: cong_nat_def)  | 
|
856  | 
||
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
857  | 
lemma binary_chinese_remainder_unique_nat:  | 
| 44872 | 858  | 
assumes a: "coprime (m1::nat) m2"  | 
859  | 
and nz: "m1 \<noteq> 0" "m2 \<noteq> 0"  | 
|
| 31719 | 860  | 
shows "EX! x. x < m1 * m2 \<and> [x = u1] (mod m1) \<and> [x = u2] (mod m2)"  | 
861  | 
proof -  | 
|
| 44872 | 862  | 
from binary_chinese_remainder_nat [OF a] obtain y where  | 
| 31719 | 863  | 
"[y = u1] (mod m1)" and "[y = u2] (mod m2)"  | 
864  | 
by blast  | 
|
865  | 
let ?x = "y mod (m1 * m2)"  | 
|
866  | 
from nz have less: "?x < m1 * m2"  | 
|
| 44872 | 867  | 
by auto  | 
| 31719 | 868  | 
have one: "[?x = u1] (mod m1)"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
869  | 
apply (rule cong_trans_nat)  | 
| 31719 | 870  | 
prefer 2  | 
871  | 
apply (rule `[y = u1] (mod m1)`)  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
872  | 
apply (rule cong_modulus_mult_nat)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
873  | 
apply (rule cong_mod_nat)  | 
| 31719 | 874  | 
using nz apply auto  | 
875  | 
done  | 
|
876  | 
have two: "[?x = u2] (mod m2)"  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
877  | 
apply (rule cong_trans_nat)  | 
| 31719 | 878  | 
prefer 2  | 
879  | 
apply (rule `[y = u2] (mod m2)`)  | 
|
880  | 
apply (subst mult_commute)  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
881  | 
apply (rule cong_modulus_mult_nat)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
882  | 
apply (rule cong_mod_nat)  | 
| 31719 | 883  | 
using nz apply auto  | 
884  | 
done  | 
|
| 44872 | 885  | 
have "ALL z. z < m1 * m2 \<and> [z = u1] (mod m1) \<and> [z = u2] (mod m2) \<longrightarrow> z = ?x"  | 
886  | 
proof clarify  | 
|
| 31719 | 887  | 
fix z  | 
888  | 
assume "z < m1 * m2"  | 
|
889  | 
assume "[z = u1] (mod m1)" and "[z = u2] (mod m2)"  | 
|
890  | 
have "[?x = z] (mod m1)"  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
891  | 
apply (rule cong_trans_nat)  | 
| 31719 | 892  | 
apply (rule `[?x = u1] (mod m1)`)  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
893  | 
apply (rule cong_sym_nat)  | 
| 31719 | 894  | 
apply (rule `[z = u1] (mod m1)`)  | 
895  | 
done  | 
|
896  | 
moreover have "[?x = z] (mod m2)"  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
897  | 
apply (rule cong_trans_nat)  | 
| 31719 | 898  | 
apply (rule `[?x = u2] (mod m2)`)  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
899  | 
apply (rule cong_sym_nat)  | 
| 31719 | 900  | 
apply (rule `[z = u2] (mod m2)`)  | 
901  | 
done  | 
|
902  | 
ultimately have "[?x = z] (mod m1 * m2)"  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
903  | 
by (auto intro: coprime_cong_mult_nat a)  | 
| 31719 | 904  | 
with `z < m1 * m2` `?x < m1 * m2` show "z = ?x"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
905  | 
apply (intro cong_less_modulus_unique_nat)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
906  | 
apply (auto, erule cong_sym_nat)  | 
| 31719 | 907  | 
done  | 
| 44872 | 908  | 
qed  | 
909  | 
with less one two show ?thesis by auto  | 
|
| 31719 | 910  | 
qed  | 
911  | 
||
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
912  | 
lemma chinese_remainder_aux_nat:  | 
| 44872 | 913  | 
fixes A :: "'a set"  | 
914  | 
and m :: "'a \<Rightarrow> nat"  | 
|
915  | 
assumes fin: "finite A"  | 
|
916  | 
and cop: "ALL i : A. (ALL j : A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))"  | 
|
917  | 
  shows "EX b. (ALL i : A. [b i = 1] (mod m i) \<and> [b i = 0] (mod (PROD j : A - {i}. m j)))"
 | 
|
| 31719 | 918  | 
proof (rule finite_set_choice, rule fin, rule ballI)  | 
919  | 
fix i  | 
|
920  | 
assume "i : A"  | 
|
921  | 
  with cop have "coprime (PROD j : A - {i}. m j) (m i)"
 | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
922  | 
by (intro setprod_coprime_nat, auto)  | 
| 44872 | 923  | 
  then have "EX x. [(PROD j : 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
 | 
924  | 
by (elim cong_solve_coprime_nat)  | 
| 31719 | 925  | 
  then obtain x where "[(PROD j : A - {i}. m j) * x = 1] (mod m i)"
 | 
926  | 
by auto  | 
|
| 44872 | 927  | 
  moreover have "[(PROD j : A - {i}. m j) * x = 0]
 | 
| 31719 | 928  | 
    (mod (PROD j : A - {i}. m j))"
 | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
929  | 
by (subst mult_commute, rule cong_mult_self_nat)  | 
| 44872 | 930  | 
ultimately show "\<exists>a. [a = 1] (mod m i) \<and> [a = 0]  | 
| 31719 | 931  | 
      (mod setprod m (A - {i}))"
 | 
932  | 
by blast  | 
|
933  | 
qed  | 
|
934  | 
||
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
935  | 
lemma chinese_remainder_nat:  | 
| 44872 | 936  | 
fixes A :: "'a set"  | 
937  | 
and m :: "'a \<Rightarrow> nat"  | 
|
938  | 
and u :: "'a \<Rightarrow> nat"  | 
|
939  | 
assumes fin: "finite A"  | 
|
940  | 
and cop: "ALL i:A. (ALL j : A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))"  | 
|
| 31719 | 941  | 
shows "EX x. (ALL i:A. [x = u i] (mod m i))"  | 
942  | 
proof -  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
943  | 
from chinese_remainder_aux_nat [OF fin cop] obtain b where  | 
| 44872 | 944  | 
bprop: "ALL i:A. [b i = 1] (mod m i) \<and>  | 
| 31719 | 945  | 
      [b i = 0] (mod (PROD j : A - {i}. m j))"
 | 
946  | 
by blast  | 
|
947  | 
let ?x = "SUM i:A. (u i) * (b i)"  | 
|
948  | 
show "?thesis"  | 
|
949  | 
proof (rule exI, clarify)  | 
|
950  | 
fix i  | 
|
951  | 
assume a: "i : A"  | 
|
| 44872 | 952  | 
show "[?x = u i] (mod m i)"  | 
| 31719 | 953  | 
proof -  | 
| 44872 | 954  | 
      from fin a have "?x = (SUM j:{i}. u j * b j) +
 | 
| 31719 | 955  | 
          (SUM j:A-{i}. u j * b j)"
 | 
956  | 
by (subst setsum_Un_disjoint [symmetric], auto intro: setsum_cong)  | 
|
| 44872 | 957  | 
      then have "[?x = u i * b i + (SUM j:A-{i}. u j * b j)] (mod m i)"
 | 
| 31719 | 958  | 
by auto  | 
959  | 
      also have "[u i * b i + (SUM j:A-{i}. u j * b j) =
 | 
|
960  | 
                  u i * 1 + (SUM j:A-{i}. u j * 0)] (mod m i)"
 | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
961  | 
apply (rule cong_add_nat)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
962  | 
apply (rule cong_scalar2_nat)  | 
| 31719 | 963  | 
using bprop a apply blast  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
964  | 
apply (rule cong_setsum_nat)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
965  | 
apply (rule cong_scalar2_nat)  | 
| 31719 | 966  | 
using bprop apply auto  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
967  | 
apply (rule cong_dvd_modulus_nat)  | 
| 31719 | 968  | 
apply (drule (1) bspec)  | 
969  | 
apply (erule conjE)  | 
|
970  | 
apply assumption  | 
|
971  | 
apply (rule dvd_setprod)  | 
|
972  | 
using fin a apply auto  | 
|
973  | 
done  | 
|
974  | 
finally show ?thesis  | 
|
975  | 
by simp  | 
|
976  | 
qed  | 
|
977  | 
qed  | 
|
978  | 
qed  | 
|
979  | 
||
| 44872 | 980  | 
lemma coprime_cong_prod_nat [rule_format]: "finite A \<Longrightarrow>  | 
| 31719 | 981  | 
(ALL i: A. (ALL j: A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))) \<longrightarrow>  | 
982  | 
(ALL i: A. [(x::nat) = y] (mod m i)) \<longrightarrow>  | 
|
| 44872 | 983  | 
[x = y] (mod (PROD i:A. m i))"  | 
| 31719 | 984  | 
apply (induct set: finite)  | 
985  | 
apply auto  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
986  | 
apply (erule (1) coprime_cong_mult_nat)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
987  | 
apply (subst gcd_commute_nat)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
988  | 
apply (rule setprod_coprime_nat)  | 
| 31719 | 989  | 
apply auto  | 
| 44872 | 990  | 
done  | 
| 31719 | 991  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
992  | 
lemma chinese_remainder_unique_nat:  | 
| 44872 | 993  | 
fixes A :: "'a set"  | 
994  | 
and m :: "'a \<Rightarrow> nat"  | 
|
995  | 
and u :: "'a \<Rightarrow> nat"  | 
|
996  | 
assumes fin: "finite A"  | 
|
997  | 
and nz: "ALL i:A. m i \<noteq> 0"  | 
|
998  | 
and cop: "ALL i:A. (ALL j : A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))"  | 
|
| 31719 | 999  | 
shows "EX! x. x < (PROD i:A. m i) \<and> (ALL i:A. [x = u i] (mod m i))"  | 
1000  | 
proof -  | 
|
| 44872 | 1001  | 
from chinese_remainder_nat [OF fin cop]  | 
1002  | 
obtain y where one: "(ALL i:A. [y = u i] (mod m i))"  | 
|
| 31719 | 1003  | 
by blast  | 
1004  | 
let ?x = "y mod (PROD i:A. m i)"  | 
|
1005  | 
from fin nz have prodnz: "(PROD i:A. m i) \<noteq> 0"  | 
|
1006  | 
by auto  | 
|
| 44872 | 1007  | 
then have less: "?x < (PROD i:A. m i)"  | 
| 31719 | 1008  | 
by auto  | 
1009  | 
have cong: "ALL i:A. [?x = u i] (mod m i)"  | 
|
1010  | 
apply auto  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
1011  | 
apply (rule cong_trans_nat)  | 
| 31719 | 1012  | 
prefer 2  | 
1013  | 
using one apply auto  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
1014  | 
apply (rule cong_dvd_modulus_nat)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
1015  | 
apply (rule cong_mod_nat)  | 
| 31719 | 1016  | 
using prodnz apply auto  | 
1017  | 
apply (rule dvd_setprod)  | 
|
1018  | 
apply (rule fin)  | 
|
1019  | 
apply assumption  | 
|
1020  | 
done  | 
|
| 44872 | 1021  | 
have unique: "ALL z. z < (PROD i:A. m i) \<and>  | 
| 31719 | 1022  | 
(ALL i:A. [z = u i] (mod m i)) \<longrightarrow> z = ?x"  | 
1023  | 
proof (clarify)  | 
|
1024  | 
fix z  | 
|
1025  | 
assume zless: "z < (PROD i:A. m i)"  | 
|
1026  | 
assume zcong: "(ALL i:A. [z = u i] (mod m i))"  | 
|
1027  | 
have "ALL i:A. [?x = z] (mod m i)"  | 
|
| 44872 | 1028  | 
apply clarify  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
1029  | 
apply (rule cong_trans_nat)  | 
| 31719 | 1030  | 
using cong apply (erule bspec)  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
1031  | 
apply (rule cong_sym_nat)  | 
| 31719 | 1032  | 
using zcong apply auto  | 
1033  | 
done  | 
|
1034  | 
with fin cop have "[?x = z] (mod (PROD i:A. m i))"  | 
|
| 44872 | 1035  | 
apply (intro coprime_cong_prod_nat)  | 
1036  | 
apply auto  | 
|
1037  | 
done  | 
|
| 31719 | 1038  | 
with zless less show "z = ?x"  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
1039  | 
apply (intro cong_less_modulus_unique_nat)  | 
| 
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31792 
diff
changeset
 | 
1040  | 
apply (auto, erule cong_sym_nat)  | 
| 31719 | 1041  | 
done  | 
| 44872 | 1042  | 
qed  | 
1043  | 
from less cong unique show ?thesis by blast  | 
|
| 31719 | 1044  | 
qed  | 
1045  | 
||
1046  | 
end  |