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