author | wenzelm |
Wed, 04 Oct 2017 12:00:53 +0200 | |
changeset 66787 | 64b47495676d |
parent 66453 | cc19f7ca2ed6 |
child 66817 | 0b12755ccbb2 |
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 |
||
63167 | 35 |
subsection \<open>Turn off \<open>One_nat_def\<close>\<close> |
31719 | 36 |
|
66380 | 37 |
lemma power_eq_one_eq_nat [simp]: "x^m = 1 \<longleftrightarrow> m = 0 \<or> x = 1" |
38 |
for x m :: nat |
|
44872 | 39 |
by (induct m) auto |
31719 | 40 |
|
41 |
declare mod_pos_pos_trivial [simp] |
|
42 |
||
43 |
||
60526 | 44 |
subsection \<open>Main definitions\<close> |
31719 | 45 |
|
46 |
class cong = |
|
66380 | 47 |
fixes cong :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" ("(1[_ = _] '(()mod _'))") |
31719 | 48 |
begin |
49 |
||
58937 | 50 |
abbreviation notcong :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" ("(1[_ \<noteq> _] '(()mod _'))") |
44872 | 51 |
where "notcong x y m \<equiv> \<not> cong x y m" |
31719 | 52 |
|
53 |
end |
|
54 |
||
66380 | 55 |
|
56 |
subsubsection \<open>Definitions for the natural numbers\<close> |
|
31719 | 57 |
|
58 |
instantiation nat :: cong |
|
44872 | 59 |
begin |
31719 | 60 |
|
44872 | 61 |
definition cong_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" |
66380 | 62 |
where "cong_nat x y m \<longleftrightarrow> x mod m = y mod m" |
31719 | 63 |
|
44872 | 64 |
instance .. |
31719 | 65 |
|
66 |
end |
|
67 |
||
68 |
||
66380 | 69 |
subsubsection \<open>Definitions for the integers\<close> |
31719 | 70 |
|
71 |
instantiation int :: cong |
|
44872 | 72 |
begin |
31719 | 73 |
|
44872 | 74 |
definition cong_int :: "int \<Rightarrow> int \<Rightarrow> int \<Rightarrow> bool" |
66380 | 75 |
where "cong_int x y m \<longleftrightarrow> x mod m = y mod m" |
31719 | 76 |
|
44872 | 77 |
instance .. |
31719 | 78 |
|
79 |
end |
|
80 |
||
81 |
||
60526 | 82 |
subsection \<open>Set up Transfer\<close> |
31719 | 83 |
|
84 |
||
85 |
lemma transfer_nat_int_cong: |
|
66380 | 86 |
"x \<ge> 0 \<Longrightarrow> y \<ge> 0 \<Longrightarrow> m \<ge> 0 \<Longrightarrow> [nat x = nat y] (mod (nat m)) \<longleftrightarrow> [x = y] (mod m)" |
87 |
for x y m :: int |
|
44872 | 88 |
unfolding cong_int_def cong_nat_def |
55130
70db8d380d62
Restored Suc rather than +1, and using Library/Binimial
paulson <lp15@cam.ac.uk>
parents:
54489
diff
changeset
|
89 |
by (metis Divides.transfer_int_nat_functions(2) nat_0_le nat_mod_distrib) |
70db8d380d62
Restored Suc rather than +1, and using Library/Binimial
paulson <lp15@cam.ac.uk>
parents:
54489
diff
changeset
|
90 |
|
31719 | 91 |
|
66380 | 92 |
declare transfer_morphism_nat_int [transfer add return: transfer_nat_int_cong] |
31719 | 93 |
|
66380 | 94 |
lemma transfer_int_nat_cong: "[int x = int y] (mod (int m)) = [x = y] (mod m)" |
95 |
by (auto simp add: cong_int_def cong_nat_def) (auto simp add: zmod_int [symmetric]) |
|
31719 | 96 |
|
66380 | 97 |
declare transfer_morphism_int_nat [transfer add return: transfer_int_nat_cong] |
31719 | 98 |
|
99 |
||
60526 | 100 |
subsection \<open>Congruence\<close> |
31719 | 101 |
|
102 |
(* was zcong_0, etc. *) |
|
66380 | 103 |
lemma cong_0_nat [simp, presburger]: "[a = b] (mod 0) \<longleftrightarrow> a = b" |
104 |
for a b :: nat |
|
105 |
by (auto simp: cong_nat_def) |
|
31719 | 106 |
|
66380 | 107 |
lemma cong_0_int [simp, presburger]: "[a = b] (mod 0) \<longleftrightarrow> a = b" |
108 |
for a b :: int |
|
109 |
by (auto simp: cong_int_def) |
|
31719 | 110 |
|
66380 | 111 |
lemma cong_1_nat [simp, presburger]: "[a = b] (mod 1)" |
112 |
for a b :: nat |
|
113 |
by (auto simp: cong_nat_def) |
|
31719 | 114 |
|
66380 | 115 |
lemma cong_Suc_0_nat [simp, presburger]: "[a = b] (mod Suc 0)" |
116 |
for a b :: nat |
|
117 |
by (auto simp: cong_nat_def) |
|
31719 | 118 |
|
66380 | 119 |
lemma cong_1_int [simp, presburger]: "[a = b] (mod 1)" |
120 |
for a b :: int |
|
121 |
by (auto simp: cong_int_def) |
|
31719 | 122 |
|
66380 | 123 |
lemma cong_refl_nat [simp]: "[k = k] (mod m)" |
124 |
for k :: nat |
|
125 |
by (auto simp: cong_nat_def) |
|
31719 | 126 |
|
66380 | 127 |
lemma cong_refl_int [simp]: "[k = k] (mod m)" |
128 |
for k :: int |
|
129 |
by (auto simp: cong_int_def) |
|
31719 | 130 |
|
66380 | 131 |
lemma cong_sym_nat: "[a = b] (mod m) \<Longrightarrow> [b = a] (mod m)" |
132 |
for a b :: nat |
|
133 |
by (auto simp: cong_nat_def) |
|
31719 | 134 |
|
66380 | 135 |
lemma cong_sym_int: "[a = b] (mod m) \<Longrightarrow> [b = a] (mod m)" |
136 |
for a b :: int |
|
137 |
by (auto simp: cong_int_def) |
|
31719 | 138 |
|
66380 | 139 |
lemma cong_sym_eq_nat: "[a = b] (mod m) = [b = a] (mod m)" |
140 |
for a b :: nat |
|
141 |
by (auto simp: cong_nat_def) |
|
31719 | 142 |
|
66380 | 143 |
lemma cong_sym_eq_int: "[a = b] (mod m) = [b = a] (mod m)" |
144 |
for a b :: int |
|
145 |
by (auto simp: cong_int_def) |
|
31719 | 146 |
|
66380 | 147 |
lemma cong_trans_nat [trans]: "[a = b] (mod m) \<Longrightarrow> [b = c] (mod m) \<Longrightarrow> [a = c] (mod m)" |
148 |
for a b c :: nat |
|
149 |
by (auto simp: cong_nat_def) |
|
31719 | 150 |
|
66380 | 151 |
lemma cong_trans_int [trans]: "[a = b] (mod m) \<Longrightarrow> [b = c] (mod m) \<Longrightarrow> [a = c] (mod m)" |
152 |
for a b c :: int |
|
153 |
by (auto simp: cong_int_def) |
|
31719 | 154 |
|
66380 | 155 |
lemma cong_add_nat: "[a = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a + c = b + d] (mod m)" |
156 |
for a b c :: nat |
|
157 |
unfolding cong_nat_def by (metis mod_add_cong) |
|
31719 | 158 |
|
66380 | 159 |
lemma cong_add_int: "[a = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a + c = b + d] (mod m)" |
160 |
for a b c :: int |
|
161 |
unfolding cong_int_def by (metis mod_add_cong) |
|
31719 | 162 |
|
66380 | 163 |
lemma cong_diff_int: "[a = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a - c = b - d] (mod m)" |
164 |
for a b c :: int |
|
165 |
unfolding cong_int_def by (metis mod_diff_cong) |
|
31719 | 166 |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
167 |
lemma cong_diff_aux_int: |
66380 | 168 |
"[a = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> |
169 |
a \<ge> c \<Longrightarrow> b \<ge> d \<Longrightarrow> [tsub a c = tsub b d] (mod m)" |
|
170 |
for a b c d :: int |
|
55130
70db8d380d62
Restored Suc rather than +1, and using Library/Binimial
paulson <lp15@cam.ac.uk>
parents:
54489
diff
changeset
|
171 |
by (metis cong_diff_int tsub_eq) |
31719 | 172 |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
173 |
lemma cong_diff_nat: |
66380 | 174 |
fixes a b c d :: nat |
175 |
assumes "[a = b] (mod m)" "[c = d] (mod m)" "a \<ge> c" "b \<ge> d" |
|
31719 | 176 |
shows "[a - c = b - d] (mod m)" |
58860 | 177 |
using assms by (rule cong_diff_aux_int [transferred]) |
31719 | 178 |
|
66380 | 179 |
lemma cong_mult_nat: "[a = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a * c = b * d] (mod m)" |
180 |
for a b c d :: nat |
|
181 |
unfolding cong_nat_def by (metis mod_mult_cong) |
|
31719 | 182 |
|
66380 | 183 |
lemma cong_mult_int: "[a = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a * c = b * d] (mod m)" |
184 |
for a b c d :: int |
|
185 |
unfolding cong_int_def by (metis mod_mult_cong) |
|
44872 | 186 |
|
66380 | 187 |
lemma cong_exp_nat: "[x = y] (mod n) \<Longrightarrow> [x^k = y^k] (mod n)" |
188 |
for x y :: nat |
|
189 |
by (induct k) (auto simp: cong_mult_nat) |
|
44872 | 190 |
|
66380 | 191 |
lemma cong_exp_int: "[x = y] (mod n) \<Longrightarrow> [x^k = y^k] (mod n)" |
192 |
for x y :: int |
|
193 |
by (induct k) (auto simp: cong_mult_int) |
|
31719 | 194 |
|
66380 | 195 |
lemma cong_sum_nat: "(\<And>x. x \<in> A \<Longrightarrow> [f x = g x] (mod m)) \<Longrightarrow> [(\<Sum>x\<in>A. f x) = (\<Sum>x\<in>A. g x)] (mod m)" |
196 |
for f g :: "'a \<Rightarrow> nat" |
|
197 |
by (induct A rule: infinite_finite_induct) (auto intro: cong_add_nat) |
|
198 |
||
199 |
lemma cong_sum_int: "(\<And>x. x \<in> A \<Longrightarrow> [f x = g x] (mod m)) \<Longrightarrow> [(\<Sum>x\<in>A. f x) = (\<Sum>x\<in>A. g x)] (mod m)" |
|
200 |
for f g :: "'a \<Rightarrow> int" |
|
201 |
by (induct A rule: infinite_finite_induct) (auto intro: cong_add_int) |
|
31719 | 202 |
|
66380 | 203 |
lemma cong_prod_nat: "(\<And>x. x \<in> A \<Longrightarrow> [f x = g x] (mod m)) \<Longrightarrow> [(\<Prod>x\<in>A. f x) = (\<Prod>x\<in>A. g x)] (mod m)" |
204 |
for f g :: "'a \<Rightarrow> nat" |
|
205 |
by (induct A rule: infinite_finite_induct) (auto intro: cong_mult_nat) |
|
31719 | 206 |
|
66380 | 207 |
lemma cong_prod_int: "(\<And>x. x \<in> A \<Longrightarrow> [f x = g x] (mod m)) \<Longrightarrow> [(\<Prod>x\<in>A. f x) = (\<Prod>x\<in>A. g x)] (mod m)" |
208 |
for f g :: "'a \<Rightarrow> int" |
|
209 |
by (induct A rule: infinite_finite_induct) (auto intro: cong_mult_int) |
|
31719 | 210 |
|
66380 | 211 |
lemma cong_scalar_nat: "[a = b] (mod m) \<Longrightarrow> [a * k = b * k] (mod m)" |
212 |
for a b k :: nat |
|
44872 | 213 |
by (rule cong_mult_nat) simp_all |
31719 | 214 |
|
66380 | 215 |
lemma cong_scalar_int: "[a = b] (mod m) \<Longrightarrow> [a * k = b * k] (mod m)" |
216 |
for a b k :: int |
|
44872 | 217 |
by (rule cong_mult_int) simp_all |
31719 | 218 |
|
66380 | 219 |
lemma cong_scalar2_nat: "[a = b] (mod m) \<Longrightarrow> [k * a = k * b] (mod m)" |
220 |
for a b k :: nat |
|
44872 | 221 |
by (rule cong_mult_nat) simp_all |
31719 | 222 |
|
66380 | 223 |
lemma cong_scalar2_int: "[a = b] (mod m) \<Longrightarrow> [k * a = k * b] (mod m)" |
224 |
for a b k :: int |
|
44872 | 225 |
by (rule cong_mult_int) simp_all |
31719 | 226 |
|
66380 | 227 |
lemma cong_mult_self_nat: "[a * m = 0] (mod m)" |
228 |
for a m :: nat |
|
229 |
by (auto simp: cong_nat_def) |
|
31719 | 230 |
|
66380 | 231 |
lemma cong_mult_self_int: "[a * m = 0] (mod m)" |
232 |
for a m :: int |
|
233 |
by (auto simp: cong_int_def) |
|
31719 | 234 |
|
66380 | 235 |
lemma cong_eq_diff_cong_0_int: "[a = b] (mod m) = [a - b = 0] (mod m)" |
236 |
for a b :: int |
|
55130
70db8d380d62
Restored Suc rather than +1, and using Library/Binimial
paulson <lp15@cam.ac.uk>
parents:
54489
diff
changeset
|
237 |
by (metis cong_add_int cong_diff_int cong_refl_int diff_add_cancel diff_self) |
31719 | 238 |
|
66380 | 239 |
lemma cong_eq_diff_cong_0_aux_int: "a \<ge> b \<Longrightarrow> [a = b] (mod m) = [tsub a b = 0] (mod m)" |
240 |
for a b :: int |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
241 |
by (subst tsub_eq, assumption, rule cong_eq_diff_cong_0_int) |
31719 | 242 |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
243 |
lemma cong_eq_diff_cong_0_nat: |
66380 | 244 |
fixes a b :: nat |
245 |
assumes "a \<ge> b" |
|
31719 | 246 |
shows "[a = b] (mod m) = [a - b = 0] (mod m)" |
41541 | 247 |
using assms by (rule cong_eq_diff_cong_0_aux_int [transferred]) |
31719 | 248 |
|
44872 | 249 |
lemma cong_diff_cong_0'_nat: |
66380 | 250 |
"[x = y] (mod n) \<longleftrightarrow> (if x \<le> y then [y - x = 0] (mod n) else [x - y = 0] (mod n))" |
251 |
for x y :: nat |
|
55130
70db8d380d62
Restored Suc rather than +1, and using Library/Binimial
paulson <lp15@cam.ac.uk>
parents:
54489
diff
changeset
|
252 |
by (metis cong_eq_diff_cong_0_nat cong_sym_nat nat_le_linear) |
31719 | 253 |
|
66380 | 254 |
lemma cong_altdef_nat: "a \<ge> b \<Longrightarrow> [a = b] (mod m) \<longleftrightarrow> m dvd (a - b)" |
255 |
for a b :: nat |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
256 |
apply (subst cong_eq_diff_cong_0_nat, assumption) |
31719 | 257 |
apply (unfold cong_nat_def) |
258 |
apply (simp add: dvd_eq_mod_eq_0 [symmetric]) |
|
44872 | 259 |
done |
31719 | 260 |
|
66380 | 261 |
lemma cong_altdef_int: "[a = b] (mod m) \<longleftrightarrow> m dvd (a - b)" |
262 |
for a b :: int |
|
64593
50c715579715
reoriented congruence rules in non-explosive direction
haftmann
parents:
64272
diff
changeset
|
263 |
by (metis cong_int_def mod_eq_dvd_iff) |
31719 | 264 |
|
66380 | 265 |
lemma cong_abs_int: "[x = y] (mod abs m) \<longleftrightarrow> [x = y] (mod m)" |
266 |
for x y :: int |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
267 |
by (simp add: cong_altdef_int) |
31719 | 268 |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
269 |
lemma cong_square_int: |
66380 | 270 |
"prime p \<Longrightarrow> 0 < a \<Longrightarrow> [a * a = 1] (mod p) \<Longrightarrow> [a = 1] (mod p) \<or> [a = - 1] (mod p)" |
271 |
for a :: int |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
272 |
apply (simp only: cong_altdef_int) |
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
273 |
apply (subst prime_dvd_mult_eq_int [symmetric], assumption) |
36350 | 274 |
apply (auto simp add: field_simps) |
44872 | 275 |
done |
31719 | 276 |
|
66380 | 277 |
lemma cong_mult_rcancel_int: "coprime k m \<Longrightarrow> [a * k = b * k] (mod m) = [a = b] (mod m)" |
278 |
for a k m :: int |
|
62353
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
haftmann
parents:
62349
diff
changeset
|
279 |
by (metis cong_altdef_int left_diff_distrib coprime_dvd_mult_iff gcd.commute) |
31719 | 280 |
|
66380 | 281 |
lemma cong_mult_rcancel_nat: "coprime k m \<Longrightarrow> [a * k = b * k] (mod m) = [a = b] (mod m)" |
282 |
for a k m :: nat |
|
55371 | 283 |
by (metis cong_mult_rcancel_int [transferred]) |
31719 | 284 |
|
66380 | 285 |
lemma cong_mult_lcancel_nat: "coprime k m \<Longrightarrow> [k * a = k * b ] (mod m) = [a = b] (mod m)" |
286 |
for a k m :: nat |
|
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57418
diff
changeset
|
287 |
by (simp add: mult.commute cong_mult_rcancel_nat) |
31719 | 288 |
|
66380 | 289 |
lemma cong_mult_lcancel_int: "coprime k m \<Longrightarrow> [k * a = k * b] (mod m) = [a = b] (mod m)" |
290 |
for a k m :: int |
|
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57418
diff
changeset
|
291 |
by (simp add: mult.commute cong_mult_rcancel_int) |
31719 | 292 |
|
293 |
(* was zcong_zgcd_zmult_zmod *) |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
294 |
lemma coprime_cong_mult_int: |
66380 | 295 |
"[a = b] (mod m) \<Longrightarrow> [a = b] (mod n) \<Longrightarrow> coprime m n \<Longrightarrow> [a = b] (mod m * n)" |
296 |
for a b :: int |
|
297 |
by (metis divides_mult cong_altdef_int) |
|
31719 | 298 |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
299 |
lemma coprime_cong_mult_nat: |
66380 | 300 |
"[a = b] (mod m) \<Longrightarrow> [a = b] (mod n) \<Longrightarrow> coprime m n \<Longrightarrow> [a = b] (mod m * n)" |
301 |
for a b :: nat |
|
302 |
by (metis coprime_cong_mult_int [transferred]) |
|
31719 | 303 |
|
66380 | 304 |
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" |
305 |
for a b :: nat |
|
41541 | 306 |
by (auto simp add: cong_nat_def) |
31719 | 307 |
|
66380 | 308 |
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" |
309 |
for a b :: int |
|
41541 | 310 |
by (auto simp add: cong_int_def) |
31719 | 311 |
|
66380 | 312 |
lemma cong_less_unique_nat: "0 < m \<Longrightarrow> (\<exists>!b. 0 \<le> b \<and> b < m \<and> [a = b] (mod m))" |
313 |
for a m :: nat |
|
55371 | 314 |
by (auto simp: cong_nat_def) (metis mod_less_divisor mod_mod_trivial) |
31719 | 315 |
|
66380 | 316 |
lemma cong_less_unique_int: "0 < m \<Longrightarrow> (\<exists>!b. 0 \<le> b \<and> b < m \<and> [a = b] (mod m))" |
317 |
for a m :: int |
|
55371 | 318 |
by (auto simp: cong_int_def) (metis mod_mod_trivial pos_mod_conj) |
31719 | 319 |
|
66380 | 320 |
lemma cong_iff_lin_int: "[a = b] (mod m) \<longleftrightarrow> (\<exists>k. b = a + m * k)" |
321 |
for a b :: int |
|
55371 | 322 |
apply (auto simp add: cong_altdef_int dvd_def) |
31719 | 323 |
apply (rule_tac [!] x = "-k" in exI, auto) |
44872 | 324 |
done |
31719 | 325 |
|
66380 | 326 |
lemma cong_iff_lin_nat: "([a = b] (mod m)) \<longleftrightarrow> (\<exists>k1 k2. b + k1 * m = a + k2 * m)" |
327 |
(is "?lhs = ?rhs") |
|
328 |
for a b :: nat |
|
329 |
proof |
|
330 |
assume ?lhs |
|
55371 | 331 |
show ?rhs |
332 |
proof (cases "b \<le> a") |
|
333 |
case True |
|
66380 | 334 |
with \<open>?lhs\<close> show ?rhs |
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57418
diff
changeset
|
335 |
by (metis cong_altdef_nat dvd_def le_add_diff_inverse add_0_right mult_0 mult.commute) |
55371 | 336 |
next |
337 |
case False |
|
66380 | 338 |
with \<open>?lhs\<close> show ?rhs |
55371 | 339 |
apply (subst (asm) cong_sym_eq_nat) |
340 |
apply (auto simp: cong_altdef_nat) |
|
341 |
apply (metis add_0_right add_diff_inverse dvd_div_mult_self less_or_eq_imp_le mult_0) |
|
342 |
done |
|
343 |
qed |
|
344 |
next |
|
345 |
assume ?rhs |
|
346 |
then show ?lhs |
|
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57418
diff
changeset
|
347 |
by (metis cong_nat_def mod_mult_self2 mult.commute) |
55371 | 348 |
qed |
31719 | 349 |
|
66380 | 350 |
lemma cong_gcd_eq_int: "[a = b] (mod m) \<Longrightarrow> gcd a m = gcd b m" |
351 |
for a b :: int |
|
55371 | 352 |
by (metis cong_int_def gcd_red_int) |
31719 | 353 |
|
66380 | 354 |
lemma cong_gcd_eq_nat: "[a = b] (mod m) \<Longrightarrow> gcd a m = gcd b m" |
355 |
for a b :: nat |
|
63092 | 356 |
by (metis cong_gcd_eq_int [transferred]) |
31719 | 357 |
|
66380 | 358 |
lemma cong_imp_coprime_nat: "[a = b] (mod m) \<Longrightarrow> coprime a m \<Longrightarrow> coprime b m" |
359 |
for a b :: nat |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
360 |
by (auto simp add: cong_gcd_eq_nat) |
31719 | 361 |
|
66380 | 362 |
lemma cong_imp_coprime_int: "[a = b] (mod m) \<Longrightarrow> coprime a m \<Longrightarrow> coprime b m" |
363 |
for a b :: int |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
364 |
by (auto simp add: cong_gcd_eq_int) |
31719 | 365 |
|
66380 | 366 |
lemma cong_cong_mod_nat: "[a = b] (mod m) \<longleftrightarrow> [a mod m = b mod m] (mod m)" |
367 |
for a b :: nat |
|
31719 | 368 |
by (auto simp add: cong_nat_def) |
369 |
||
66380 | 370 |
lemma cong_cong_mod_int: "[a = b] (mod m) \<longleftrightarrow> [a mod m = b mod m] (mod m)" |
371 |
for a b :: int |
|
31719 | 372 |
by (auto simp add: cong_int_def) |
373 |
||
66380 | 374 |
lemma cong_minus_int [iff]: "[a = b] (mod - m) \<longleftrightarrow> [a = b] (mod m)" |
375 |
for a b :: int |
|
55371 | 376 |
by (metis cong_iff_lin_int minus_equation_iff mult_minus_left mult_minus_right) |
31719 | 377 |
|
378 |
(* |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
379 |
lemma mod_dvd_mod_int: |
31719 | 380 |
"0 < (m::int) \<Longrightarrow> m dvd b \<Longrightarrow> (a mod b mod m) = (a mod m)" |
381 |
apply (unfold dvd_def, auto) |
|
382 |
apply (rule mod_mod_cancel) |
|
383 |
apply auto |
|
44872 | 384 |
done |
31719 | 385 |
|
386 |
lemma mod_dvd_mod: |
|
387 |
assumes "0 < (m::nat)" and "m dvd b" |
|
388 |
shows "(a mod b mod m) = (a mod m)" |
|
389 |
||
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
390 |
apply (rule mod_dvd_mod_int [transferred]) |
41541 | 391 |
using assms apply auto |
392 |
done |
|
31719 | 393 |
*) |
394 |
||
66380 | 395 |
lemma cong_add_lcancel_nat: "[a + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)" |
396 |
for a x y :: nat |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
397 |
by (simp add: cong_iff_lin_nat) |
31719 | 398 |
|
66380 | 399 |
lemma cong_add_lcancel_int: "[a + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)" |
400 |
for a x y :: int |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
401 |
by (simp add: cong_iff_lin_int) |
31719 | 402 |
|
66380 | 403 |
lemma cong_add_rcancel_nat: "[x + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)" |
404 |
for a x y :: nat |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
405 |
by (simp add: cong_iff_lin_nat) |
31719 | 406 |
|
66380 | 407 |
lemma cong_add_rcancel_int: "[x + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)" |
408 |
for a x y :: int |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
409 |
by (simp add: cong_iff_lin_int) |
31719 | 410 |
|
66380 | 411 |
lemma cong_add_lcancel_0_nat: "[a + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)" |
412 |
for a x :: nat |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
413 |
by (simp add: cong_iff_lin_nat) |
31719 | 414 |
|
66380 | 415 |
lemma cong_add_lcancel_0_int: "[a + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)" |
416 |
for a x :: int |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
417 |
by (simp add: cong_iff_lin_int) |
31719 | 418 |
|
66380 | 419 |
lemma cong_add_rcancel_0_nat: "[x + a = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)" |
420 |
for a x :: nat |
|
421 |
by (simp add: cong_iff_lin_nat) |
|
422 |
||
423 |
lemma cong_add_rcancel_0_int: "[x + a = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)" |
|
424 |
for a x :: int |
|
425 |
by (simp add: cong_iff_lin_int) |
|
426 |
||
427 |
lemma cong_dvd_modulus_nat: "[x = y] (mod m) \<Longrightarrow> n dvd m \<Longrightarrow> [x = y] (mod n)" |
|
428 |
for x y :: nat |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
429 |
apply (auto simp add: cong_iff_lin_nat dvd_def) |
66380 | 430 |
apply (rule_tac x= "k1 * k" in exI) |
431 |
apply (rule_tac x= "k2 * k" in exI) |
|
36350 | 432 |
apply (simp add: field_simps) |
44872 | 433 |
done |
31719 | 434 |
|
66380 | 435 |
lemma cong_dvd_modulus_int: "[x = y] (mod m) \<Longrightarrow> n dvd m \<Longrightarrow> [x = y] (mod n)" |
436 |
for x y :: int |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
437 |
by (auto simp add: cong_altdef_int dvd_def) |
31719 | 438 |
|
66380 | 439 |
lemma cong_dvd_eq_nat: "[x = y] (mod n) \<Longrightarrow> n dvd x \<longleftrightarrow> n dvd y" |
440 |
for x y :: nat |
|
441 |
by (auto simp: cong_nat_def dvd_eq_mod_eq_0) |
|
31719 | 442 |
|
66380 | 443 |
lemma cong_dvd_eq_int: "[x = y] (mod n) \<Longrightarrow> n dvd x \<longleftrightarrow> n dvd y" |
444 |
for x y :: int |
|
445 |
by (auto simp: cong_int_def dvd_eq_mod_eq_0) |
|
31719 | 446 |
|
66380 | 447 |
lemma cong_mod_nat: "n \<noteq> 0 \<Longrightarrow> [a mod n = a] (mod n)" |
448 |
for a n :: nat |
|
31719 | 449 |
by (simp add: cong_nat_def) |
450 |
||
66380 | 451 |
lemma cong_mod_int: "n \<noteq> 0 \<Longrightarrow> [a mod n = a] (mod n)" |
452 |
for a n :: int |
|
31719 | 453 |
by (simp add: cong_int_def) |
454 |
||
66380 | 455 |
lemma mod_mult_cong_nat: "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> [x mod (a * b) = y] (mod a) \<longleftrightarrow> [x = y] (mod a)" |
456 |
for a b :: nat |
|
31719 | 457 |
by (simp add: cong_nat_def mod_mult2_eq mod_add_left_eq) |
458 |
||
66380 | 459 |
lemma neg_cong_int: "[a = b] (mod m) \<longleftrightarrow> [- a = - b] (mod m)" |
460 |
for a b :: int |
|
64593
50c715579715
reoriented congruence rules in non-explosive direction
haftmann
parents:
64272
diff
changeset
|
461 |
by (metis cong_int_def minus_minus mod_minus_cong) |
31719 | 462 |
|
66380 | 463 |
lemma cong_modulus_neg_int: "[a = b] (mod m) \<longleftrightarrow> [a = b] (mod - m)" |
464 |
for a b :: int |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
465 |
by (auto simp add: cong_altdef_int) |
31719 | 466 |
|
66380 | 467 |
lemma mod_mult_cong_int: "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> [x mod (a * b) = y] (mod a) \<longleftrightarrow> [x = y] (mod a)" |
468 |
for a b :: int |
|
469 |
proof (cases "b > 0") |
|
470 |
case True |
|
471 |
then show ?thesis |
|
472 |
by (simp add: cong_int_def mod_mod_cancel mod_add_left_eq) |
|
473 |
next |
|
474 |
case False |
|
475 |
then show ?thesis |
|
476 |
apply (subst (1 2) cong_modulus_neg_int) |
|
477 |
apply (unfold cong_int_def) |
|
478 |
apply (subgoal_tac "a * b = (- a * - b)") |
|
479 |
apply (erule ssubst) |
|
480 |
apply (subst zmod_zmult2_eq) |
|
481 |
apply (auto simp add: mod_add_left_eq mod_minus_right div_minus_right) |
|
482 |
apply (metis mod_diff_left_eq mod_diff_right_eq mod_mult_self1_is_0 diff_zero)+ |
|
483 |
done |
|
484 |
qed |
|
485 |
||
486 |
lemma cong_to_1_nat: |
|
487 |
fixes a :: nat |
|
488 |
assumes "[a = 1] (mod n)" |
|
489 |
shows "n dvd (a - 1)" |
|
490 |
proof (cases "a = 0") |
|
491 |
case True |
|
492 |
then show ?thesis by force |
|
493 |
next |
|
494 |
case False |
|
495 |
with assms show ?thesis by (metis cong_altdef_nat leI less_one) |
|
496 |
qed |
|
497 |
||
498 |
lemma cong_0_1_nat': "[0 = Suc 0] (mod n) \<longleftrightarrow> n = Suc 0" |
|
499 |
by (auto simp: cong_nat_def) |
|
500 |
||
501 |
lemma cong_0_1_nat: "[0 = 1] (mod n) \<longleftrightarrow> n = 1" |
|
502 |
for n :: nat |
|
503 |
by (auto simp: cong_nat_def) |
|
504 |
||
505 |
lemma cong_0_1_int: "[0 = 1] (mod n) \<longleftrightarrow> n = 1 \<or> n = - 1" |
|
506 |
for n :: int |
|
507 |
by (auto simp: cong_int_def zmult_eq_1_iff) |
|
508 |
||
509 |
lemma cong_to_1'_nat: "[a = 1] (mod n) \<longleftrightarrow> a = 0 \<and> n = 1 \<or> (\<exists>m. a = 1 + m * n)" |
|
510 |
for a :: nat |
|
511 |
by (metis add.right_neutral cong_0_1_nat cong_iff_lin_nat cong_to_1_nat |
|
512 |
dvd_div_mult_self leI le_add_diff_inverse less_one mult_eq_if) |
|
513 |
||
514 |
lemma cong_le_nat: "y \<le> x \<Longrightarrow> [x = y] (mod n) \<longleftrightarrow> (\<exists>q. x = q * n + y)" |
|
515 |
for x y :: nat |
|
516 |
by (metis cong_altdef_nat Nat.le_imp_diff_is_add dvd_def mult.commute) |
|
517 |
||
518 |
lemma cong_solve_nat: |
|
519 |
fixes a :: nat |
|
520 |
assumes "a \<noteq> 0" |
|
521 |
shows "\<exists>x. [a * x = gcd a n] (mod n)" |
|
522 |
proof (cases "n = 0") |
|
523 |
case True |
|
524 |
then show ?thesis by force |
|
525 |
next |
|
526 |
case False |
|
527 |
then show ?thesis |
|
528 |
using bezout_nat [of a n, OF \<open>a \<noteq> 0\<close>] |
|
529 |
by auto (metis cong_add_rcancel_0_nat cong_mult_self_nat mult.commute) |
|
530 |
qed |
|
531 |
||
532 |
lemma cong_solve_int: "a \<noteq> 0 \<Longrightarrow> \<exists>x. [a * x = gcd a n] (mod n)" |
|
533 |
for a :: int |
|
534 |
apply (cases "n = 0") |
|
535 |
apply (cases "a \<ge> 0") |
|
536 |
apply auto |
|
537 |
apply (rule_tac x = "-1" in exI) |
|
538 |
apply auto |
|
539 |
apply (insert bezout_int [of a n], auto) |
|
540 |
apply (metis cong_iff_lin_int mult.commute) |
|
44872 | 541 |
done |
31719 | 542 |
|
44872 | 543 |
lemma cong_solve_dvd_nat: |
66380 | 544 |
fixes a :: nat |
545 |
assumes a: "a \<noteq> 0" and b: "gcd a n dvd d" |
|
546 |
shows "\<exists>x. [a * x = d] (mod n)" |
|
31719 | 547 |
proof - |
44872 | 548 |
from cong_solve_nat [OF a] obtain x where "[a * x = gcd a n](mod n)" |
31719 | 549 |
by auto |
44872 | 550 |
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
|
551 |
by (elim cong_scalar2_nat) |
31719 | 552 |
also from b have "(d div gcd a n) * gcd a n = d" |
553 |
by (rule dvd_div_mult_self) |
|
554 |
also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)" |
|
555 |
by auto |
|
556 |
finally show ?thesis |
|
557 |
by auto |
|
558 |
qed |
|
559 |
||
44872 | 560 |
lemma cong_solve_dvd_int: |
31719 | 561 |
assumes a: "(a::int) \<noteq> 0" and b: "gcd a n dvd d" |
66380 | 562 |
shows "\<exists>x. [a * x = d] (mod n)" |
31719 | 563 |
proof - |
44872 | 564 |
from cong_solve_int [OF a] obtain x where "[a * x = gcd a n](mod n)" |
31719 | 565 |
by auto |
44872 | 566 |
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
|
567 |
by (elim cong_scalar2_int) |
31719 | 568 |
also from b have "(d div gcd a n) * gcd a n = d" |
569 |
by (rule dvd_div_mult_self) |
|
570 |
also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)" |
|
571 |
by auto |
|
572 |
finally show ?thesis |
|
573 |
by auto |
|
574 |
qed |
|
575 |
||
66380 | 576 |
lemma cong_solve_coprime_nat: |
577 |
fixes a :: nat |
|
578 |
assumes "coprime a n" |
|
579 |
shows "\<exists>x. [a * x = 1] (mod n)" |
|
580 |
proof (cases "a = 0") |
|
581 |
case True |
|
582 |
with assms show ?thesis by force |
|
583 |
next |
|
584 |
case False |
|
585 |
with assms show ?thesis by (metis cong_solve_nat) |
|
586 |
qed |
|
31719 | 587 |
|
66380 | 588 |
lemma cong_solve_coprime_int: "coprime (a::int) n \<Longrightarrow> \<exists>x. [a * x = 1] (mod n)" |
44872 | 589 |
apply (cases "a = 0") |
66380 | 590 |
apply auto |
591 |
apply (cases "n \<ge> 0") |
|
592 |
apply auto |
|
55161 | 593 |
apply (metis cong_solve_int) |
594 |
done |
|
595 |
||
62349
7c23469b5118
cleansed junk-producing interpretations for gcd/lcm on nat altogether
haftmann
parents:
62348
diff
changeset
|
596 |
lemma coprime_iff_invertible_nat: |
66380 | 597 |
"m > 0 \<Longrightarrow> coprime a m = (\<exists>x. [a * x = Suc 0] (mod m))" |
62429
25271ff79171
Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents:
62353
diff
changeset
|
598 |
by (metis One_nat_def cong_gcd_eq_nat cong_solve_coprime_nat coprime_lmult gcd.commute gcd_Suc_0) |
66380 | 599 |
|
600 |
lemma coprime_iff_invertible_int: "m > 0 \<Longrightarrow> coprime a m \<longleftrightarrow> (\<exists>x. [a * x = 1] (mod m))" |
|
601 |
for m :: int |
|
55161 | 602 |
apply (auto intro: cong_solve_coprime_int) |
62429
25271ff79171
Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents:
62353
diff
changeset
|
603 |
apply (metis cong_int_def coprime_mul_eq gcd_1_int gcd.commute gcd_red_int) |
44872 | 604 |
done |
31719 | 605 |
|
66380 | 606 |
lemma coprime_iff_invertible'_nat: |
607 |
"m > 0 \<Longrightarrow> coprime a m \<longleftrightarrow> (\<exists>x. 0 \<le> x \<and> x < m \<and> [a * x = Suc 0] (mod m))" |
|
55161 | 608 |
apply (subst coprime_iff_invertible_nat) |
66380 | 609 |
apply auto |
55161 | 610 |
apply (auto simp add: cong_nat_def) |
611 |
apply (metis mod_less_divisor mod_mult_right_eq) |
|
44872 | 612 |
done |
31719 | 613 |
|
66380 | 614 |
lemma coprime_iff_invertible'_int: |
615 |
"m > 0 \<Longrightarrow> coprime a m \<longleftrightarrow> (\<exists>x. 0 \<le> x \<and> x < m \<and> [a * x = 1] (mod m))" |
|
616 |
for m :: int |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
617 |
apply (subst coprime_iff_invertible_int) |
66380 | 618 |
apply (auto simp add: cong_int_def) |
55371 | 619 |
apply (metis mod_mult_right_eq pos_mod_conj) |
44872 | 620 |
done |
31719 | 621 |
|
66380 | 622 |
lemma cong_cong_lcm_nat: "[x = y] (mod a) \<Longrightarrow> [x = y] (mod b) \<Longrightarrow> [x = y] (mod lcm a b)" |
623 |
for x y :: nat |
|
44872 | 624 |
apply (cases "y \<le> x") |
62348 | 625 |
apply (metis cong_altdef_nat lcm_least) |
62349
7c23469b5118
cleansed junk-producing interpretations for gcd/lcm on nat altogether
haftmann
parents:
62348
diff
changeset
|
626 |
apply (meson cong_altdef_nat cong_sym_nat lcm_least_iff nat_le_linear) |
44872 | 627 |
done |
31719 | 628 |
|
66380 | 629 |
lemma cong_cong_lcm_int: "[x = y] (mod a) \<Longrightarrow> [x = y] (mod b) \<Longrightarrow> [x = y] (mod lcm a b)" |
630 |
for x y :: int |
|
631 |
by (auto simp add: cong_altdef_int lcm_least) |
|
31719 | 632 |
|
64272 | 633 |
lemma cong_cong_prod_coprime_nat [rule_format]: "finite A \<Longrightarrow> |
61954 | 634 |
(\<forall>i\<in>A. (\<forall>j\<in>A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))) \<longrightarrow> |
635 |
(\<forall>i\<in>A. [(x::nat) = y] (mod m i)) \<longrightarrow> |
|
636 |
[x = y] (mod (\<Prod>i\<in>A. m i))" |
|
31719 | 637 |
apply (induct set: finite) |
638 |
apply auto |
|
64272 | 639 |
apply (metis One_nat_def coprime_cong_mult_nat gcd.commute prod_coprime) |
44872 | 640 |
done |
31719 | 641 |
|
64272 | 642 |
lemma cong_cong_prod_coprime_int [rule_format]: "finite A \<Longrightarrow> |
61954 | 643 |
(\<forall>i\<in>A. (\<forall>j\<in>A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))) \<longrightarrow> |
644 |
(\<forall>i\<in>A. [(x::int) = y] (mod m i)) \<longrightarrow> |
|
645 |
[x = y] (mod (\<Prod>i\<in>A. m i))" |
|
31719 | 646 |
apply (induct set: finite) |
647 |
apply auto |
|
64272 | 648 |
apply (metis coprime_cong_mult_int gcd.commute prod_coprime) |
44872 | 649 |
done |
31719 | 650 |
|
44872 | 651 |
lemma binary_chinese_remainder_aux_nat: |
66380 | 652 |
fixes m1 m2 :: nat |
653 |
assumes a: "coprime m1 m2" |
|
654 |
shows "\<exists>b1 b2. [b1 = 1] (mod m1) \<and> [b1 = 0] (mod m2) \<and> [b2 = 0] (mod m1) \<and> [b2 = 1] (mod m2)" |
|
31719 | 655 |
proof - |
66380 | 656 |
from cong_solve_coprime_nat [OF a] obtain x1 where 1: "[m1 * x1 = 1] (mod m2)" |
31719 | 657 |
by auto |
44872 | 658 |
from a have b: "coprime m2 m1" |
62348 | 659 |
by (subst gcd.commute) |
66380 | 660 |
from cong_solve_coprime_nat [OF b] obtain x2 where 2: "[m2 * x2 = 1] (mod m1)" |
31719 | 661 |
by auto |
662 |
have "[m1 * x1 = 0] (mod m1)" |
|
66380 | 663 |
by (subst mult.commute) (rule cong_mult_self_nat) |
31719 | 664 |
moreover have "[m2 * x2 = 0] (mod m2)" |
66380 | 665 |
by (subst mult.commute) (rule cong_mult_self_nat) |
666 |
ultimately show ?thesis |
|
667 |
using 1 2 by blast |
|
31719 | 668 |
qed |
669 |
||
44872 | 670 |
lemma binary_chinese_remainder_aux_int: |
66380 | 671 |
fixes m1 m2 :: int |
672 |
assumes a: "coprime m1 m2" |
|
673 |
shows "\<exists>b1 b2. [b1 = 1] (mod m1) \<and> [b1 = 0] (mod m2) \<and> [b2 = 0] (mod m1) \<and> [b2 = 1] (mod m2)" |
|
31719 | 674 |
proof - |
66380 | 675 |
from cong_solve_coprime_int [OF a] obtain x1 where 1: "[m1 * x1 = 1] (mod m2)" |
31719 | 676 |
by auto |
44872 | 677 |
from a have b: "coprime m2 m1" |
62348 | 678 |
by (subst gcd.commute) |
66380 | 679 |
from cong_solve_coprime_int [OF b] obtain x2 where 2: "[m2 * x2 = 1] (mod m1)" |
31719 | 680 |
by auto |
681 |
have "[m1 * x1 = 0] (mod m1)" |
|
66380 | 682 |
by (subst mult.commute) (rule cong_mult_self_int) |
31719 | 683 |
moreover have "[m2 * x2 = 0] (mod m2)" |
66380 | 684 |
by (subst mult.commute) (rule cong_mult_self_int) |
685 |
ultimately show ?thesis |
|
686 |
using 1 2 by blast |
|
31719 | 687 |
qed |
688 |
||
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
689 |
lemma binary_chinese_remainder_nat: |
66380 | 690 |
fixes m1 m2 :: nat |
691 |
assumes a: "coprime m1 m2" |
|
692 |
shows "\<exists>x. [x = u1] (mod m1) \<and> [x = u2] (mod m2)" |
|
31719 | 693 |
proof - |
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
694 |
from binary_chinese_remainder_aux_nat [OF a] obtain b1 b2 |
66380 | 695 |
where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)" |
696 |
and "[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)" |
|
31719 | 697 |
by blast |
698 |
let ?x = "u1 * b1 + u2 * b2" |
|
699 |
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
|
700 |
apply (rule cong_add_nat) |
66380 | 701 |
apply (rule cong_scalar2_nat) |
702 |
apply (rule \<open>[b1 = 1] (mod m1)\<close>) |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
703 |
apply (rule cong_scalar2_nat) |
60526 | 704 |
apply (rule \<open>[b2 = 0] (mod m1)\<close>) |
31719 | 705 |
done |
44872 | 706 |
then have "[?x = u1] (mod m1)" by simp |
31719 | 707 |
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
|
708 |
apply (rule cong_add_nat) |
66380 | 709 |
apply (rule cong_scalar2_nat) |
710 |
apply (rule \<open>[b1 = 0] (mod m2)\<close>) |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
711 |
apply (rule cong_scalar2_nat) |
60526 | 712 |
apply (rule \<open>[b2 = 1] (mod m2)\<close>) |
31719 | 713 |
done |
66380 | 714 |
then have "[?x = u2] (mod m2)" |
715 |
by simp |
|
716 |
with \<open>[?x = u1] (mod m1)\<close> show ?thesis |
|
717 |
by blast |
|
31719 | 718 |
qed |
719 |
||
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
720 |
lemma binary_chinese_remainder_int: |
66380 | 721 |
fixes m1 m2 :: int |
722 |
assumes a: "coprime m1 m2" |
|
723 |
shows "\<exists>x. [x = u1] (mod m1) \<and> [x = u2] (mod m2)" |
|
31719 | 724 |
proof - |
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
725 |
from binary_chinese_remainder_aux_int [OF a] obtain b1 b2 |
66380 | 726 |
where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)" |
727 |
and "[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)" |
|
31719 | 728 |
by blast |
729 |
let ?x = "u1 * b1 + u2 * b2" |
|
730 |
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
|
731 |
apply (rule cong_add_int) |
66380 | 732 |
apply (rule cong_scalar2_int) |
733 |
apply (rule \<open>[b1 = 1] (mod m1)\<close>) |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
734 |
apply (rule cong_scalar2_int) |
60526 | 735 |
apply (rule \<open>[b2 = 0] (mod m1)\<close>) |
31719 | 736 |
done |
44872 | 737 |
then have "[?x = u1] (mod m1)" by simp |
31719 | 738 |
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
|
739 |
apply (rule cong_add_int) |
66380 | 740 |
apply (rule cong_scalar2_int) |
741 |
apply (rule \<open>[b1 = 0] (mod m2)\<close>) |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
742 |
apply (rule cong_scalar2_int) |
60526 | 743 |
apply (rule \<open>[b2 = 1] (mod m2)\<close>) |
31719 | 744 |
done |
44872 | 745 |
then have "[?x = u2] (mod m2)" by simp |
66380 | 746 |
with \<open>[?x = u1] (mod m1)\<close> show ?thesis |
747 |
by blast |
|
31719 | 748 |
qed |
749 |
||
66380 | 750 |
lemma cong_modulus_mult_nat: "[x = y] (mod m * n) \<Longrightarrow> [x = y] (mod m)" |
751 |
for x y :: nat |
|
44872 | 752 |
apply (cases "y \<le> x") |
66380 | 753 |
apply (simp add: cong_altdef_nat) |
754 |
apply (erule dvd_mult_left) |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
755 |
apply (rule cong_sym_nat) |
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
756 |
apply (subst (asm) cong_sym_eq_nat) |
44872 | 757 |
apply (simp add: cong_altdef_nat) |
31719 | 758 |
apply (erule dvd_mult_left) |
44872 | 759 |
done |
31719 | 760 |
|
66380 | 761 |
lemma cong_modulus_mult_int: "[x = y] (mod m * n) \<Longrightarrow> [x = y] (mod m)" |
762 |
for x y :: int |
|
44872 | 763 |
apply (simp add: cong_altdef_int) |
31719 | 764 |
apply (erule dvd_mult_left) |
44872 | 765 |
done |
31719 | 766 |
|
66380 | 767 |
lemma cong_less_modulus_unique_nat: "[x = y] (mod m) \<Longrightarrow> x < m \<Longrightarrow> y < m \<Longrightarrow> x = y" |
768 |
for x y :: nat |
|
31719 | 769 |
by (simp add: cong_nat_def) |
770 |
||
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
771 |
lemma binary_chinese_remainder_unique_nat: |
66380 | 772 |
fixes m1 m2 :: nat |
773 |
assumes a: "coprime m1 m2" |
|
44872 | 774 |
and nz: "m1 \<noteq> 0" "m2 \<noteq> 0" |
63901 | 775 |
shows "\<exists>!x. x < m1 * m2 \<and> [x = u1] (mod m1) \<and> [x = u2] (mod m2)" |
31719 | 776 |
proof - |
66380 | 777 |
from binary_chinese_remainder_nat [OF a] obtain y |
778 |
where "[y = u1] (mod m1)" and "[y = u2] (mod m2)" |
|
31719 | 779 |
by blast |
780 |
let ?x = "y mod (m1 * m2)" |
|
781 |
from nz have less: "?x < m1 * m2" |
|
44872 | 782 |
by auto |
66380 | 783 |
have 1: "[?x = u1] (mod m1)" |
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
784 |
apply (rule cong_trans_nat) |
66380 | 785 |
prefer 2 |
786 |
apply (rule \<open>[y = u1] (mod m1)\<close>) |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
787 |
apply (rule cong_modulus_mult_nat) |
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
788 |
apply (rule cong_mod_nat) |
31719 | 789 |
using nz apply auto |
790 |
done |
|
66380 | 791 |
have 2: "[?x = u2] (mod m2)" |
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
792 |
apply (rule cong_trans_nat) |
66380 | 793 |
prefer 2 |
794 |
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
|
795 |
apply (subst mult.commute) |
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
796 |
apply (rule cong_modulus_mult_nat) |
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
797 |
apply (rule cong_mod_nat) |
31719 | 798 |
using nz apply auto |
799 |
done |
|
66380 | 800 |
have "\<forall>z. z < m1 * m2 \<and> [z = u1] (mod m1) \<and> [z = u2] (mod m2) \<longrightarrow> z = ?x" |
44872 | 801 |
proof clarify |
31719 | 802 |
fix z |
803 |
assume "z < m1 * m2" |
|
804 |
assume "[z = u1] (mod m1)" and "[z = u2] (mod m2)" |
|
805 |
have "[?x = z] (mod m1)" |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
806 |
apply (rule cong_trans_nat) |
66380 | 807 |
apply (rule \<open>[?x = u1] (mod m1)\<close>) |
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
808 |
apply (rule cong_sym_nat) |
60526 | 809 |
apply (rule \<open>[z = u1] (mod m1)\<close>) |
31719 | 810 |
done |
811 |
moreover have "[?x = z] (mod m2)" |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
812 |
apply (rule cong_trans_nat) |
66380 | 813 |
apply (rule \<open>[?x = u2] (mod m2)\<close>) |
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
814 |
apply (rule cong_sym_nat) |
60526 | 815 |
apply (rule \<open>[z = u2] (mod m2)\<close>) |
31719 | 816 |
done |
817 |
ultimately have "[?x = z] (mod m1 * m2)" |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
818 |
by (auto intro: coprime_cong_mult_nat a) |
60526 | 819 |
with \<open>z < m1 * m2\<close> \<open>?x < m1 * m2\<close> show "z = ?x" |
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
820 |
apply (intro cong_less_modulus_unique_nat) |
66380 | 821 |
apply (auto, erule cong_sym_nat) |
31719 | 822 |
done |
44872 | 823 |
qed |
66380 | 824 |
with less 1 2 show ?thesis by auto |
31719 | 825 |
qed |
826 |
||
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
827 |
lemma chinese_remainder_aux_nat: |
44872 | 828 |
fixes A :: "'a set" |
829 |
and m :: "'a \<Rightarrow> nat" |
|
830 |
assumes fin: "finite A" |
|
66380 | 831 |
and cop: "\<forall>i \<in> A. (\<forall>j \<in> A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))" |
832 |
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 | 833 |
proof (rule finite_set_choice, rule fin, rule ballI) |
834 |
fix i |
|
66380 | 835 |
assume "i \<in> A" |
61954 | 836 |
with cop have "coprime (\<Prod>j \<in> A - {i}. m j) (m i)" |
66380 | 837 |
by (intro prod_coprime) auto |
838 |
then have "\<exists>x. [(\<Prod>j \<in> A - {i}. m j) * x = 1] (mod m i)" |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
839 |
by (elim cong_solve_coprime_nat) |
61954 | 840 |
then obtain x where "[(\<Prod>j \<in> A - {i}. m j) * x = 1] (mod m i)" |
31719 | 841 |
by auto |
66380 | 842 |
moreover have "[(\<Prod>j \<in> A - {i}. m j) * x = 0] (mod (\<Prod>j \<in> A - {i}. m j))" |
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57418
diff
changeset
|
843 |
by (subst mult.commute, rule cong_mult_self_nat) |
66380 | 844 |
ultimately show "\<exists>a. [a = 1] (mod m i) \<and> [a = 0] (mod prod m (A - {i}))" |
31719 | 845 |
by blast |
846 |
qed |
|
847 |
||
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
848 |
lemma chinese_remainder_nat: |
44872 | 849 |
fixes A :: "'a set" |
850 |
and m :: "'a \<Rightarrow> nat" |
|
851 |
and u :: "'a \<Rightarrow> nat" |
|
852 |
assumes fin: "finite A" |
|
66380 | 853 |
and cop: "\<forall>i \<in> A. \<forall>j \<in> A. i \<noteq> j \<longrightarrow> coprime (m i) (m j)" |
854 |
shows "\<exists>x. \<forall>i \<in> A. [x = u i] (mod m i)" |
|
31719 | 855 |
proof - |
66380 | 856 |
from chinese_remainder_aux_nat [OF fin cop] |
857 |
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 | 858 |
by blast |
61954 | 859 |
let ?x = "\<Sum>i\<in>A. (u i) * (b i)" |
66380 | 860 |
show ?thesis |
31719 | 861 |
proof (rule exI, clarify) |
862 |
fix i |
|
66380 | 863 |
assume a: "i \<in> A" |
44872 | 864 |
show "[?x = u i] (mod m i)" |
31719 | 865 |
proof - |
66380 | 866 |
from fin a have "?x = (\<Sum>j \<in> {i}. u j * b j) + (\<Sum>j \<in> A - {i}. u j * b j)" |
867 |
by (subst sum.union_disjoint [symmetric]) (auto intro: sum.cong) |
|
61954 | 868 |
then have "[?x = u i * b i + (\<Sum>j \<in> A - {i}. u j * b j)] (mod m i)" |
31719 | 869 |
by auto |
61954 | 870 |
also have "[u i * b i + (\<Sum>j \<in> A - {i}. u j * b j) = |
871 |
u i * 1 + (\<Sum>j \<in> 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
|
872 |
apply (rule cong_add_nat) |
66380 | 873 |
apply (rule cong_scalar2_nat) |
874 |
using b a apply blast |
|
64267 | 875 |
apply (rule cong_sum_nat) |
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
876 |
apply (rule cong_scalar2_nat) |
66380 | 877 |
using b apply auto |
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
878 |
apply (rule cong_dvd_modulus_nat) |
66380 | 879 |
apply (drule (1) bspec) |
880 |
apply (erule conjE) |
|
881 |
apply assumption |
|
59010 | 882 |
apply rule |
31719 | 883 |
using fin a apply auto |
884 |
done |
|
885 |
finally show ?thesis |
|
886 |
by simp |
|
887 |
qed |
|
888 |
qed |
|
889 |
qed |
|
890 |
||
44872 | 891 |
lemma coprime_cong_prod_nat [rule_format]: "finite A \<Longrightarrow> |
61954 | 892 |
(\<forall>i\<in>A. (\<forall>j\<in>A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))) \<longrightarrow> |
893 |
(\<forall>i\<in>A. [(x::nat) = y] (mod m i)) \<longrightarrow> |
|
894 |
[x = y] (mod (\<Prod>i\<in>A. m i))" |
|
31719 | 895 |
apply (induct set: finite) |
896 |
apply auto |
|
64272 | 897 |
apply (metis One_nat_def coprime_cong_mult_nat gcd.commute prod_coprime) |
44872 | 898 |
done |
31719 | 899 |
|
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
900 |
lemma chinese_remainder_unique_nat: |
44872 | 901 |
fixes A :: "'a set" |
902 |
and m :: "'a \<Rightarrow> nat" |
|
903 |
and u :: "'a \<Rightarrow> nat" |
|
904 |
assumes fin: "finite A" |
|
61954 | 905 |
and nz: "\<forall>i\<in>A. m i \<noteq> 0" |
66380 | 906 |
and cop: "\<forall>i\<in>A. \<forall>j\<in>A. i \<noteq> j \<longrightarrow> coprime (m i) (m j)" |
63901 | 907 |
shows "\<exists>!x. x < (\<Prod>i\<in>A. m i) \<and> (\<forall>i\<in>A. [x = u i] (mod m i))" |
31719 | 908 |
proof - |
44872 | 909 |
from chinese_remainder_nat [OF fin cop] |
66380 | 910 |
obtain y where one: "(\<forall>i\<in>A. [y = u i] (mod m i))" |
31719 | 911 |
by blast |
61954 | 912 |
let ?x = "y mod (\<Prod>i\<in>A. m i)" |
913 |
from fin nz have prodnz: "(\<Prod>i\<in>A. m i) \<noteq> 0" |
|
31719 | 914 |
by auto |
61954 | 915 |
then have less: "?x < (\<Prod>i\<in>A. m i)" |
31719 | 916 |
by auto |
66380 | 917 |
have cong: "\<forall>i\<in>A. [?x = u i] (mod m i)" |
31719 | 918 |
apply auto |
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
919 |
apply (rule cong_trans_nat) |
66380 | 920 |
prefer 2 |
31719 | 921 |
using one apply auto |
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
922 |
apply (rule cong_dvd_modulus_nat) |
66380 | 923 |
apply (rule cong_mod_nat) |
31719 | 924 |
using prodnz apply auto |
59010 | 925 |
apply rule |
66380 | 926 |
apply (rule fin) |
31719 | 927 |
apply assumption |
928 |
done |
|
66380 | 929 |
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" |
930 |
proof clarify |
|
31719 | 931 |
fix z |
61954 | 932 |
assume zless: "z < (\<Prod>i\<in>A. m i)" |
66380 | 933 |
assume zcong: "(\<forall>i\<in>A. [z = u i] (mod m i))" |
934 |
have "\<forall>i\<in>A. [?x = z] (mod m i)" |
|
44872 | 935 |
apply clarify |
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
936 |
apply (rule cong_trans_nat) |
31719 | 937 |
using cong apply (erule bspec) |
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
938 |
apply (rule cong_sym_nat) |
31719 | 939 |
using zcong apply auto |
940 |
done |
|
61954 | 941 |
with fin cop have "[?x = z] (mod (\<Prod>i\<in>A. m i))" |
44872 | 942 |
apply (intro coprime_cong_prod_nat) |
66380 | 943 |
apply auto |
44872 | 944 |
done |
31719 | 945 |
with zless less show "z = ?x" |
31952
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents:
31792
diff
changeset
|
946 |
apply (intro cong_less_modulus_unique_nat) |
66380 | 947 |
apply auto |
948 |
apply (erule cong_sym_nat) |
|
31719 | 949 |
done |
44872 | 950 |
qed |
66380 | 951 |
from less cong unique show ?thesis |
952 |
by blast |
|
31719 | 953 |
qed |
954 |
||
955 |
end |