author | berghofe |
Thu, 06 Oct 2005 10:13:34 +0200 | |
changeset 17771 | 1e07f6ab3118 |
parent 16663 | 13e9c402308b |
child 18369 | 694ea14ab4f2 |
permissions | -rw-r--r-- |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
1 |
(* Title: HOL/NumberTheory/IntPrimes.thy |
9508
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
paulson
parents:
diff
changeset
|
2 |
ID: $Id$ |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
3 |
Author: Thomas M. Rasmussen |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
4 |
Copyright 2000 University of Cambridge |
13833 | 5 |
|
6 |
Changes by Jeremy Avigad, 2003/02/21: |
|
7 |
Repaired definition of zprime_def, added "0 <= m &" |
|
8 |
Added lemma zgcd_geq_zero |
|
9 |
Repaired proof of zprime_imp_zrelprime |
|
9508
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
paulson
parents:
diff
changeset
|
10 |
*) |
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
paulson
parents:
diff
changeset
|
11 |
|
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
12 |
header {* Divisibility and prime numbers (on integers) *} |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
13 |
|
16417 | 14 |
theory IntPrimes imports Primes begin |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
15 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
16 |
text {* |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
17 |
The @{text dvd} relation, GCD, Euclid's extended algorithm, primes, |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
18 |
congruences (all on the Integers). Comparable to theory @{text |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
19 |
Primes}, but @{text dvd} is included here as it is not present in |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
20 |
main HOL. Also includes extended GCD and congruences not present in |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
21 |
@{text Primes}. |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
22 |
*} |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
23 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
24 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
25 |
subsection {* Definitions *} |
9508
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
paulson
parents:
diff
changeset
|
26 |
|
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
paulson
parents:
diff
changeset
|
27 |
consts |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
28 |
xzgcda :: "int * int * int * int * int * int * int * int => int * int * int" |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
29 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
30 |
recdef xzgcda |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
31 |
"measure ((\<lambda>(m, n, r', r, s', s, t', t). nat r) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
32 |
:: int * int * int * int *int * int * int * int => nat)" |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
33 |
"xzgcda (m, n, r', r, s', s, t', t) = |
13833 | 34 |
(if r \<le> 0 then (r', s', t') |
35 |
else xzgcda (m, n, r, r' mod r, |
|
36 |
s, s' - (r' div r) * s, |
|
37 |
t, t' - (r' div r) * t))" |
|
9508
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
paulson
parents:
diff
changeset
|
38 |
|
9943 | 39 |
constdefs |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
40 |
zgcd :: "int * int => int" |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
41 |
"zgcd == \<lambda>(x,y). int (gcd (nat (abs x), nat (abs y)))" |
9943 | 42 |
|
16663 | 43 |
zprime :: "int \<Rightarrow> bool" |
44 |
"zprime p == 1 < p \<and> (\<forall>m. 0 <= m & m dvd p --> m = 1 \<or> m = p)" |
|
13833 | 45 |
|
46 |
xzgcd :: "int => int => int * int * int" |
|
47 |
"xzgcd m n == xzgcda (m, n, m, n, 1, 0, 0, 1)" |
|
48 |
||
49 |
zcong :: "int => int => int => bool" ("(1[_ = _] '(mod _'))") |
|
50 |
"[a = b] (mod m) == m dvd (a - b)" |
|
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
51 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
52 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
53 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
54 |
text {* \medskip @{term gcd} lemmas *} |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
55 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
56 |
lemma gcd_add1_eq: "gcd (m + k, k) = gcd (m + k, m)" |
13833 | 57 |
by (simp add: gcd_commute) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
58 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
59 |
lemma gcd_diff2: "m \<le> n ==> gcd (n, n - m) = gcd (n, m)" |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
60 |
apply (subgoal_tac "n = m + (n - m)") |
13833 | 61 |
apply (erule ssubst, rule gcd_add1_eq, simp) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
62 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
63 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
64 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
65 |
subsection {* Euclid's Algorithm and GCD *} |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
66 |
|
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
67 |
lemma zgcd_0 [simp]: "zgcd (m, 0) = abs m" |
15003 | 68 |
by (simp add: zgcd_def abs_if) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
69 |
|
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
70 |
lemma zgcd_0_left [simp]: "zgcd (0, m) = abs m" |
15003 | 71 |
by (simp add: zgcd_def abs_if) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
72 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
73 |
lemma zgcd_zminus [simp]: "zgcd (-m, n) = zgcd (m, n)" |
13833 | 74 |
by (simp add: zgcd_def) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
75 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
76 |
lemma zgcd_zminus2 [simp]: "zgcd (m, -n) = zgcd (m, n)" |
13833 | 77 |
by (simp add: zgcd_def) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
78 |
|
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
79 |
lemma zgcd_non_0: "0 < n ==> zgcd (m, n) = zgcd (n, m mod n)" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
80 |
apply (frule_tac b = n and a = m in pos_mod_sign) |
15003 | 81 |
apply (simp del: pos_mod_sign add: zgcd_def abs_if nat_mod_distrib) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
82 |
apply (auto simp add: gcd_non_0 nat_mod_distrib [symmetric] zmod_zminus1_eq_if) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
83 |
apply (frule_tac a = m in pos_mod_bound) |
13833 | 84 |
apply (simp del: pos_mod_bound add: nat_diff_distrib gcd_diff2 nat_le_eq_zle) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
85 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
86 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
87 |
lemma zgcd_eq: "zgcd (m, n) = zgcd (n, m mod n)" |
13183 | 88 |
apply (case_tac "n = 0", simp add: DIVISION_BY_ZERO) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
89 |
apply (auto simp add: linorder_neq_iff zgcd_non_0) |
13833 | 90 |
apply (cut_tac m = "-m" and n = "-n" in zgcd_non_0, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
91 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
92 |
|
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
93 |
lemma zgcd_1 [simp]: "zgcd (m, 1) = 1" |
15003 | 94 |
by (simp add: zgcd_def abs_if) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
95 |
|
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
96 |
lemma zgcd_0_1_iff [simp]: "(zgcd (0, m) = 1) = (abs m = 1)" |
15003 | 97 |
by (simp add: zgcd_def abs_if) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
98 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
99 |
lemma zgcd_zdvd1 [iff]: "zgcd (m, n) dvd m" |
15003 | 100 |
by (simp add: zgcd_def abs_if int_dvd_iff) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
101 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
102 |
lemma zgcd_zdvd2 [iff]: "zgcd (m, n) dvd n" |
15003 | 103 |
by (simp add: zgcd_def abs_if int_dvd_iff) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
104 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
105 |
lemma zgcd_greatest_iff: "k dvd zgcd (m, n) = (k dvd m \<and> k dvd n)" |
15003 | 106 |
by (simp add: zgcd_def abs_if int_dvd_iff dvd_int_iff nat_dvd_iff) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
107 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
108 |
lemma zgcd_commute: "zgcd (m, n) = zgcd (n, m)" |
13833 | 109 |
by (simp add: zgcd_def gcd_commute) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
110 |
|
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
111 |
lemma zgcd_1_left [simp]: "zgcd (1, m) = 1" |
13833 | 112 |
by (simp add: zgcd_def gcd_1_left) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
113 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
114 |
lemma zgcd_assoc: "zgcd (zgcd (k, m), n) = zgcd (k, zgcd (m, n))" |
13833 | 115 |
by (simp add: zgcd_def gcd_assoc) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
116 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
117 |
lemma zgcd_left_commute: "zgcd (k, zgcd (m, n)) = zgcd (m, zgcd (k, n))" |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
118 |
apply (rule zgcd_commute [THEN trans]) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
119 |
apply (rule zgcd_assoc [THEN trans]) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
120 |
apply (rule zgcd_commute [THEN arg_cong]) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
121 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
122 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
123 |
lemmas zgcd_ac = zgcd_assoc zgcd_commute zgcd_left_commute |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
124 |
-- {* addition is an AC-operator *} |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
125 |
|
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
126 |
lemma zgcd_zmult_distrib2: "0 \<le> k ==> k * zgcd (m, n) = zgcd (k * m, k * n)" |
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
127 |
by (simp del: minus_mult_right [symmetric] |
15003 | 128 |
add: minus_mult_right nat_mult_distrib zgcd_def abs_if |
14353
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents:
14271
diff
changeset
|
129 |
mult_less_0_iff gcd_mult_distrib2 [symmetric] zmult_int [symmetric]) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
130 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
131 |
lemma zgcd_zmult_distrib2_abs: "zgcd (k * m, k * n) = abs k * zgcd (m, n)" |
15003 | 132 |
by (simp add: abs_if zgcd_zmult_distrib2) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
133 |
|
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
134 |
lemma zgcd_self [simp]: "0 \<le> m ==> zgcd (m, m) = m" |
13833 | 135 |
by (cut_tac k = m and m = 1 and n = 1 in zgcd_zmult_distrib2, simp_all) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
136 |
|
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
137 |
lemma zgcd_zmult_eq_self [simp]: "0 \<le> k ==> zgcd (k, k * n) = k" |
13833 | 138 |
by (cut_tac k = k and m = 1 and n = n in zgcd_zmult_distrib2, simp_all) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
139 |
|
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
140 |
lemma zgcd_zmult_eq_self2 [simp]: "0 \<le> k ==> zgcd (k * n, k) = k" |
13833 | 141 |
by (cut_tac k = k and m = n and n = 1 in zgcd_zmult_distrib2, simp_all) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
142 |
|
13833 | 143 |
lemma zrelprime_zdvd_zmult_aux: |
144 |
"zgcd (n, k) = 1 ==> k dvd m * n ==> 0 \<le> m ==> k dvd m" |
|
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
145 |
apply (subgoal_tac "m = zgcd (m * n, m * k)") |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
146 |
apply (erule ssubst, rule zgcd_greatest_iff [THEN iffD2]) |
14353
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents:
14271
diff
changeset
|
147 |
apply (simp_all add: zgcd_zmult_distrib2 [symmetric] zero_le_mult_iff) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
148 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
149 |
|
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
150 |
lemma zrelprime_zdvd_zmult: "zgcd (n, k) = 1 ==> k dvd m * n ==> k dvd m" |
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
151 |
apply (case_tac "0 \<le> m") |
13524 | 152 |
apply (blast intro: zrelprime_zdvd_zmult_aux) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
153 |
apply (subgoal_tac "k dvd -m") |
13833 | 154 |
apply (rule_tac [2] zrelprime_zdvd_zmult_aux, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
155 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
156 |
|
13833 | 157 |
lemma zgcd_geq_zero: "0 <= zgcd(x,y)" |
158 |
by (auto simp add: zgcd_def) |
|
159 |
||
13837
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents:
13833
diff
changeset
|
160 |
text{*This is merely a sanity check on zprime, since the previous version |
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents:
13833
diff
changeset
|
161 |
denoted the empty set.*} |
16663 | 162 |
lemma "zprime 2" |
13837
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents:
13833
diff
changeset
|
163 |
apply (auto simp add: zprime_def) |
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents:
13833
diff
changeset
|
164 |
apply (frule zdvd_imp_le, simp) |
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents:
13833
diff
changeset
|
165 |
apply (auto simp add: order_le_less dvd_def) |
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents:
13833
diff
changeset
|
166 |
done |
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents:
13833
diff
changeset
|
167 |
|
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
168 |
lemma zprime_imp_zrelprime: |
16663 | 169 |
"zprime p ==> \<not> p dvd n ==> zgcd (n, p) = 1" |
13833 | 170 |
apply (auto simp add: zprime_def) |
171 |
apply (drule_tac x = "zgcd(n, p)" in allE) |
|
172 |
apply (auto simp add: zgcd_zdvd2 [of n p] zgcd_geq_zero) |
|
173 |
apply (insert zgcd_zdvd1 [of n p], auto) |
|
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
174 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
175 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
176 |
lemma zless_zprime_imp_zrelprime: |
16663 | 177 |
"zprime p ==> 0 < n ==> n < p ==> zgcd (n, p) = 1" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
178 |
apply (erule zprime_imp_zrelprime) |
13833 | 179 |
apply (erule zdvd_not_zless, assumption) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
180 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
181 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
182 |
lemma zprime_zdvd_zmult: |
16663 | 183 |
"0 \<le> (m::int) ==> zprime p ==> p dvd m * n ==> p dvd m \<or> p dvd n" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
184 |
apply safe |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
185 |
apply (rule zrelprime_zdvd_zmult) |
13833 | 186 |
apply (rule zprime_imp_zrelprime, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
187 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
188 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
189 |
lemma zgcd_zadd_zmult [simp]: "zgcd (m + n * k, n) = zgcd (m, n)" |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
190 |
apply (rule zgcd_eq [THEN trans]) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
191 |
apply (simp add: zmod_zadd1_eq) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
192 |
apply (rule zgcd_eq [symmetric]) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
193 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
194 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
195 |
lemma zgcd_zdvd_zgcd_zmult: "zgcd (m, n) dvd zgcd (k * m, n)" |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
196 |
apply (simp add: zgcd_greatest_iff) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
197 |
apply (blast intro: zdvd_trans) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
198 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
199 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
200 |
lemma zgcd_zmult_zdvd_zgcd: |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
201 |
"zgcd (k, n) = 1 ==> zgcd (k * m, n) dvd zgcd (m, n)" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
202 |
apply (simp add: zgcd_greatest_iff) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
203 |
apply (rule_tac n = k in zrelprime_zdvd_zmult) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
204 |
prefer 2 |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
205 |
apply (simp add: zmult_commute) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
206 |
apply (subgoal_tac "zgcd (k, zgcd (k * m, n)) = zgcd (k * m, zgcd (k, n))") |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
207 |
apply simp |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
208 |
apply (simp (no_asm) add: zgcd_ac) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
209 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
210 |
|
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
211 |
lemma zgcd_zmult_cancel: "zgcd (k, n) = 1 ==> zgcd (k * m, n) = zgcd (m, n)" |
13833 | 212 |
by (simp add: zgcd_def nat_abs_mult_distrib gcd_mult_cancel) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
213 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
214 |
lemma zgcd_zgcd_zmult: |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
215 |
"zgcd (k, m) = 1 ==> zgcd (n, m) = 1 ==> zgcd (k * n, m) = 1" |
13833 | 216 |
by (simp add: zgcd_zmult_cancel) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
217 |
|
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
218 |
lemma zdvd_iff_zgcd: "0 < m ==> (m dvd n) = (zgcd (n, m) = m)" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
219 |
apply safe |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
220 |
apply (rule_tac [2] n = "zgcd (n, m)" in zdvd_trans) |
13833 | 221 |
apply (rule_tac [3] zgcd_zdvd1, simp_all) |
222 |
apply (unfold dvd_def, auto) |
|
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
223 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
224 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
225 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
226 |
subsection {* Congruences *} |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
227 |
|
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
228 |
lemma zcong_1 [simp]: "[a = b] (mod 1)" |
13833 | 229 |
by (unfold zcong_def, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
230 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
231 |
lemma zcong_refl [simp]: "[k = k] (mod m)" |
13833 | 232 |
by (unfold zcong_def, auto) |
9508
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
paulson
parents:
diff
changeset
|
233 |
|
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
234 |
lemma zcong_sym: "[a = b] (mod m) = [b = a] (mod m)" |
13833 | 235 |
apply (unfold zcong_def dvd_def, auto) |
236 |
apply (rule_tac [!] x = "-k" in exI, auto) |
|
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
237 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
238 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
239 |
lemma zcong_zadd: |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
240 |
"[a = b] (mod m) ==> [c = d] (mod m) ==> [a + c = b + d] (mod m)" |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
241 |
apply (unfold zcong_def) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
242 |
apply (rule_tac s = "(a - b) + (c - d)" in subst) |
13833 | 243 |
apply (rule_tac [2] zdvd_zadd, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
244 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
245 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
246 |
lemma zcong_zdiff: |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
247 |
"[a = b] (mod m) ==> [c = d] (mod m) ==> [a - c = b - d] (mod m)" |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
248 |
apply (unfold zcong_def) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
249 |
apply (rule_tac s = "(a - b) - (c - d)" in subst) |
13833 | 250 |
apply (rule_tac [2] zdvd_zdiff, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
251 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
252 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
253 |
lemma zcong_trans: |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
254 |
"[a = b] (mod m) ==> [b = c] (mod m) ==> [a = c] (mod m)" |
13833 | 255 |
apply (unfold zcong_def dvd_def, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
256 |
apply (rule_tac x = "k + ka" in exI) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
257 |
apply (simp add: zadd_ac zadd_zmult_distrib2) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
258 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
259 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
260 |
lemma zcong_zmult: |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
261 |
"[a = b] (mod m) ==> [c = d] (mod m) ==> [a * c = b * d] (mod m)" |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
262 |
apply (rule_tac b = "b * c" in zcong_trans) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
263 |
apply (unfold zcong_def) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
264 |
apply (rule_tac s = "c * (a - b)" in subst) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
265 |
apply (rule_tac [3] s = "b * (c - d)" in subst) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
266 |
prefer 4 |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
267 |
apply (blast intro: zdvd_zmult) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
268 |
prefer 2 |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
269 |
apply (blast intro: zdvd_zmult) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
270 |
apply (simp_all add: zdiff_zmult_distrib2 zmult_commute) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
271 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
272 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
273 |
lemma zcong_scalar: "[a = b] (mod m) ==> [a * k = b * k] (mod m)" |
13833 | 274 |
by (rule zcong_zmult, simp_all) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
275 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
276 |
lemma zcong_scalar2: "[a = b] (mod m) ==> [k * a = k * b] (mod m)" |
13833 | 277 |
by (rule zcong_zmult, simp_all) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
278 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
279 |
lemma zcong_zmult_self: "[a * m = b * m] (mod m)" |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
280 |
apply (unfold zcong_def) |
13833 | 281 |
apply (rule zdvd_zdiff, simp_all) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
282 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
283 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
284 |
lemma zcong_square: |
16663 | 285 |
"[| zprime p; 0 < a; [a * a = 1] (mod p)|] |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
286 |
==> [a = 1] (mod p) \<or> [a = p - 1] (mod p)" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
287 |
apply (unfold zcong_def) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
288 |
apply (rule zprime_zdvd_zmult) |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
289 |
apply (rule_tac [3] s = "a * a - 1 + p * (1 - a)" in subst) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
290 |
prefer 4 |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
291 |
apply (simp add: zdvd_reduce) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
292 |
apply (simp_all add: zdiff_zmult_distrib zmult_commute zdiff_zmult_distrib2) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
293 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
294 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
295 |
lemma zcong_cancel: |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
296 |
"0 \<le> m ==> |
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
297 |
zgcd (k, m) = 1 ==> [a * k = b * k] (mod m) = [a = b] (mod m)" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
298 |
apply safe |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
299 |
prefer 2 |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
300 |
apply (blast intro: zcong_scalar) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
301 |
apply (case_tac "b < a") |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
302 |
prefer 2 |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
303 |
apply (subst zcong_sym) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
304 |
apply (unfold zcong_def) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
305 |
apply (rule_tac [!] zrelprime_zdvd_zmult) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
306 |
apply (simp_all add: zdiff_zmult_distrib) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
307 |
apply (subgoal_tac "m dvd (-(a * k - b * k))") |
14271 | 308 |
apply simp |
13833 | 309 |
apply (subst zdvd_zminus_iff, assumption) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
310 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
311 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
312 |
lemma zcong_cancel2: |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
313 |
"0 \<le> m ==> |
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
314 |
zgcd (k, m) = 1 ==> [k * a = k * b] (mod m) = [a = b] (mod m)" |
13833 | 315 |
by (simp add: zmult_commute zcong_cancel) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
316 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
317 |
lemma zcong_zgcd_zmult_zmod: |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
318 |
"[a = b] (mod m) ==> [a = b] (mod n) ==> zgcd (m, n) = 1 |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
319 |
==> [a = b] (mod m * n)" |
13833 | 320 |
apply (unfold zcong_def dvd_def, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
321 |
apply (subgoal_tac "m dvd n * ka") |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
322 |
apply (subgoal_tac "m dvd ka") |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
323 |
apply (case_tac [2] "0 \<le> ka") |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
324 |
prefer 3 |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
325 |
apply (subst zdvd_zminus_iff [symmetric]) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
326 |
apply (rule_tac n = n in zrelprime_zdvd_zmult) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
327 |
apply (simp add: zgcd_commute) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
328 |
apply (simp add: zmult_commute zdvd_zminus_iff) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
329 |
prefer 2 |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
330 |
apply (rule_tac n = n in zrelprime_zdvd_zmult) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
331 |
apply (simp add: zgcd_commute) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
332 |
apply (simp add: zmult_commute) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
333 |
apply (auto simp add: dvd_def) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
334 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
335 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
336 |
lemma zcong_zless_imp_eq: |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
337 |
"0 \<le> a ==> |
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
338 |
a < m ==> 0 \<le> b ==> b < m ==> [a = b] (mod m) ==> a = b" |
13833 | 339 |
apply (unfold zcong_def dvd_def, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
340 |
apply (drule_tac f = "\<lambda>z. z mod m" in arg_cong) |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14353
diff
changeset
|
341 |
apply (cut_tac x = a and y = b in linorder_less_linear, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
342 |
apply (subgoal_tac [2] "(a - b) mod m = a - b") |
13833 | 343 |
apply (rule_tac [3] mod_pos_pos_trivial, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
344 |
apply (subgoal_tac "(m + (a - b)) mod m = m + (a - b)") |
13833 | 345 |
apply (rule_tac [2] mod_pos_pos_trivial, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
346 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
347 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
348 |
lemma zcong_square_zless: |
16663 | 349 |
"zprime p ==> 0 < a ==> a < p ==> |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
350 |
[a * a = 1] (mod p) ==> a = 1 \<or> a = p - 1" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
351 |
apply (cut_tac p = p and a = a in zcong_square) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
352 |
apply (simp add: zprime_def) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
353 |
apply (auto intro: zcong_zless_imp_eq) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
354 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
355 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
356 |
lemma zcong_not: |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
357 |
"0 < a ==> a < m ==> 0 < b ==> b < a ==> \<not> [a = b] (mod m)" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
358 |
apply (unfold zcong_def) |
13833 | 359 |
apply (rule zdvd_not_zless, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
360 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
361 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
362 |
lemma zcong_zless_0: |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
363 |
"0 \<le> a ==> a < m ==> [a = 0] (mod m) ==> a = 0" |
13833 | 364 |
apply (unfold zcong_def dvd_def, auto) |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
365 |
apply (subgoal_tac "0 < m") |
14353
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents:
14271
diff
changeset
|
366 |
apply (simp add: zero_le_mult_iff) |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
367 |
apply (subgoal_tac "m * k < m * 1") |
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
368 |
apply (drule mult_less_cancel_left [THEN iffD1]) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
369 |
apply (auto simp add: linorder_neq_iff) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
370 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
371 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
372 |
lemma zcong_zless_unique: |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
373 |
"0 < m ==> (\<exists>!b. 0 \<le> b \<and> b < m \<and> [a = b] (mod m))" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
374 |
apply auto |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
375 |
apply (subgoal_tac [2] "[b = y] (mod m)") |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
376 |
apply (case_tac [2] "b = 0") |
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
377 |
apply (case_tac [3] "y = 0") |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
378 |
apply (auto intro: zcong_trans zcong_zless_0 zcong_zless_imp_eq order_less_le |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
379 |
simp add: zcong_sym) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
380 |
apply (unfold zcong_def dvd_def) |
13833 | 381 |
apply (rule_tac x = "a mod m" in exI, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
382 |
apply (rule_tac x = "-(a div m)" in exI) |
14271 | 383 |
apply (simp add: diff_eq_eq eq_diff_eq add_commute) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
384 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
385 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
386 |
lemma zcong_iff_lin: "([a = b] (mod m)) = (\<exists>k. b = a + m * k)" |
13833 | 387 |
apply (unfold zcong_def dvd_def, auto) |
388 |
apply (rule_tac [!] x = "-k" in exI, auto) |
|
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
389 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
390 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
391 |
lemma zgcd_zcong_zgcd: |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
392 |
"0 < m ==> |
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
393 |
zgcd (a, m) = 1 ==> [a = b] (mod m) ==> zgcd (b, m) = 1" |
13833 | 394 |
by (auto simp add: zcong_iff_lin) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
395 |
|
13833 | 396 |
lemma zcong_zmod_aux: |
397 |
"a - b = (m::int) * (a div m - b div m) + (a mod m - b mod m)" |
|
14271 | 398 |
by(simp add: zdiff_zmult_distrib2 add_diff_eq eq_diff_eq add_ac) |
13517 | 399 |
|
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
400 |
lemma zcong_zmod: "[a = b] (mod m) = [a mod m = b mod m] (mod m)" |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
401 |
apply (unfold zcong_def) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
402 |
apply (rule_tac t = "a - b" in ssubst) |
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13837
diff
changeset
|
403 |
apply (rule_tac m = m in zcong_zmod_aux) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
404 |
apply (rule trans) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
405 |
apply (rule_tac [2] k = m and m = "a div m - b div m" in zdvd_reduce) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
406 |
apply (simp add: zadd_commute) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
407 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
408 |
|
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
409 |
lemma zcong_zmod_eq: "0 < m ==> [a = b] (mod m) = (a mod m = b mod m)" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
410 |
apply auto |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
411 |
apply (rule_tac m = m in zcong_zless_imp_eq) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
412 |
prefer 5 |
13833 | 413 |
apply (subst zcong_zmod [symmetric], simp_all) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
414 |
apply (unfold zcong_def dvd_def) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
415 |
apply (rule_tac x = "a div m - b div m" in exI) |
13833 | 416 |
apply (rule_tac m1 = m in zcong_zmod_aux [THEN trans], auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
417 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
418 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
419 |
lemma zcong_zminus [iff]: "[a = b] (mod -m) = [a = b] (mod m)" |
13833 | 420 |
by (auto simp add: zcong_def) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
421 |
|
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
422 |
lemma zcong_zero [iff]: "[a = b] (mod 0) = (a = b)" |
13833 | 423 |
by (auto simp add: zcong_def) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
424 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
425 |
lemma "[a = b] (mod m) = (a mod m = b mod m)" |
13183 | 426 |
apply (case_tac "m = 0", simp add: DIVISION_BY_ZERO) |
13193 | 427 |
apply (simp add: linorder_neq_iff) |
428 |
apply (erule disjE) |
|
429 |
prefer 2 apply (simp add: zcong_zmod_eq) |
|
430 |
txt{*Remainding case: @{term "m<0"}*} |
|
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
431 |
apply (rule_tac t = m in zminus_zminus [THEN subst]) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
432 |
apply (subst zcong_zminus) |
13833 | 433 |
apply (subst zcong_zmod_eq, arith) |
13193 | 434 |
apply (frule neg_mod_bound [of _ a], frule neg_mod_bound [of _ b]) |
13788 | 435 |
apply (simp add: zmod_zminus2_eq_if del: neg_mod_bound) |
13193 | 436 |
done |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
437 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
438 |
subsection {* Modulo *} |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
439 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
440 |
lemma zmod_zdvd_zmod: |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
441 |
"0 < (m::int) ==> m dvd b ==> (a mod b mod m) = (a mod m)" |
13833 | 442 |
apply (unfold dvd_def, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
443 |
apply (subst zcong_zmod_eq [symmetric]) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
444 |
prefer 2 |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
445 |
apply (subst zcong_iff_lin) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
446 |
apply (rule_tac x = "k * (a div (m * k))" in exI) |
13833 | 447 |
apply (simp add:zmult_assoc [symmetric], assumption) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
448 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
449 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
450 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
451 |
subsection {* Extended GCD *} |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
452 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
453 |
declare xzgcda.simps [simp del] |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
454 |
|
13524 | 455 |
lemma xzgcd_correct_aux1: |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
456 |
"zgcd (r', r) = k --> 0 < r --> |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
457 |
(\<exists>sn tn. xzgcda (m, n, r', r, s', s, t', t) = (k, sn, tn))" |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
458 |
apply (rule_tac u = m and v = n and w = r' and x = r and y = s' and |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
459 |
z = s and aa = t' and ab = t in xzgcda.induct) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
460 |
apply (subst zgcd_eq) |
13833 | 461 |
apply (subst xzgcda.simps, auto) |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
462 |
apply (case_tac "r' mod r = 0") |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
463 |
prefer 2 |
13833 | 464 |
apply (frule_tac a = "r'" in pos_mod_sign, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
465 |
apply (rule exI) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
466 |
apply (rule exI) |
13833 | 467 |
apply (subst xzgcda.simps, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
468 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
469 |
|
13524 | 470 |
lemma xzgcd_correct_aux2: |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
471 |
"(\<exists>sn tn. xzgcda (m, n, r', r, s', s, t', t) = (k, sn, tn)) --> 0 < r --> |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
472 |
zgcd (r', r) = k" |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
473 |
apply (rule_tac u = m and v = n and w = r' and x = r and y = s' and |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
474 |
z = s and aa = t' and ab = t in xzgcda.induct) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
475 |
apply (subst zgcd_eq) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
476 |
apply (subst xzgcda.simps) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
477 |
apply (auto simp add: linorder_not_le) |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
478 |
apply (case_tac "r' mod r = 0") |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
479 |
prefer 2 |
13833 | 480 |
apply (frule_tac a = "r'" in pos_mod_sign, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
481 |
apply (erule_tac P = "xzgcda ?u = ?v" in rev_mp) |
13833 | 482 |
apply (subst xzgcda.simps, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
483 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
484 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
485 |
lemma xzgcd_correct: |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
486 |
"0 < n ==> (zgcd (m, n) = k) = (\<exists>s t. xzgcd m n = (k, s, t))" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
487 |
apply (unfold xzgcd_def) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
488 |
apply (rule iffI) |
13524 | 489 |
apply (rule_tac [2] xzgcd_correct_aux2 [THEN mp, THEN mp]) |
13833 | 490 |
apply (rule xzgcd_correct_aux1 [THEN mp, THEN mp], auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
491 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
492 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
493 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
494 |
text {* \medskip @{term xzgcd} linear *} |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
495 |
|
13524 | 496 |
lemma xzgcda_linear_aux1: |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
497 |
"(a - r * b) * m + (c - r * d) * (n::int) = |
13833 | 498 |
(a * m + c * n) - r * (b * m + d * n)" |
499 |
by (simp add: zdiff_zmult_distrib zadd_zmult_distrib2 zmult_assoc) |
|
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
500 |
|
13524 | 501 |
lemma xzgcda_linear_aux2: |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
502 |
"r' = s' * m + t' * n ==> r = s * m + t * n |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
503 |
==> (r' mod r) = (s' - (r' div r) * s) * m + (t' - (r' div r) * t) * (n::int)" |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
504 |
apply (rule trans) |
13524 | 505 |
apply (rule_tac [2] xzgcda_linear_aux1 [symmetric]) |
14271 | 506 |
apply (simp add: eq_diff_eq mult_commute) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
507 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
508 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
509 |
lemma order_le_neq_implies_less: "(x::'a::order) \<le> y ==> x \<noteq> y ==> x < y" |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
510 |
by (rule iffD2 [OF order_less_le conjI]) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
511 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
512 |
lemma xzgcda_linear [rule_format]: |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
513 |
"0 < r --> xzgcda (m, n, r', r, s', s, t', t) = (rn, sn, tn) --> |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
514 |
r' = s' * m + t' * n --> r = s * m + t * n --> rn = sn * m + tn * n" |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
515 |
apply (rule_tac u = m and v = n and w = r' and x = r and y = s' and |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
516 |
z = s and aa = t' and ab = t in xzgcda.induct) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
517 |
apply (subst xzgcda.simps) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
518 |
apply (simp (no_asm)) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
519 |
apply (rule impI)+ |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
520 |
apply (case_tac "r' mod r = 0") |
13833 | 521 |
apply (simp add: xzgcda.simps, clarify) |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
522 |
apply (subgoal_tac "0 < r' mod r") |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
523 |
apply (rule_tac [2] order_le_neq_implies_less) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
524 |
apply (rule_tac [2] pos_mod_sign) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
525 |
apply (cut_tac m = m and n = n and r' = r' and r = r and s' = s' and |
13833 | 526 |
s = s and t' = t' and t = t in xzgcda_linear_aux2, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
527 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
528 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
529 |
lemma xzgcd_linear: |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
530 |
"0 < n ==> xzgcd m n = (r, s, t) ==> r = s * m + t * n" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
531 |
apply (unfold xzgcd_def) |
13837
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents:
13833
diff
changeset
|
532 |
apply (erule xzgcda_linear, assumption, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
533 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
534 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
535 |
lemma zgcd_ex_linear: |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
536 |
"0 < n ==> zgcd (m, n) = k ==> (\<exists>s t. k = s * m + t * n)" |
13833 | 537 |
apply (simp add: xzgcd_correct, safe) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
538 |
apply (rule exI)+ |
13833 | 539 |
apply (erule xzgcd_linear, auto) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
540 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
541 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
542 |
lemma zcong_lineq_ex: |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
543 |
"0 < n ==> zgcd (a, n) = 1 ==> \<exists>x. [a * x = 1] (mod n)" |
13833 | 544 |
apply (cut_tac m = a and n = n and k = 1 in zgcd_ex_linear, safe) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
545 |
apply (rule_tac x = s in exI) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
546 |
apply (rule_tac b = "s * a + t * n" in zcong_trans) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
547 |
prefer 2 |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
548 |
apply simp |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
549 |
apply (unfold zcong_def) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
550 |
apply (simp (no_asm) add: zmult_commute zdvd_zminus_iff) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
551 |
done |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
552 |
|
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
553 |
lemma zcong_lineq_unique: |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
554 |
"0 < n ==> |
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
555 |
zgcd (a, n) = 1 ==> \<exists>!x. 0 \<le> x \<and> x < n \<and> [a * x = b] (mod n)" |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
556 |
apply auto |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
557 |
apply (rule_tac [2] zcong_zless_imp_eq) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
558 |
apply (tactic {* stac (thm "zcong_cancel2" RS sym) 6 *}) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
559 |
apply (rule_tac [8] zcong_trans) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
560 |
apply (simp_all (no_asm_simp)) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
561 |
prefer 2 |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
562 |
apply (simp add: zcong_sym) |
13833 | 563 |
apply (cut_tac a = a and n = n in zcong_lineq_ex, auto) |
564 |
apply (rule_tac x = "x * b mod n" in exI, safe) |
|
13788 | 565 |
apply (simp_all (no_asm_simp)) |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
566 |
apply (subst zcong_zmod) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
567 |
apply (subst zmod_zmult1_eq [symmetric]) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
568 |
apply (subst zcong_zmod [symmetric]) |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
569 |
apply (subgoal_tac "[a * x * b = 1 * b] (mod n)") |
11049
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
570 |
apply (rule_tac [2] zcong_zmult) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
571 |
apply (simp_all add: zmult_assoc) |
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
wenzelm
parents:
10147
diff
changeset
|
572 |
done |
9508
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
paulson
parents:
diff
changeset
|
573 |
|
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
paulson
parents:
diff
changeset
|
574 |
end |