| author | haftmann | 
| Mon, 06 Jul 2009 16:49:51 +0200 | |
| changeset 31950 | 7300186d745a | 
| parent 30242 | aea5d7fa7ef5 | 
| permissions | -rw-r--r-- | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 1 | (* Title: HOL/NumberTheory/Chinese.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: 
9508diff
changeset | 3 | Author: Thomas M. Rasmussen | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
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: 
9508diff
changeset | 7 | header {* The Chinese Remainder Theorem *}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 8 | |
| 27556 | 9 | theory Chinese | 
| 10 | imports IntPrimes | |
| 11 | begin | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 12 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 13 | text {*
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 14 | The Chinese Remainder Theorem for an arbitrary finite number of | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 15 |   equations.  (The one-equation case is included in theory @{text
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 16 |   IntPrimes}.  Uses functions for indexing.\footnote{Maybe @{term
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 17 |   funprod} and @{term funsum} should be based on general @{term fold}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 18 | on indices?} | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 19 | *} | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 20 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 21 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
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: 
9508diff
changeset | 25 | funprod :: "(nat => int) => nat => nat => int" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 26 | funsum :: "(nat => int) => nat => nat => int" | 
| 9508 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 27 | |
| 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 28 | primrec | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 29 | "funprod f i 0 = f i" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 30 | "funprod f i (Suc n) = f (Suc (i + n)) * funprod f i n" | 
| 9508 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 31 | |
| 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 32 | primrec | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 33 | "funsum f i 0 = f i" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 34 | "funsum f i (Suc n) = f (Suc (i + n)) + funsum f i n" | 
| 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: 
20432diff
changeset | 37 | m_cond :: "nat => (nat => int) => bool" where | 
| 19670 | 38 | "m_cond n mf = | 
| 39 | ((\<forall>i. i \<le> n --> 0 < mf i) \<and> | |
| 27556 | 40 | (\<forall>i j. i \<le> n \<and> j \<le> n \<and> i \<noteq> j --> zgcd (mf i) (mf j) = 1))" | 
| 19670 | 41 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20432diff
changeset | 42 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20432diff
changeset | 43 | km_cond :: "nat => (nat => int) => (nat => int) => bool" where | 
| 27556 | 44 | "km_cond n kf mf = (\<forall>i. i \<le> n --> zgcd (kf i) (mf i) = 1)" | 
| 19670 | 45 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20432diff
changeset | 46 | definition | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 47 | lincong_sol :: | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20432diff
changeset | 48 | "nat => (nat => int) => (nat => int) => (nat => int) => int => bool" where | 
| 19670 | 49 | "lincong_sol n kf bf mf x = (\<forall>i. i \<le> n --> zcong (kf i * x) (bf i) (mf i))" | 
| 9508 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 50 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20432diff
changeset | 51 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20432diff
changeset | 52 | mhf :: "(nat => int) => nat => nat => int" where | 
| 19670 | 53 | "mhf mf n i = | 
| 54 | (if i = 0 then funprod mf (Suc 0) (n - Suc 0) | |
| 55 | else if i = n then funprod mf 0 (n - Suc 0) | |
| 56 | else funprod mf 0 (i - Suc 0) * funprod mf (Suc i) (n - Suc 0 - i))" | |
| 57 | ||
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20432diff
changeset | 58 | definition | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 59 | xilin_sol :: | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20432diff
changeset | 60 | "nat => nat => (nat => int) => (nat => int) => (nat => int) => int" where | 
| 19670 | 61 | "xilin_sol i n kf bf mf = | 
| 62 | (if 0 < n \<and> i \<le> n \<and> m_cond n mf \<and> km_cond n kf mf then | |
| 63 | (SOME x. 0 \<le> x \<and> x < mf i \<and> zcong (kf i * mhf mf n i * x) (bf i) (mf i)) | |
| 64 | else 0)" | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 65 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20432diff
changeset | 66 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20432diff
changeset | 67 | x_sol :: "nat => (nat => int) => (nat => int) => (nat => int) => int" where | 
| 19670 | 68 | "x_sol n kf bf mf = funsum (\<lambda>i. xilin_sol i n kf bf mf * mhf mf n i) 0 n" | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 69 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 70 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 71 | text {* \medskip @{term funprod} and @{term funsum} *}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 72 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 73 | lemma funprod_pos: "(\<forall>i. i \<le> n --> 0 < mf i) ==> 0 < funprod mf 0 n" | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 74 | apply (induct n) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 75 | apply auto | 
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
13524diff
changeset | 76 | apply (simp add: zero_less_mult_iff) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 77 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 78 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 79 | lemma funprod_zgcd [rule_format (no_asm)]: | 
| 27556 | 80 | "(\<forall>i. k \<le> i \<and> i \<le> k + l --> zgcd (mf i) (mf m) = 1) --> | 
| 81 | zgcd (funprod mf k l) (mf m) = 1" | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 82 | apply (induct l) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 83 | apply simp_all | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 84 | apply (rule impI)+ | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 85 | apply (subst zgcd_zmult_cancel) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 86 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 87 | done | 
| 9508 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 88 | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 89 | lemma funprod_zdvd [rule_format]: | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 90 | "k \<le> i --> i \<le> k + l --> mf i dvd funprod mf k l" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 91 | apply (induct l) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 92 | apply auto | 
| 30042 | 93 | apply (subgoal_tac "i = Suc (k + l)") | 
| 94 | apply (simp_all (no_asm_simp)) | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 95 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 96 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 97 | lemma funsum_mod: | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 98 | "funsum f k l mod m = funsum (\<lambda>i. (f i) mod m) k l mod m" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 99 | apply (induct l) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 100 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 101 | apply (rule trans) | 
| 29948 | 102 | apply (rule mod_add_eq) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 103 | apply simp | 
| 30034 | 104 | apply (rule mod_add_right_eq [symmetric]) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 105 | done | 
| 9508 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 106 | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 107 | lemma funsum_zero [rule_format (no_asm)]: | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 108 | "(\<forall>i. k \<le> i \<and> i \<le> k + l --> f i = 0) --> (funsum f k l) = 0" | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 109 | apply (induct l) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 110 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 111 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 112 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 113 | lemma funsum_oneelem [rule_format (no_asm)]: | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 114 | "k \<le> j --> j \<le> k + l --> | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 115 | (\<forall>i. k \<le> i \<and> i \<le> k + l \<and> i \<noteq> j --> f i = 0) --> | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 116 | funsum f k l = f j" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 117 | apply (induct l) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 118 | prefer 2 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 119 | apply clarify | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 120 | defer | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 121 | apply clarify | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 122 | apply (subgoal_tac "k = j") | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 123 | apply (simp_all (no_asm_simp)) | 
| 15236 
f289e8ba2bb3
Proofs needed to be updated because induction now preserves name of
 nipkow parents: 
15197diff
changeset | 124 | apply (case_tac "Suc (k + l) = j") | 
| 
f289e8ba2bb3
Proofs needed to be updated because induction now preserves name of
 nipkow parents: 
15197diff
changeset | 125 | apply (subgoal_tac "funsum f k l = 0") | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 126 | apply (rule_tac [2] funsum_zero) | 
| 15236 
f289e8ba2bb3
Proofs needed to be updated because induction now preserves name of
 nipkow parents: 
15197diff
changeset | 127 | apply (subgoal_tac [3] "f (Suc (k + l)) = 0") | 
| 
f289e8ba2bb3
Proofs needed to be updated because induction now preserves name of
 nipkow parents: 
15197diff
changeset | 128 | apply (subgoal_tac [3] "j \<le> k + l") | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 129 | prefer 4 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 130 | apply arith | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 131 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 132 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 133 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 134 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 135 | subsection {* Chinese: uniqueness *}
 | 
| 9508 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 136 | |
| 13524 | 137 | lemma zcong_funprod_aux: | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 138 | "m_cond n mf ==> km_cond n kf mf | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 139 | ==> lincong_sol n kf bf mf x ==> lincong_sol n kf bf mf y | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 140 | ==> [x = y] (mod mf n)" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 141 | apply (unfold m_cond_def km_cond_def lincong_sol_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 142 | apply (rule iffD1) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 143 | apply (rule_tac k = "kf n" in zcong_cancel2) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 144 | apply (rule_tac [3] b = "bf n" in zcong_trans) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 145 | prefer 4 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 146 | apply (subst zcong_sym) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 147 | defer | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 148 | apply (rule order_less_imp_le) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 149 | apply simp_all | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 150 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 151 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 152 | lemma zcong_funprod [rule_format]: | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 153 | "m_cond n mf --> km_cond n kf mf --> | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 154 | lincong_sol n kf bf mf x --> lincong_sol n kf bf mf y --> | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 155 | [x = y] (mod funprod mf 0 n)" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 156 | apply (induct n) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 157 | apply (simp_all (no_asm)) | 
| 13524 | 158 | apply (blast intro: zcong_funprod_aux) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 159 | apply (rule impI)+ | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 160 | apply (rule zcong_zgcd_zmult_zmod) | 
| 13524 | 161 | apply (blast intro: zcong_funprod_aux) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 162 | prefer 2 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 163 | apply (subst zgcd_commute) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 164 | apply (rule funprod_zgcd) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 165 | apply (auto simp add: m_cond_def km_cond_def lincong_sol_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 166 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 167 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 168 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 169 | subsection {* Chinese: existence *}
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 170 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 171 | lemma unique_xi_sol: | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 172 | "0 < n ==> i \<le> n ==> m_cond n mf ==> km_cond n kf mf | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 173 | ==> \<exists>!x. 0 \<le> x \<and> x < mf i \<and> [kf i * mhf mf n i * x = bf i] (mod mf i)" | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 174 | apply (rule zcong_lineq_unique) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 175 |    apply (tactic {* stac (thm "zgcd_zmult_cancel") 2 *})
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 176 | apply (unfold m_cond_def km_cond_def mhf_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 177 | apply (simp_all (no_asm_simp)) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 178 | apply safe | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 179 |     apply (tactic {* stac (thm "zgcd_zmult_cancel") 3 *})
 | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 180 | apply (rule_tac [!] funprod_zgcd) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 181 | apply safe | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 182 | apply simp_all | 
| 20432 
07ec57376051
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20272diff
changeset | 183 | apply (subgoal_tac "i<n") | 
| 
07ec57376051
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20272diff
changeset | 184 | prefer 2 | 
| 
07ec57376051
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20272diff
changeset | 185 | apply arith | 
| 
07ec57376051
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20272diff
changeset | 186 | apply (case_tac [2] i) | 
| 
07ec57376051
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20272diff
changeset | 187 | apply simp_all | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 188 | done | 
| 9508 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 189 | |
| 13524 | 190 | lemma x_sol_lin_aux: | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 191 | "0 < n ==> i \<le> n ==> j \<le> n ==> j \<noteq> i ==> mf j dvd mhf mf n i" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 192 | apply (unfold mhf_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 193 | apply (case_tac "i = 0") | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 194 | apply (case_tac [2] "i = n") | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 195 | apply (simp_all (no_asm_simp)) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 196 | apply (case_tac [3] "j < i") | 
| 30042 | 197 | apply (rule_tac [3] dvd_mult2) | 
| 198 | apply (rule_tac [4] dvd_mult) | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 199 | apply (rule_tac [!] funprod_zdvd) | 
| 23315 | 200 | apply arith | 
| 201 | apply arith | |
| 202 | apply arith | |
| 203 | apply arith | |
| 204 | apply arith | |
| 205 | apply arith | |
| 206 | apply arith | |
| 207 | apply arith | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 208 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 209 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 210 | lemma x_sol_lin: | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 211 | "0 < n ==> i \<le> n | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 212 | ==> x_sol n kf bf mf mod mf i = | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 213 | xilin_sol i n kf bf mf * mhf mf n i mod mf i" | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 214 | apply (unfold x_sol_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 215 | apply (subst funsum_mod) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 216 | apply (subst funsum_oneelem) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 217 | apply auto | 
| 30042 | 218 | apply (subst dvd_eq_mod_eq_0 [symmetric]) | 
| 219 | apply (rule dvd_mult) | |
| 13524 | 220 | apply (rule x_sol_lin_aux) | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 221 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 222 | done | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 223 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 224 | |
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 225 | subsection {* Chinese *}
 | 
| 9508 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 226 | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 227 | lemma chinese_remainder: | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 228 | "0 < n ==> m_cond n mf ==> km_cond n kf mf | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 229 | ==> \<exists>!x. 0 \<le> x \<and> x < funprod mf 0 n \<and> lincong_sol n kf bf mf x" | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 230 | apply safe | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 231 | apply (rule_tac [2] m = "funprod mf 0 n" in zcong_zless_imp_eq) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 232 | apply (rule_tac [6] zcong_funprod) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 233 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 234 | apply (rule_tac x = "x_sol n kf bf mf mod funprod mf 0 n" in exI) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 235 | apply (unfold lincong_sol_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 236 | apply safe | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 237 |     apply (tactic {* stac (thm "zcong_zmod") 3 *})
 | 
| 29948 | 238 |     apply (tactic {* stac (thm "mod_mult_eq") 3 *})
 | 
| 30034 | 239 |     apply (tactic {* stac (thm "mod_mod_cancel") 3 *})
 | 
| 240 |       apply (tactic {* stac (thm "x_sol_lin") 4 *})
 | |
| 241 |         apply (tactic {* stac (thm "mod_mult_eq" RS sym) 6 *})
 | |
| 242 |         apply (tactic {* stac (thm "zcong_zmod" RS sym) 6 *})
 | |
| 243 | apply (subgoal_tac [6] | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 244 | "0 \<le> xilin_sol i n kf bf mf \<and> xilin_sol i n kf bf mf < mf i | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 245 | \<and> [kf i * mhf mf n i * xilin_sol i n kf bf mf = bf i] (mod mf i)") | 
| 30034 | 246 | prefer 6 | 
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 247 | apply (simp add: zmult_ac) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 248 | apply (unfold xilin_sol_def) | 
| 30034 | 249 |         apply (tactic {* asm_simp_tac @{simpset} 6 *})
 | 
| 250 | apply (rule_tac [6] ex1_implies_ex [THEN someI_ex]) | |
| 251 | apply (rule_tac [6] unique_xi_sol) | |
| 252 | apply (rule_tac [3] funprod_zdvd) | |
| 11049 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 253 | apply (unfold m_cond_def) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 254 | apply (rule funprod_pos [THEN pos_mod_sign]) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 255 | apply (rule_tac [2] funprod_pos [THEN pos_mod_bound]) | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 256 | apply auto | 
| 
7eef34adb852
HOL-NumberTheory: converted to new-style format and proper document setup;
 wenzelm parents: 
9508diff
changeset | 257 | done | 
| 9508 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 258 | |
| 
4d01dbf6ded7
Chinese Remainder Theorem, Wilsons Theorem, etc., by T M Masmussen
 paulson parents: diff
changeset | 259 | end |