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