| author | huffman | 
| Wed, 24 Dec 2008 08:16:45 -0800 | |
| changeset 29166 | c23b2d108612 | 
| parent 27651 | 16a26996c30e | 
| child 29412 | 4085a531153d | 
| permissions | -rw-r--r-- | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
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: 
10147diff
changeset | 3 | Author: Thomas M. Rasmussen | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
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: 
10147diff
changeset | 7 | header {* Divisibility and prime numbers (on integers) *}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 8 | |
| 25596 | 9 | theory IntPrimes | 
| 27368 | 10 | imports Main Primes | 
| 25596 | 11 | begin | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 12 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 13 | text {*
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 14 |   The @{text dvd} relation, GCD, Euclid's extended algorithm, primes,
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 15 |   congruences (all on the Integers).  Comparable to theory @{text
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 16 |   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: 
10147diff
changeset | 17 | 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: 
10147diff
changeset | 18 |   @{text Primes}.
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 19 | *} | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 20 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 21 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 22 | subsection {* Definitions *}
 | 
| 9508 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 23 | |
| 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 24 | consts | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 25 | 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: 
10147diff
changeset | 26 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 27 | recdef xzgcda | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 28 | "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: 
10147diff
changeset | 29 | :: int * int * int * int *int * int * int * int => nat)" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 30 | "xzgcda (m, n, r', r, s', s, t', t) = | 
| 13833 | 31 | (if r \<le> 0 then (r', s', t') | 
| 32 | else xzgcda (m, n, r, r' mod r, | |
| 33 | s, s' - (r' div r) * s, | |
| 34 | t, t' - (r' div r) * t))" | |
| 9508 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 35 | |
| 19670 | 36 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19670diff
changeset | 37 | zprime :: "int \<Rightarrow> bool" where | 
| 19670 | 38 | "zprime p = (1 < p \<and> (\<forall>m. 0 <= m & m dvd p --> m = 1 \<or> m = p))" | 
| 13833 | 39 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19670diff
changeset | 40 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19670diff
changeset | 41 | xzgcd :: "int => int => int * int * int" where | 
| 19670 | 42 | "xzgcd m n = xzgcda (m, n, m, n, 1, 0, 0, 1)" | 
| 13833 | 43 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19670diff
changeset | 44 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19670diff
changeset | 45 |   zcong :: "int => int => int => bool"  ("(1[_ = _] '(mod _'))") where
 | 
| 19670 | 46 | "[a = b] (mod m) = (m dvd (a - b))" | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 47 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 48 | subsection {* Euclid's Algorithm and GCD *}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 49 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 50 | |
| 13833 | 51 | lemma zrelprime_zdvd_zmult_aux: | 
| 27556 | 52 | "zgcd n k = 1 ==> k dvd m * n ==> 0 \<le> m ==> k dvd m" | 
| 27569 | 53 | by (metis abs_of_nonneg zdvd_triv_right zgcd_greatest_iff zgcd_zmult_distrib2_abs zmult_1_right) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 54 | |
| 27556 | 55 | lemma zrelprime_zdvd_zmult: "zgcd n k = 1 ==> k dvd m * n ==> k dvd m" | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 56 | apply (case_tac "0 \<le> m") | 
| 13524 | 57 | apply (blast intro: zrelprime_zdvd_zmult_aux) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 58 | apply (subgoal_tac "k dvd -m") | 
| 13833 | 59 | apply (rule_tac [2] zrelprime_zdvd_zmult_aux, auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 60 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 61 | |
| 27556 | 62 | lemma zgcd_geq_zero: "0 <= zgcd x y" | 
| 13833 | 63 | by (auto simp add: zgcd_def) | 
| 64 | ||
| 13837 
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
 paulson parents: 
13833diff
changeset | 65 | 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: 
13833diff
changeset | 66 | denoted the empty set.*} | 
| 16663 | 67 | lemma "zprime 2" | 
| 13837 
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
 paulson parents: 
13833diff
changeset | 68 | apply (auto simp add: zprime_def) | 
| 
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
 paulson parents: 
13833diff
changeset | 69 | apply (frule zdvd_imp_le, simp) | 
| 
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
 paulson parents: 
13833diff
changeset | 70 | apply (auto simp add: order_le_less dvd_def) | 
| 
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
 paulson parents: 
13833diff
changeset | 71 | done | 
| 
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
 paulson parents: 
13833diff
changeset | 72 | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 73 | lemma zprime_imp_zrelprime: | 
| 27556 | 74 | "zprime p ==> \<not> p dvd n ==> zgcd n p = 1" | 
| 13833 | 75 | apply (auto simp add: zprime_def) | 
| 23839 | 76 | apply (metis zgcd_commute zgcd_geq_zero zgcd_zdvd1 zgcd_zdvd2) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 77 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 78 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 79 | lemma zless_zprime_imp_zrelprime: | 
| 27556 | 80 | "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: 
10147diff
changeset | 81 | apply (erule zprime_imp_zrelprime) | 
| 13833 | 82 | apply (erule zdvd_not_zless, assumption) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 83 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 84 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 85 | lemma zprime_zdvd_zmult: | 
| 16663 | 86 | "0 \<le> (m::int) ==> zprime p ==> p dvd m * n ==> p dvd m \<or> p dvd n" | 
| 27569 | 87 | by (metis zgcd_zdvd1 zgcd_zdvd2 zgcd_pos zprime_def zrelprime_dvd_mult) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 88 | |
| 27556 | 89 | lemma zgcd_zadd_zmult [simp]: "zgcd (m + n * k) n = zgcd m n" | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 90 | apply (rule zgcd_eq [THEN trans]) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 91 | apply (simp add: zmod_zadd1_eq) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 92 | apply (rule zgcd_eq [symmetric]) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 93 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 94 | |
| 27556 | 95 | lemma zgcd_zdvd_zgcd_zmult: "zgcd m n dvd zgcd (k * m) n" | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 96 | apply (simp add: zgcd_greatest_iff) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 97 | apply (blast intro: zdvd_trans) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 98 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 99 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 100 | lemma zgcd_zmult_zdvd_zgcd: | 
| 27569 | 101 | "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: 
10147diff
changeset | 102 | apply (simp add: zgcd_greatest_iff) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 103 | apply (rule_tac n = k in zrelprime_zdvd_zmult) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 104 | prefer 2 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 105 | apply (simp add: zmult_commute) | 
| 23839 | 106 | apply (metis zgcd_1 zgcd_commute zgcd_left_commute) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 107 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 108 | |
| 27556 | 109 | lemma zgcd_zmult_cancel: "zgcd k n = 1 ==> zgcd (k * m) n = zgcd m n" | 
| 13833 | 110 | 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: 
10147diff
changeset | 111 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 112 | lemma zgcd_zgcd_zmult: | 
| 27569 | 113 | "zgcd k m = 1 ==> zgcd n m = 1 ==> zgcd (k * n) m = 1" | 
| 13833 | 114 | by (simp add: zgcd_zmult_cancel) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 115 | |
| 27556 | 116 | lemma zdvd_iff_zgcd: "0 < m ==> m dvd n \<longleftrightarrow> zgcd n m = m" | 
| 23839 | 117 | by (metis abs_of_pos zdvd_mult_div_cancel zgcd_0 zgcd_commute zgcd_geq_zero zgcd_zdvd2 zgcd_zmult_eq_self) | 
| 118 | ||
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 119 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 120 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 121 | subsection {* Congruences *}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 122 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 123 | lemma zcong_1 [simp]: "[a = b] (mod 1)" | 
| 13833 | 124 | by (unfold zcong_def, auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 125 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 126 | lemma zcong_refl [simp]: "[k = k] (mod m)" | 
| 13833 | 127 | by (unfold zcong_def, auto) | 
| 9508 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 128 | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 129 | lemma zcong_sym: "[a = b] (mod m) = [b = a] (mod m)" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27569diff
changeset | 130 | unfolding zcong_def minus_diff_eq [of a, symmetric] zdvd_zminus_iff .. | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 131 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 132 | lemma zcong_zadd: | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 133 | "[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: 
10147diff
changeset | 134 | apply (unfold zcong_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 135 | apply (rule_tac s = "(a - b) + (c - d)" in subst) | 
| 13833 | 136 | apply (rule_tac [2] zdvd_zadd, auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 137 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 138 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 139 | lemma zcong_zdiff: | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 140 | "[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: 
10147diff
changeset | 141 | apply (unfold zcong_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 142 | apply (rule_tac s = "(a - b) - (c - d)" in subst) | 
| 13833 | 143 | apply (rule_tac [2] zdvd_zdiff, auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 144 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 145 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 146 | lemma zcong_trans: | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 147 | "[a = b] (mod m) ==> [b = c] (mod m) ==> [a = c] (mod m)" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27569diff
changeset | 148 | unfolding zcong_def | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27569diff
changeset | 149 | apply (auto elim!: dvdE simp add: ring_simps) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27569diff
changeset | 150 | unfolding left_distrib [symmetric] | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27569diff
changeset | 151 | apply (rule dvd_mult dvd_refl)+ | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 152 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 153 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 154 | lemma zcong_zmult: | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 155 | "[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: 
10147diff
changeset | 156 | apply (rule_tac b = "b * c" in zcong_trans) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 157 | apply (unfold zcong_def) | 
| 23839 | 158 | apply (metis zdiff_zmult_distrib2 zdvd_zmult zmult_commute) | 
| 159 | apply (metis zdiff_zmult_distrib2 zdvd_zmult) | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 160 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 161 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 162 | lemma zcong_scalar: "[a = b] (mod m) ==> [a * k = b * k] (mod m)" | 
| 13833 | 163 | by (rule zcong_zmult, simp_all) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 164 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 165 | lemma zcong_scalar2: "[a = b] (mod m) ==> [k * a = k * b] (mod m)" | 
| 13833 | 166 | by (rule zcong_zmult, simp_all) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 167 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 168 | lemma zcong_zmult_self: "[a * m = b * m] (mod m)" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 169 | apply (unfold zcong_def) | 
| 13833 | 170 | apply (rule zdvd_zdiff, simp_all) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 171 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 172 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 173 | lemma zcong_square: | 
| 16663 | 174 | "[| 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: 
11701diff
changeset | 175 | ==> [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: 
10147diff
changeset | 176 | apply (unfold zcong_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 177 | apply (rule zprime_zdvd_zmult) | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 178 | 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: 
10147diff
changeset | 179 | prefer 4 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 180 | apply (simp add: zdvd_reduce) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 181 | 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: 
10147diff
changeset | 182 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 183 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 184 | lemma zcong_cancel: | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 185 | "0 \<le> m ==> | 
| 27556 | 186 | 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: 
10147diff
changeset | 187 | apply safe | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 188 | prefer 2 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 189 | apply (blast intro: zcong_scalar) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 190 | apply (case_tac "b < a") | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 191 | prefer 2 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 192 | apply (subst zcong_sym) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 193 | apply (unfold zcong_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 194 | apply (rule_tac [!] zrelprime_zdvd_zmult) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 195 | apply (simp_all add: zdiff_zmult_distrib) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 196 | apply (subgoal_tac "m dvd (-(a * k - b * k))") | 
| 14271 | 197 | apply simp | 
| 13833 | 198 | apply (subst zdvd_zminus_iff, assumption) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 199 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 200 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 201 | lemma zcong_cancel2: | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 202 | "0 \<le> m ==> | 
| 27556 | 203 | zgcd k m = 1 ==> [k * a = k * b] (mod m) = [a = b] (mod m)" | 
| 13833 | 204 | by (simp add: zmult_commute zcong_cancel) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 205 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 206 | lemma zcong_zgcd_zmult_zmod: | 
| 27556 | 207 | "[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: 
10147diff
changeset | 208 | ==> [a = b] (mod m * n)" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27569diff
changeset | 209 | apply (auto simp add: zcong_def dvd_def) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 210 | apply (subgoal_tac "m dvd n * ka") | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 211 | apply (subgoal_tac "m dvd ka") | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 212 | apply (case_tac [2] "0 \<le> ka") | 
| 23839 | 213 | apply (metis zdvd_mult_div_cancel zdvd_refl zdvd_zminus2_iff zdvd_zmultD2 zgcd_zminus zmult_commute zmult_zminus zrelprime_zdvd_zmult) | 
| 214 | apply (metis IntDiv.zdvd_abs1 abs_of_nonneg zadd_0 zgcd_0_left zgcd_commute zgcd_zadd_zmult zgcd_zdvd_zgcd_zmult zgcd_zmult_distrib2_abs zmult_1_right zmult_commute) | |
| 215 | apply (metis abs_eq_0 int_0_neq_1 mult_le_0_iff zdvd_mono zdvd_mult_cancel zdvd_mult_cancel1 zdvd_refl zdvd_triv_left zdvd_zmult2 zero_le_mult_iff zgcd_greatest_iff zle_anti_sym zle_linear zle_refl zmult_commute zrelprime_zdvd_zmult) | |
| 216 | apply (metis zdvd_triv_left) | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 217 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 218 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 219 | lemma zcong_zless_imp_eq: | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 220 | "0 \<le> a ==> | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 221 | a < m ==> 0 \<le> b ==> b < m ==> [a = b] (mod m) ==> a = b" | 
| 13833 | 222 | apply (unfold zcong_def dvd_def, auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 223 | apply (drule_tac f = "\<lambda>z. z mod m" in arg_cong) | 
| 23839 | 224 | apply (metis diff_add_cancel mod_pos_pos_trivial zadd_0 zadd_commute zmod_eq_0_iff zmod_zadd_right_eq) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 225 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 226 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 227 | lemma zcong_square_zless: | 
| 16663 | 228 | "zprime p ==> 0 < a ==> a < p ==> | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 229 | [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: 
10147diff
changeset | 230 | 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: 
10147diff
changeset | 231 | apply (simp add: zprime_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 232 | apply (auto intro: zcong_zless_imp_eq) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 233 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 234 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 235 | lemma zcong_not: | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 236 | "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: 
10147diff
changeset | 237 | apply (unfold zcong_def) | 
| 13833 | 238 | apply (rule zdvd_not_zless, auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 239 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 240 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 241 | lemma zcong_zless_0: | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 242 | "0 \<le> a ==> a < m ==> [a = 0] (mod m) ==> a = 0" | 
| 13833 | 243 | apply (unfold zcong_def dvd_def, auto) | 
| 23839 | 244 | apply (metis div_pos_pos_trivial linorder_not_less zdiv_zmult_self2 zle_refl zle_trans) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 245 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 246 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 247 | lemma zcong_zless_unique: | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 248 | "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: 
10147diff
changeset | 249 | apply auto | 
| 23839 | 250 | prefer 2 apply (metis zcong_sym zcong_trans zcong_zless_imp_eq) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 251 | apply (unfold zcong_def dvd_def) | 
| 13833 | 252 | apply (rule_tac x = "a mod m" in exI, auto) | 
| 23839 | 253 | apply (metis zmult_div_cancel) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 254 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 255 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 256 | lemma zcong_iff_lin: "([a = b] (mod m)) = (\<exists>k. b = a + m * k)" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27569diff
changeset | 257 | unfolding zcong_def | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27569diff
changeset | 258 | apply (auto elim!: dvdE simp add: ring_simps) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27569diff
changeset | 259 | apply (rule_tac x = "-k" in exI) apply simp | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 260 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 261 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 262 | lemma zgcd_zcong_zgcd: | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 263 | "0 < m ==> | 
| 27556 | 264 | zgcd a m = 1 ==> [a = b] (mod m) ==> zgcd b m = 1" | 
| 13833 | 265 | by (auto simp add: zcong_iff_lin) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 266 | |
| 13833 | 267 | lemma zcong_zmod_aux: | 
| 268 | "a - b = (m::int) * (a div m - b div m) + (a mod m - b mod m)" | |
| 14271 | 269 | by(simp add: zdiff_zmult_distrib2 add_diff_eq eq_diff_eq add_ac) | 
| 13517 | 270 | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 271 | 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: 
10147diff
changeset | 272 | apply (unfold zcong_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 273 | apply (rule_tac t = "a - b" in ssubst) | 
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13837diff
changeset | 274 | apply (rule_tac m = m in zcong_zmod_aux) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 275 | apply (rule trans) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 276 | 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: 
10147diff
changeset | 277 | apply (simp add: zadd_commute) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 278 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 279 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 280 | 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: 
10147diff
changeset | 281 | apply auto | 
| 23839 | 282 | apply (metis pos_mod_conj zcong_zless_imp_eq zcong_zmod) | 
| 283 | apply (metis zcong_refl zcong_zmod) | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 284 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 285 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 286 | lemma zcong_zminus [iff]: "[a = b] (mod -m) = [a = b] (mod m)" | 
| 13833 | 287 | by (auto simp add: zcong_def) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 288 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 289 | lemma zcong_zero [iff]: "[a = b] (mod 0) = (a = b)" | 
| 13833 | 290 | by (auto simp add: zcong_def) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 291 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 292 | lemma "[a = b] (mod m) = (a mod m = b mod m)" | 
| 13183 | 293 | apply (case_tac "m = 0", simp add: DIVISION_BY_ZERO) | 
| 13193 | 294 | apply (simp add: linorder_neq_iff) | 
| 295 | apply (erule disjE) | |
| 296 | prefer 2 apply (simp add: zcong_zmod_eq) | |
| 297 |   txt{*Remainding case: @{term "m<0"}*}
 | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 298 | apply (rule_tac t = m in zminus_zminus [THEN subst]) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 299 | apply (subst zcong_zminus) | 
| 13833 | 300 | apply (subst zcong_zmod_eq, arith) | 
| 13193 | 301 | apply (frule neg_mod_bound [of _ a], frule neg_mod_bound [of _ b]) | 
| 13788 | 302 | apply (simp add: zmod_zminus2_eq_if del: neg_mod_bound) | 
| 13193 | 303 | done | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 304 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 305 | subsection {* Modulo *}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 306 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 307 | lemma zmod_zdvd_zmod: | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 308 | "0 < (m::int) ==> m dvd b ==> (a mod b mod m) = (a mod m)" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27569diff
changeset | 309 | by (rule zmod_zmod_cancel) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 310 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 311 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 312 | subsection {* Extended GCD *}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 313 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 314 | declare xzgcda.simps [simp del] | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 315 | |
| 13524 | 316 | lemma xzgcd_correct_aux1: | 
| 27556 | 317 | "zgcd r' r = k --> 0 < r --> | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 318 | (\<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: 
10147diff
changeset | 319 | 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: 
10147diff
changeset | 320 | 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: 
10147diff
changeset | 321 | apply (subst zgcd_eq) | 
| 13833 | 322 | apply (subst xzgcda.simps, auto) | 
| 24759 | 323 | apply (case_tac "r' mod r = 0") | 
| 324 | prefer 2 | |
| 325 | apply (frule_tac a = "r'" in pos_mod_sign, auto) | |
| 326 | apply (rule exI) | |
| 327 | apply (rule exI) | |
| 328 | apply (subst xzgcda.simps, auto) | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 329 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 330 | |
| 13524 | 331 | lemma xzgcd_correct_aux2: | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 332 | "(\<exists>sn tn. xzgcda (m, n, r', r, s', s, t', t) = (k, sn, tn)) --> 0 < r --> | 
| 27556 | 333 | zgcd r' r = k" | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 334 | 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: 
10147diff
changeset | 335 | 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: 
10147diff
changeset | 336 | apply (subst zgcd_eq) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 337 | apply (subst xzgcda.simps) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 338 | 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: 
11701diff
changeset | 339 | apply (case_tac "r' mod r = 0") | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 340 | prefer 2 | 
| 13833 | 341 | apply (frule_tac a = "r'" in pos_mod_sign, auto) | 
| 23839 | 342 | apply (metis Pair_eq simps zle_refl) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 343 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 344 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 345 | lemma xzgcd_correct: | 
| 27569 | 346 | "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: 
10147diff
changeset | 347 | apply (unfold xzgcd_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 348 | apply (rule iffI) | 
| 13524 | 349 | apply (rule_tac [2] xzgcd_correct_aux2 [THEN mp, THEN mp]) | 
| 13833 | 350 | 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: 
10147diff
changeset | 351 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 352 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 353 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 354 | text {* \medskip @{term xzgcd} linear *}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 355 | |
| 13524 | 356 | lemma xzgcda_linear_aux1: | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 357 | "(a - r * b) * m + (c - r * d) * (n::int) = | 
| 13833 | 358 | (a * m + c * n) - r * (b * m + d * n)" | 
| 359 | 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: 
10147diff
changeset | 360 | |
| 13524 | 361 | lemma xzgcda_linear_aux2: | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 362 | "r' = s' * m + t' * n ==> r = s * m + t * n | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 363 | ==> (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: 
10147diff
changeset | 364 | apply (rule trans) | 
| 13524 | 365 | apply (rule_tac [2] xzgcda_linear_aux1 [symmetric]) | 
| 14271 | 366 | apply (simp add: eq_diff_eq mult_commute) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 367 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 368 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 369 | 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: 
10147diff
changeset | 370 | by (rule iffD2 [OF order_less_le conjI]) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 371 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 372 | lemma xzgcda_linear [rule_format]: | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 373 | "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: 
10147diff
changeset | 374 | 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: 
10147diff
changeset | 375 | 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: 
10147diff
changeset | 376 | 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: 
10147diff
changeset | 377 | apply (subst xzgcda.simps) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 378 | apply (simp (no_asm)) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 379 | apply (rule impI)+ | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 380 | apply (case_tac "r' mod r = 0") | 
| 13833 | 381 | apply (simp add: xzgcda.simps, clarify) | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 382 | apply (subgoal_tac "0 < r' mod r") | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 383 | apply (rule_tac [2] order_le_neq_implies_less) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 384 | apply (rule_tac [2] pos_mod_sign) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 385 | apply (cut_tac m = m and n = n and r' = r' and r = r and s' = s' and | 
| 13833 | 386 | 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: 
10147diff
changeset | 387 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 388 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 389 | lemma xzgcd_linear: | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 390 | "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: 
10147diff
changeset | 391 | apply (unfold xzgcd_def) | 
| 13837 
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
 paulson parents: 
13833diff
changeset | 392 | apply (erule xzgcda_linear, assumption, auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 393 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 394 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 395 | lemma zgcd_ex_linear: | 
| 27556 | 396 | "0 < n ==> zgcd m n = k ==> (\<exists>s t. k = s * m + t * n)" | 
| 13833 | 397 | apply (simp add: xzgcd_correct, safe) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 398 | apply (rule exI)+ | 
| 13833 | 399 | apply (erule xzgcd_linear, auto) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 400 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 401 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 402 | lemma zcong_lineq_ex: | 
| 27556 | 403 | "0 < n ==> zgcd a n = 1 ==> \<exists>x. [a * x = 1] (mod n)" | 
| 13833 | 404 | 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: 
10147diff
changeset | 405 | apply (rule_tac x = s in exI) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 406 | 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: 
10147diff
changeset | 407 | prefer 2 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 408 | apply simp | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 409 | apply (unfold zcong_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 410 | apply (simp (no_asm) add: zmult_commute zdvd_zminus_iff) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 411 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 412 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 413 | lemma zcong_lineq_unique: | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 414 | "0 < n ==> | 
| 27556 | 415 | 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: 
10147diff
changeset | 416 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 417 | apply (rule_tac [2] zcong_zless_imp_eq) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 418 |        apply (tactic {* stac (thm "zcong_cancel2" RS sym) 6 *})
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 419 | apply (rule_tac [8] zcong_trans) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 420 | apply (simp_all (no_asm_simp)) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 421 | prefer 2 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 422 | apply (simp add: zcong_sym) | 
| 13833 | 423 | apply (cut_tac a = a and n = n in zcong_lineq_ex, auto) | 
| 424 | apply (rule_tac x = "x * b mod n" in exI, safe) | |
| 13788 | 425 | apply (simp_all (no_asm_simp)) | 
| 23839 | 426 | apply (metis zcong_scalar zcong_zmod zmod_zmult1_eq zmult_1 zmult_assoc) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
10147diff
changeset | 427 | done | 
| 9508 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 428 | |
| 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 429 | end |