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