| author | huffman | 
| Thu, 11 Jun 2009 09:03:24 -0700 | |
| changeset 31563 | ded2364d14d4 | 
| parent 31068 | f591144b0f17 | 
| child 31662 | 57f7ef0dba8e | 
| permissions | -rw-r--r-- | 
| 23164 | 1 | (* Title: HOL/IntDiv.thy | 
| 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 3 | Copyright 1999 University of Cambridge | |
| 4 | ||
| 5 | *) | |
| 6 | ||
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 7 | header{* The Division Operators div and mod *}
 | 
| 23164 | 8 | |
| 9 | theory IntDiv | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25571diff
changeset | 10 | imports Int Divides FunDef | 
| 23164 | 11 | begin | 
| 12 | ||
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 13 | definition divmod_rel :: "int \<Rightarrow> int \<Rightarrow> int \<times> int \<Rightarrow> bool" where | 
| 23164 | 14 |     --{*definition of quotient and remainder*}
 | 
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 15 | [code]: "divmod_rel a b = (\<lambda>(q, r). a = b * q + r \<and> | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 16 | (if 0 < b then 0 \<le> r \<and> r < b else b < r \<and> r \<le> 0))" | 
| 23164 | 17 | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 18 | definition adjust :: "int \<Rightarrow> int \<times> int \<Rightarrow> int \<times> int" where | 
| 23164 | 19 |     --{*for the division algorithm*}
 | 
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 20 | [code]: "adjust b = (\<lambda>(q, r). if 0 \<le> r - b then (2 * q + 1, r - b) | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 21 | else (2 * q, r))" | 
| 23164 | 22 | |
| 23 | text{*algorithm for the case @{text "a\<ge>0, b>0"}*}
 | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 24 | function posDivAlg :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 25 | "posDivAlg a b = (if a < b \<or> b \<le> 0 then (0, a) | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 26 | else adjust b (posDivAlg a (2 * b)))" | 
| 23164 | 27 | by auto | 
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 28 | termination by (relation "measure (\<lambda>(a, b). nat (a - b + 1))") auto | 
| 23164 | 29 | |
| 30 | text{*algorithm for the case @{text "a<0, b>0"}*}
 | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 31 | function negDivAlg :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 32 | "negDivAlg a b = (if 0 \<le>a + b \<or> b \<le> 0 then (-1, a + b) | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 33 | else adjust b (negDivAlg a (2 * b)))" | 
| 23164 | 34 | by auto | 
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 35 | termination by (relation "measure (\<lambda>(a, b). nat (- a - b))") auto | 
| 23164 | 36 | |
| 37 | text{*algorithm for the general case @{term "b\<noteq>0"}*}
 | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 38 | definition negateSnd :: "int \<times> int \<Rightarrow> int \<times> int" where | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 39 | [code inline]: "negateSnd = apsnd uminus" | 
| 23164 | 40 | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 41 | definition divmod :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where | 
| 23164 | 42 |     --{*The full division algorithm considers all possible signs for a, b
 | 
| 43 |        including the special case @{text "a=0, b<0"} because 
 | |
| 44 |        @{term negDivAlg} requires @{term "a<0"}.*}
 | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 45 | "divmod a b = (if 0 \<le> a then if 0 \<le> b then posDivAlg a b | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 46 | else if a = 0 then (0, 0) | 
| 23164 | 47 | else negateSnd (negDivAlg (-a) (-b)) | 
| 48 | else | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 49 | if 0 < b then negDivAlg a b | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 50 | else negateSnd (posDivAlg (-a) (-b)))" | 
| 23164 | 51 | |
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25134diff
changeset | 52 | instantiation int :: Divides.div | 
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25134diff
changeset | 53 | begin | 
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25134diff
changeset | 54 | |
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25134diff
changeset | 55 | definition | 
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 56 | div_def: "a div b = fst (divmod a b)" | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25134diff
changeset | 57 | |
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25134diff
changeset | 58 | definition | 
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 59 | mod_def: "a mod b = snd (divmod a b)" | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25134diff
changeset | 60 | |
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25134diff
changeset | 61 | instance .. | 
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25134diff
changeset | 62 | |
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25134diff
changeset | 63 | end | 
| 23164 | 64 | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 65 | lemma divmod_mod_div: | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 66 | "divmod p q = (p div q, p mod q)" | 
| 23164 | 67 | by (auto simp add: div_def mod_def) | 
| 68 | ||
| 69 | text{*
 | |
| 70 | Here is the division algorithm in ML: | |
| 71 | ||
| 72 | \begin{verbatim}
 | |
| 73 | fun posDivAlg (a,b) = | |
| 74 | if a<b then (0,a) | |
| 75 | else let val (q,r) = posDivAlg(a, 2*b) | |
| 76 | in if 0\<le>r-b then (2*q+1, r-b) else (2*q, r) | |
| 77 | end | |
| 78 | ||
| 79 | fun negDivAlg (a,b) = | |
| 80 | if 0\<le>a+b then (~1,a+b) | |
| 81 | else let val (q,r) = negDivAlg(a, 2*b) | |
| 82 | in if 0\<le>r-b then (2*q+1, r-b) else (2*q, r) | |
| 83 | end; | |
| 84 | ||
| 85 | fun negateSnd (q,r:int) = (q,~r); | |
| 86 | ||
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 87 | fun divmod (a,b) = if 0\<le>a then | 
| 23164 | 88 | if b>0 then posDivAlg (a,b) | 
| 89 | else if a=0 then (0,0) | |
| 90 | else negateSnd (negDivAlg (~a,~b)) | |
| 91 | else | |
| 92 | if 0<b then negDivAlg (a,b) | |
| 93 | else negateSnd (posDivAlg (~a,~b)); | |
| 94 | \end{verbatim}
 | |
| 95 | *} | |
| 96 | ||
| 97 | ||
| 98 | ||
| 99 | subsection{*Uniqueness and Monotonicity of Quotients and Remainders*}
 | |
| 100 | ||
| 101 | lemma unique_quotient_lemma: | |
| 102 | "[| b*q' + r' \<le> b*q + r; 0 \<le> r'; r' < b; r < b |] | |
| 103 | ==> q' \<le> (q::int)" | |
| 104 | apply (subgoal_tac "r' + b * (q'-q) \<le> r") | |
| 105 | prefer 2 apply (simp add: right_diff_distrib) | |
| 106 | apply (subgoal_tac "0 < b * (1 + q - q') ") | |
| 107 | apply (erule_tac [2] order_le_less_trans) | |
| 108 | prefer 2 apply (simp add: right_diff_distrib right_distrib) | |
| 109 | apply (subgoal_tac "b * q' < b * (1 + q) ") | |
| 110 | prefer 2 apply (simp add: right_diff_distrib right_distrib) | |
| 111 | apply (simp add: mult_less_cancel_left) | |
| 112 | done | |
| 113 | ||
| 114 | lemma unique_quotient_lemma_neg: | |
| 115 | "[| b*q' + r' \<le> b*q + r; r \<le> 0; b < r; b < r' |] | |
| 116 | ==> q \<le> (q'::int)" | |
| 117 | by (rule_tac b = "-b" and r = "-r'" and r' = "-r" in unique_quotient_lemma, | |
| 118 | auto) | |
| 119 | ||
| 120 | lemma unique_quotient: | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 121 | "[| divmod_rel a b (q, r); divmod_rel a b (q', r'); b \<noteq> 0 |] | 
| 23164 | 122 | ==> q = q'" | 
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 123 | apply (simp add: divmod_rel_def linorder_neq_iff split: split_if_asm) | 
| 23164 | 124 | apply (blast intro: order_antisym | 
| 125 | dest: order_eq_refl [THEN unique_quotient_lemma] | |
| 126 | order_eq_refl [THEN unique_quotient_lemma_neg] sym)+ | |
| 127 | done | |
| 128 | ||
| 129 | ||
| 130 | lemma unique_remainder: | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 131 | "[| divmod_rel a b (q, r); divmod_rel a b (q', r'); b \<noteq> 0 |] | 
| 23164 | 132 | ==> r = r'" | 
| 133 | apply (subgoal_tac "q = q'") | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 134 | apply (simp add: divmod_rel_def) | 
| 23164 | 135 | apply (blast intro: unique_quotient) | 
| 136 | done | |
| 137 | ||
| 138 | ||
| 139 | subsection{*Correctness of @{term posDivAlg}, the Algorithm for Non-Negative Dividends*}
 | |
| 140 | ||
| 141 | text{*And positive divisors*}
 | |
| 142 | ||
| 143 | lemma adjust_eq [simp]: | |
| 144 | "adjust b (q,r) = | |
| 145 | (let diff = r-b in | |
| 146 | if 0 \<le> diff then (2*q + 1, diff) | |
| 147 | else (2*q, r))" | |
| 148 | by (simp add: Let_def adjust_def) | |
| 149 | ||
| 150 | declare posDivAlg.simps [simp del] | |
| 151 | ||
| 152 | text{*use with a simproc to avoid repeatedly proving the premise*}
 | |
| 153 | lemma posDivAlg_eqn: | |
| 154 | "0 < b ==> | |
| 155 | posDivAlg a b = (if a<b then (0,a) else adjust b (posDivAlg a (2*b)))" | |
| 156 | by (rule posDivAlg.simps [THEN trans], simp) | |
| 157 | ||
| 158 | text{*Correctness of @{term posDivAlg}: it computes quotients correctly*}
 | |
| 159 | theorem posDivAlg_correct: | |
| 160 | assumes "0 \<le> a" and "0 < b" | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 161 | shows "divmod_rel a b (posDivAlg a b)" | 
| 23164 | 162 | using prems apply (induct a b rule: posDivAlg.induct) | 
| 163 | apply auto | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 164 | apply (simp add: divmod_rel_def) | 
| 23164 | 165 | apply (subst posDivAlg_eqn, simp add: right_distrib) | 
| 166 | apply (case_tac "a < b") | |
| 167 | apply simp_all | |
| 168 | apply (erule splitE) | |
| 169 | apply (auto simp add: right_distrib Let_def) | |
| 170 | done | |
| 171 | ||
| 172 | ||
| 173 | subsection{*Correctness of @{term negDivAlg}, the Algorithm for Negative Dividends*}
 | |
| 174 | ||
| 175 | text{*And positive divisors*}
 | |
| 176 | ||
| 177 | declare negDivAlg.simps [simp del] | |
| 178 | ||
| 179 | text{*use with a simproc to avoid repeatedly proving the premise*}
 | |
| 180 | lemma negDivAlg_eqn: | |
| 181 | "0 < b ==> | |
| 182 | negDivAlg a b = | |
| 183 | (if 0\<le>a+b then (-1,a+b) else adjust b (negDivAlg a (2*b)))" | |
| 184 | by (rule negDivAlg.simps [THEN trans], simp) | |
| 185 | ||
| 186 | (*Correctness of negDivAlg: it computes quotients correctly | |
| 187 | It doesn't work if a=0 because the 0/b equals 0, not -1*) | |
| 188 | lemma negDivAlg_correct: | |
| 189 | assumes "a < 0" and "b > 0" | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 190 | shows "divmod_rel a b (negDivAlg a b)" | 
| 23164 | 191 | using prems apply (induct a b rule: negDivAlg.induct) | 
| 192 | apply (auto simp add: linorder_not_le) | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 193 | apply (simp add: divmod_rel_def) | 
| 23164 | 194 | apply (subst negDivAlg_eqn, assumption) | 
| 195 | apply (case_tac "a + b < (0\<Colon>int)") | |
| 196 | apply simp_all | |
| 197 | apply (erule splitE) | |
| 198 | apply (auto simp add: right_distrib Let_def) | |
| 199 | done | |
| 200 | ||
| 201 | ||
| 202 | subsection{*Existence Shown by Proving the Division Algorithm to be Correct*}
 | |
| 203 | ||
| 204 | (*the case a=0*) | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 205 | lemma divmod_rel_0: "b \<noteq> 0 ==> divmod_rel 0 b (0, 0)" | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 206 | by (auto simp add: divmod_rel_def linorder_neq_iff) | 
| 23164 | 207 | |
| 208 | lemma posDivAlg_0 [simp]: "posDivAlg 0 b = (0, 0)" | |
| 209 | by (subst posDivAlg.simps, auto) | |
| 210 | ||
| 211 | lemma negDivAlg_minus1 [simp]: "negDivAlg -1 b = (-1, b - 1)" | |
| 212 | by (subst negDivAlg.simps, auto) | |
| 213 | ||
| 214 | lemma negateSnd_eq [simp]: "negateSnd(q,r) = (q,-r)" | |
| 215 | by (simp add: negateSnd_def) | |
| 216 | ||
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 217 | lemma divmod_rel_neg: "divmod_rel (-a) (-b) qr ==> divmod_rel a b (negateSnd qr)" | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 218 | by (auto simp add: split_ifs divmod_rel_def) | 
| 23164 | 219 | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 220 | lemma divmod_correct: "b \<noteq> 0 ==> divmod_rel a b (divmod a b)" | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 221 | by (force simp add: linorder_neq_iff divmod_rel_0 divmod_def divmod_rel_neg | 
| 23164 | 222 | posDivAlg_correct negDivAlg_correct) | 
| 223 | ||
| 224 | text{*Arbitrary definitions for division by zero.  Useful to simplify 
 | |
| 225 | certain equations.*} | |
| 226 | ||
| 227 | lemma DIVISION_BY_ZERO [simp]: "a div (0::int) = 0 & a mod (0::int) = a" | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 228 | by (simp add: div_def mod_def divmod_def posDivAlg.simps) | 
| 23164 | 229 | |
| 230 | ||
| 231 | text{*Basic laws about division and remainder*}
 | |
| 232 | ||
| 233 | lemma zmod_zdiv_equality: "(a::int) = b * (a div b) + (a mod b)" | |
| 234 | apply (case_tac "b = 0", simp) | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 235 | apply (cut_tac a = a and b = b in divmod_correct) | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 236 | apply (auto simp add: divmod_rel_def div_def mod_def) | 
| 23164 | 237 | done | 
| 238 | ||
| 239 | lemma zdiv_zmod_equality: "(b * (a div b) + (a mod b)) + k = (a::int)+k" | |
| 240 | by(simp add: zmod_zdiv_equality[symmetric]) | |
| 241 | ||
| 242 | lemma zdiv_zmod_equality2: "((a div b) * b + (a mod b)) + k = (a::int)+k" | |
| 243 | by(simp add: mult_commute zmod_zdiv_equality[symmetric]) | |
| 244 | ||
| 245 | text {* Tool setup *}
 | |
| 246 | ||
| 26480 | 247 | ML {*
 | 
| 30934 | 248 | local | 
| 23164 | 249 | |
| 30934 | 250 | structure CancelDivMod = CancelDivModFun(struct | 
| 251 | ||
| 252 |   val div_name = @{const_name div};
 | |
| 253 |   val mod_name = @{const_name mod};
 | |
| 23164 | 254 | val mk_binop = HOLogic.mk_binop; | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31065diff
changeset | 255 | val mk_sum = Numeral_Simprocs.mk_sum HOLogic.intT; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31065diff
changeset | 256 | val dest_sum = Numeral_Simprocs.dest_sum; | 
| 30934 | 257 | |
| 258 |   val div_mod_eqs = map mk_meta_eq [@{thm zdiv_zmod_equality}, @{thm zdiv_zmod_equality2}];
 | |
| 259 | ||
| 23164 | 260 | val trans = trans; | 
| 30934 | 261 | |
| 262 | val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac | |
| 263 |     (@{thm diff_minus} :: @{thms add_0s} @ @{thms add_ac}))
 | |
| 264 | ||
| 23164 | 265 | end) | 
| 266 | ||
| 267 | in | |
| 268 | ||
| 30934 | 269 | val cancel_div_mod_int_proc = Simplifier.simproc (the_context ()) | 
| 270 | "cancel_zdiv_zmod" ["(k::int) + l"] (K CancelDivMod.proc); | |
| 23164 | 271 | |
| 30934 | 272 | val _ = Addsimprocs [cancel_div_mod_int_proc]; | 
| 23164 | 273 | |
| 30934 | 274 | end | 
| 23164 | 275 | *} | 
| 276 | ||
| 277 | lemma pos_mod_conj : "(0::int) < b ==> 0 \<le> a mod b & a mod b < b" | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 278 | apply (cut_tac a = a and b = b in divmod_correct) | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 279 | apply (auto simp add: divmod_rel_def mod_def) | 
| 23164 | 280 | done | 
| 281 | ||
| 282 | lemmas pos_mod_sign [simp] = pos_mod_conj [THEN conjunct1, standard] | |
| 283 | and pos_mod_bound [simp] = pos_mod_conj [THEN conjunct2, standard] | |
| 284 | ||
| 285 | lemma neg_mod_conj : "b < (0::int) ==> a mod b \<le> 0 & b < a mod b" | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 286 | apply (cut_tac a = a and b = b in divmod_correct) | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 287 | apply (auto simp add: divmod_rel_def div_def mod_def) | 
| 23164 | 288 | done | 
| 289 | ||
| 290 | lemmas neg_mod_sign [simp] = neg_mod_conj [THEN conjunct1, standard] | |
| 291 | and neg_mod_bound [simp] = neg_mod_conj [THEN conjunct2, standard] | |
| 292 | ||
| 293 | ||
| 294 | ||
| 295 | subsection{*General Properties of div and mod*}
 | |
| 296 | ||
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 297 | lemma divmod_rel_div_mod: "b \<noteq> 0 ==> divmod_rel a b (a div b, a mod b)" | 
| 23164 | 298 | apply (cut_tac a = a and b = b in zmod_zdiv_equality) | 
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 299 | apply (force simp add: divmod_rel_def linorder_neq_iff) | 
| 23164 | 300 | done | 
| 301 | ||
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 302 | lemma divmod_rel_div: "[| divmod_rel a b (q, r); b \<noteq> 0 |] ==> a div b = q" | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 303 | by (simp add: divmod_rel_div_mod [THEN unique_quotient]) | 
| 23164 | 304 | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 305 | lemma divmod_rel_mod: "[| divmod_rel a b (q, r); b \<noteq> 0 |] ==> a mod b = r" | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 306 | by (simp add: divmod_rel_div_mod [THEN unique_remainder]) | 
| 23164 | 307 | |
| 308 | lemma div_pos_pos_trivial: "[| (0::int) \<le> a; a < b |] ==> a div b = 0" | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 309 | apply (rule divmod_rel_div) | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 310 | apply (auto simp add: divmod_rel_def) | 
| 23164 | 311 | done | 
| 312 | ||
| 313 | lemma div_neg_neg_trivial: "[| a \<le> (0::int); b < a |] ==> a div b = 0" | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 314 | apply (rule divmod_rel_div) | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 315 | apply (auto simp add: divmod_rel_def) | 
| 23164 | 316 | done | 
| 317 | ||
| 318 | lemma div_pos_neg_trivial: "[| (0::int) < a; a+b \<le> 0 |] ==> a div b = -1" | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 319 | apply (rule divmod_rel_div) | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 320 | apply (auto simp add: divmod_rel_def) | 
| 23164 | 321 | done | 
| 322 | ||
| 323 | (*There is no div_neg_pos_trivial because 0 div b = 0 would supersede it*) | |
| 324 | ||
| 325 | lemma mod_pos_pos_trivial: "[| (0::int) \<le> a; a < b |] ==> a mod b = a" | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 326 | apply (rule_tac q = 0 in divmod_rel_mod) | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 327 | apply (auto simp add: divmod_rel_def) | 
| 23164 | 328 | done | 
| 329 | ||
| 330 | lemma mod_neg_neg_trivial: "[| a \<le> (0::int); b < a |] ==> a mod b = a" | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 331 | apply (rule_tac q = 0 in divmod_rel_mod) | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 332 | apply (auto simp add: divmod_rel_def) | 
| 23164 | 333 | done | 
| 334 | ||
| 335 | lemma mod_pos_neg_trivial: "[| (0::int) < a; a+b \<le> 0 |] ==> a mod b = a+b" | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 336 | apply (rule_tac q = "-1" in divmod_rel_mod) | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 337 | apply (auto simp add: divmod_rel_def) | 
| 23164 | 338 | done | 
| 339 | ||
| 340 | text{*There is no @{text mod_neg_pos_trivial}.*}
 | |
| 341 | ||
| 342 | ||
| 343 | (*Simpler laws such as -a div b = -(a div b) FAIL, but see just below*) | |
| 344 | lemma zdiv_zminus_zminus [simp]: "(-a) div (-b) = a div (b::int)" | |
| 345 | apply (case_tac "b = 0", simp) | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 346 | apply (simp add: divmod_rel_div_mod [THEN divmod_rel_neg, simplified, | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 347 | THEN divmod_rel_div, THEN sym]) | 
| 23164 | 348 | |
| 349 | done | |
| 350 | ||
| 351 | (*Simpler laws such as -a mod b = -(a mod b) FAIL, but see just below*) | |
| 352 | lemma zmod_zminus_zminus [simp]: "(-a) mod (-b) = - (a mod (b::int))" | |
| 353 | apply (case_tac "b = 0", simp) | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 354 | apply (subst divmod_rel_div_mod [THEN divmod_rel_neg, simplified, THEN divmod_rel_mod], | 
| 23164 | 355 | auto) | 
| 356 | done | |
| 357 | ||
| 358 | ||
| 359 | subsection{*Laws for div and mod with Unary Minus*}
 | |
| 360 | ||
| 361 | lemma zminus1_lemma: | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 362 | "divmod_rel a b (q, r) | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 363 | ==> divmod_rel (-a) b (if r=0 then -q else -q - 1, | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 364 | if r=0 then 0 else b-r)" | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 365 | by (force simp add: split_ifs divmod_rel_def linorder_neq_iff right_diff_distrib) | 
| 23164 | 366 | |
| 367 | ||
| 368 | lemma zdiv_zminus1_eq_if: | |
| 369 | "b \<noteq> (0::int) | |
| 370 | ==> (-a) div b = | |
| 371 | (if a mod b = 0 then - (a div b) else - (a div b) - 1)" | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 372 | by (blast intro: divmod_rel_div_mod [THEN zminus1_lemma, THEN divmod_rel_div]) | 
| 23164 | 373 | |
| 374 | lemma zmod_zminus1_eq_if: | |
| 375 | "(-a::int) mod b = (if a mod b = 0 then 0 else b - (a mod b))" | |
| 376 | apply (case_tac "b = 0", simp) | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 377 | apply (blast intro: divmod_rel_div_mod [THEN zminus1_lemma, THEN divmod_rel_mod]) | 
| 23164 | 378 | done | 
| 379 | ||
| 29936 | 380 | lemma zmod_zminus1_not_zero: | 
| 381 | fixes k l :: int | |
| 382 | shows "- k mod l \<noteq> 0 \<Longrightarrow> k mod l \<noteq> 0" | |
| 383 | unfolding zmod_zminus1_eq_if by auto | |
| 384 | ||
| 23164 | 385 | lemma zdiv_zminus2: "a div (-b) = (-a::int) div b" | 
| 386 | by (cut_tac a = "-a" in zdiv_zminus_zminus, auto) | |
| 387 | ||
| 388 | lemma zmod_zminus2: "a mod (-b) = - ((-a::int) mod b)" | |
| 389 | by (cut_tac a = "-a" and b = b in zmod_zminus_zminus, auto) | |
| 390 | ||
| 391 | lemma zdiv_zminus2_eq_if: | |
| 392 | "b \<noteq> (0::int) | |
| 393 | ==> a div (-b) = | |
| 394 | (if a mod b = 0 then - (a div b) else - (a div b) - 1)" | |
| 395 | by (simp add: zdiv_zminus1_eq_if zdiv_zminus2) | |
| 396 | ||
| 397 | lemma zmod_zminus2_eq_if: | |
| 398 | "a mod (-b::int) = (if a mod b = 0 then 0 else (a mod b) - b)" | |
| 399 | by (simp add: zmod_zminus1_eq_if zmod_zminus2) | |
| 400 | ||
| 29936 | 401 | lemma zmod_zminus2_not_zero: | 
| 402 | fixes k l :: int | |
| 403 | shows "k mod - l \<noteq> 0 \<Longrightarrow> k mod l \<noteq> 0" | |
| 404 | unfolding zmod_zminus2_eq_if by auto | |
| 405 | ||
| 23164 | 406 | |
| 407 | subsection{*Division of a Number by Itself*}
 | |
| 408 | ||
| 409 | lemma self_quotient_aux1: "[| (0::int) < a; a = r + a*q; r < a |] ==> 1 \<le> q" | |
| 410 | apply (subgoal_tac "0 < a*q") | |
| 411 | apply (simp add: zero_less_mult_iff, arith) | |
| 412 | done | |
| 413 | ||
| 414 | lemma self_quotient_aux2: "[| (0::int) < a; a = r + a*q; 0 \<le> r |] ==> q \<le> 1" | |
| 415 | apply (subgoal_tac "0 \<le> a* (1-q) ") | |
| 416 | apply (simp add: zero_le_mult_iff) | |
| 417 | apply (simp add: right_diff_distrib) | |
| 418 | done | |
| 419 | ||
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 420 | lemma self_quotient: "[| divmod_rel a a (q, r); a \<noteq> (0::int) |] ==> q = 1" | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 421 | apply (simp add: split_ifs divmod_rel_def linorder_neq_iff) | 
| 23164 | 422 | apply (rule order_antisym, safe, simp_all) | 
| 423 | apply (rule_tac [3] a = "-a" and r = "-r" in self_quotient_aux1) | |
| 424 | apply (rule_tac a = "-a" and r = "-r" in self_quotient_aux2) | |
| 425 | apply (force intro: self_quotient_aux1 self_quotient_aux2 simp add: add_commute)+ | |
| 426 | done | |
| 427 | ||
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 428 | lemma self_remainder: "[| divmod_rel a a (q, r); a \<noteq> (0::int) |] ==> r = 0" | 
| 23164 | 429 | apply (frule self_quotient, assumption) | 
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 430 | apply (simp add: divmod_rel_def) | 
| 23164 | 431 | done | 
| 432 | ||
| 433 | lemma zdiv_self [simp]: "a \<noteq> 0 ==> a div a = (1::int)" | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 434 | by (simp add: divmod_rel_div_mod [THEN self_quotient]) | 
| 23164 | 435 | |
| 436 | (*Here we have 0 mod 0 = 0, also assumed by Knuth (who puts m mod 0 = 0) *) | |
| 437 | lemma zmod_self [simp]: "a mod a = (0::int)" | |
| 438 | apply (case_tac "a = 0", simp) | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 439 | apply (simp add: divmod_rel_div_mod [THEN self_remainder]) | 
| 23164 | 440 | done | 
| 441 | ||
| 442 | ||
| 443 | subsection{*Computation of Division and Remainder*}
 | |
| 444 | ||
| 445 | lemma zdiv_zero [simp]: "(0::int) div b = 0" | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 446 | by (simp add: div_def divmod_def) | 
| 23164 | 447 | |
| 448 | lemma div_eq_minus1: "(0::int) < b ==> -1 div b = -1" | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 449 | by (simp add: div_def divmod_def) | 
| 23164 | 450 | |
| 451 | lemma zmod_zero [simp]: "(0::int) mod b = 0" | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 452 | by (simp add: mod_def divmod_def) | 
| 23164 | 453 | |
| 454 | lemma zmod_minus1: "(0::int) < b ==> -1 mod b = b - 1" | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 455 | by (simp add: mod_def divmod_def) | 
| 23164 | 456 | |
| 457 | text{*a positive, b positive *}
 | |
| 458 | ||
| 459 | lemma div_pos_pos: "[| 0 < a; 0 \<le> b |] ==> a div b = fst (posDivAlg a b)" | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 460 | by (simp add: div_def divmod_def) | 
| 23164 | 461 | |
| 462 | lemma mod_pos_pos: "[| 0 < a; 0 \<le> b |] ==> a mod b = snd (posDivAlg a b)" | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 463 | by (simp add: mod_def divmod_def) | 
| 23164 | 464 | |
| 465 | text{*a negative, b positive *}
 | |
| 466 | ||
| 467 | lemma div_neg_pos: "[| a < 0; 0 < b |] ==> a div b = fst (negDivAlg a b)" | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 468 | by (simp add: div_def divmod_def) | 
| 23164 | 469 | |
| 470 | lemma mod_neg_pos: "[| a < 0; 0 < b |] ==> a mod b = snd (negDivAlg a b)" | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 471 | by (simp add: mod_def divmod_def) | 
| 23164 | 472 | |
| 473 | text{*a positive, b negative *}
 | |
| 474 | ||
| 475 | lemma div_pos_neg: | |
| 476 | "[| 0 < a; b < 0 |] ==> a div b = fst (negateSnd (negDivAlg (-a) (-b)))" | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 477 | by (simp add: div_def divmod_def) | 
| 23164 | 478 | |
| 479 | lemma mod_pos_neg: | |
| 480 | "[| 0 < a; b < 0 |] ==> a mod b = snd (negateSnd (negDivAlg (-a) (-b)))" | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 481 | by (simp add: mod_def divmod_def) | 
| 23164 | 482 | |
| 483 | text{*a negative, b negative *}
 | |
| 484 | ||
| 485 | lemma div_neg_neg: | |
| 486 | "[| a < 0; b \<le> 0 |] ==> a div b = fst (negateSnd (posDivAlg (-a) (-b)))" | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 487 | by (simp add: div_def divmod_def) | 
| 23164 | 488 | |
| 489 | lemma mod_neg_neg: | |
| 490 | "[| a < 0; b \<le> 0 |] ==> a mod b = snd (negateSnd (posDivAlg (-a) (-b)))" | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 491 | by (simp add: mod_def divmod_def) | 
| 23164 | 492 | |
| 493 | text {*Simplify expresions in which div and mod combine numerical constants*}
 | |
| 494 | ||
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 495 | lemma divmod_relI: | 
| 24481 | 496 | "\<lbrakk>a == b * q + r; if 0 < b then 0 \<le> r \<and> r < b else b < r \<and> r \<le> 0\<rbrakk> | 
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 497 | \<Longrightarrow> divmod_rel a b (q, r)" | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 498 | unfolding divmod_rel_def by simp | 
| 24481 | 499 | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 500 | lemmas divmod_rel_div_eq = divmod_relI [THEN divmod_rel_div, THEN eq_reflection] | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 501 | lemmas divmod_rel_mod_eq = divmod_relI [THEN divmod_rel_mod, THEN eq_reflection] | 
| 24481 | 502 | lemmas arithmetic_simps = | 
| 503 | arith_simps | |
| 504 | add_special | |
| 505 | OrderedGroup.add_0_left | |
| 506 | OrderedGroup.add_0_right | |
| 507 | mult_zero_left | |
| 508 | mult_zero_right | |
| 509 | mult_1_left | |
| 510 | mult_1_right | |
| 511 | ||
| 512 | (* simprocs adapted from HOL/ex/Binary.thy *) | |
| 513 | ML {*
 | |
| 514 | local | |
| 30517 | 515 | val mk_number = HOLogic.mk_number HOLogic.intT; | 
| 516 |   fun mk_cert u k l = @{term "plus :: int \<Rightarrow> int \<Rightarrow> int"} $
 | |
| 517 |     (@{term "times :: int \<Rightarrow> int \<Rightarrow> int"} $ u $ mk_number k) $
 | |
| 518 | mk_number l; | |
| 519 | fun prove ctxt prop = Goal.prove ctxt [] [] prop | |
| 520 |     (K (ALLGOALS (full_simp_tac (HOL_basic_ss addsimps @{thms arithmetic_simps}))));
 | |
| 24481 | 521 | fun binary_proc proc ss ct = | 
| 522 | (case Thm.term_of ct of | |
| 523 | _ $ t $ u => | |
| 524 | (case try (pairself (`(snd o HOLogic.dest_number))) (t, u) of | |
| 525 | SOME args => proc (Simplifier.the_context ss) args | |
| 526 | | NONE => NONE) | |
| 527 | | _ => NONE); | |
| 528 | in | |
| 30517 | 529 | fun divmod_proc rule = binary_proc (fn ctxt => fn ((m, t), (n, u)) => | 
| 530 | if n = 0 then NONE | |
| 531 | else let val (k, l) = Integer.div_mod m n; | |
| 532 | in SOME (rule OF [prove ctxt (Logic.mk_equals (t, mk_cert u k l))]) end); | |
| 533 | end | |
| 24481 | 534 | *} | 
| 535 | ||
| 536 | simproc_setup binary_int_div ("number_of m div number_of n :: int") =
 | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 537 |   {* K (divmod_proc (@{thm divmod_rel_div_eq})) *}
 | 
| 24481 | 538 | |
| 539 | simproc_setup binary_int_mod ("number_of m mod number_of n :: int") =
 | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 540 |   {* K (divmod_proc (@{thm divmod_rel_mod_eq})) *}
 | 
| 24481 | 541 | |
| 23164 | 542 | lemmas posDivAlg_eqn_number_of [simp] = | 
| 543 | posDivAlg_eqn [of "number_of v" "number_of w", standard] | |
| 544 | ||
| 545 | lemmas negDivAlg_eqn_number_of [simp] = | |
| 546 | negDivAlg_eqn [of "number_of v" "number_of w", standard] | |
| 547 | ||
| 548 | ||
| 549 | text{*Special-case simplification *}
 | |
| 550 | ||
| 551 | lemma zmod_minus1_right [simp]: "a mod (-1::int) = 0" | |
| 552 | apply (cut_tac a = a and b = "-1" in neg_mod_sign) | |
| 553 | apply (cut_tac [2] a = a and b = "-1" in neg_mod_bound) | |
| 554 | apply (auto simp del: neg_mod_sign neg_mod_bound) | |
| 555 | done | |
| 556 | ||
| 557 | lemma zdiv_minus1_right [simp]: "a div (-1::int) = -a" | |
| 558 | by (cut_tac a = a and b = "-1" in zmod_zdiv_equality, auto) | |
| 559 | ||
| 560 | (** The last remaining special cases for constant arithmetic: | |
| 561 | 1 div z and 1 mod z **) | |
| 562 | ||
| 563 | lemmas div_pos_pos_1_number_of [simp] = | |
| 564 | div_pos_pos [OF int_0_less_1, of "number_of w", standard] | |
| 565 | ||
| 566 | lemmas div_pos_neg_1_number_of [simp] = | |
| 567 | div_pos_neg [OF int_0_less_1, of "number_of w", standard] | |
| 568 | ||
| 569 | lemmas mod_pos_pos_1_number_of [simp] = | |
| 570 | mod_pos_pos [OF int_0_less_1, of "number_of w", standard] | |
| 571 | ||
| 572 | lemmas mod_pos_neg_1_number_of [simp] = | |
| 573 | mod_pos_neg [OF int_0_less_1, of "number_of w", standard] | |
| 574 | ||
| 575 | ||
| 576 | lemmas posDivAlg_eqn_1_number_of [simp] = | |
| 577 | posDivAlg_eqn [of concl: 1 "number_of w", standard] | |
| 578 | ||
| 579 | lemmas negDivAlg_eqn_1_number_of [simp] = | |
| 580 | negDivAlg_eqn [of concl: 1 "number_of w", standard] | |
| 581 | ||
| 582 | ||
| 583 | ||
| 584 | subsection{*Monotonicity in the First Argument (Dividend)*}
 | |
| 585 | ||
| 586 | lemma zdiv_mono1: "[| a \<le> a'; 0 < (b::int) |] ==> a div b \<le> a' div b" | |
| 587 | apply (cut_tac a = a and b = b in zmod_zdiv_equality) | |
| 588 | apply (cut_tac a = a' and b = b in zmod_zdiv_equality) | |
| 589 | apply (rule unique_quotient_lemma) | |
| 590 | apply (erule subst) | |
| 591 | apply (erule subst, simp_all) | |
| 592 | done | |
| 593 | ||
| 594 | lemma zdiv_mono1_neg: "[| a \<le> a'; (b::int) < 0 |] ==> a' div b \<le> a div b" | |
| 595 | apply (cut_tac a = a and b = b in zmod_zdiv_equality) | |
| 596 | apply (cut_tac a = a' and b = b in zmod_zdiv_equality) | |
| 597 | apply (rule unique_quotient_lemma_neg) | |
| 598 | apply (erule subst) | |
| 599 | apply (erule subst, simp_all) | |
| 600 | done | |
| 601 | ||
| 602 | ||
| 603 | subsection{*Monotonicity in the Second Argument (Divisor)*}
 | |
| 604 | ||
| 605 | lemma q_pos_lemma: | |
| 606 | "[| 0 \<le> b'*q' + r'; r' < b'; 0 < b' |] ==> 0 \<le> (q'::int)" | |
| 607 | apply (subgoal_tac "0 < b'* (q' + 1) ") | |
| 608 | apply (simp add: zero_less_mult_iff) | |
| 609 | apply (simp add: right_distrib) | |
| 610 | done | |
| 611 | ||
| 612 | lemma zdiv_mono2_lemma: | |
| 613 | "[| b*q + r = b'*q' + r'; 0 \<le> b'*q' + r'; | |
| 614 | r' < b'; 0 \<le> r; 0 < b'; b' \<le> b |] | |
| 615 | ==> q \<le> (q'::int)" | |
| 616 | apply (frule q_pos_lemma, assumption+) | |
| 617 | apply (subgoal_tac "b*q < b* (q' + 1) ") | |
| 618 | apply (simp add: mult_less_cancel_left) | |
| 619 | apply (subgoal_tac "b*q = r' - r + b'*q'") | |
| 620 | prefer 2 apply simp | |
| 621 | apply (simp (no_asm_simp) add: right_distrib) | |
| 622 | apply (subst add_commute, rule zadd_zless_mono, arith) | |
| 623 | apply (rule mult_right_mono, auto) | |
| 624 | done | |
| 625 | ||
| 626 | lemma zdiv_mono2: | |
| 627 | "[| (0::int) \<le> a; 0 < b'; b' \<le> b |] ==> a div b \<le> a div b'" | |
| 628 | apply (subgoal_tac "b \<noteq> 0") | |
| 629 | prefer 2 apply arith | |
| 630 | apply (cut_tac a = a and b = b in zmod_zdiv_equality) | |
| 631 | apply (cut_tac a = a and b = b' in zmod_zdiv_equality) | |
| 632 | apply (rule zdiv_mono2_lemma) | |
| 633 | apply (erule subst) | |
| 634 | apply (erule subst, simp_all) | |
| 635 | done | |
| 636 | ||
| 637 | lemma q_neg_lemma: | |
| 638 | "[| b'*q' + r' < 0; 0 \<le> r'; 0 < b' |] ==> q' \<le> (0::int)" | |
| 639 | apply (subgoal_tac "b'*q' < 0") | |
| 640 | apply (simp add: mult_less_0_iff, arith) | |
| 641 | done | |
| 642 | ||
| 643 | lemma zdiv_mono2_neg_lemma: | |
| 644 | "[| b*q + r = b'*q' + r'; b'*q' + r' < 0; | |
| 645 | r < b; 0 \<le> r'; 0 < b'; b' \<le> b |] | |
| 646 | ==> q' \<le> (q::int)" | |
| 647 | apply (frule q_neg_lemma, assumption+) | |
| 648 | apply (subgoal_tac "b*q' < b* (q + 1) ") | |
| 649 | apply (simp add: mult_less_cancel_left) | |
| 650 | apply (simp add: right_distrib) | |
| 651 | apply (subgoal_tac "b*q' \<le> b'*q'") | |
| 652 | prefer 2 apply (simp add: mult_right_mono_neg, arith) | |
| 653 | done | |
| 654 | ||
| 655 | lemma zdiv_mono2_neg: | |
| 656 | "[| a < (0::int); 0 < b'; b' \<le> b |] ==> a div b' \<le> a div b" | |
| 657 | apply (cut_tac a = a and b = b in zmod_zdiv_equality) | |
| 658 | apply (cut_tac a = a and b = b' in zmod_zdiv_equality) | |
| 659 | apply (rule zdiv_mono2_neg_lemma) | |
| 660 | apply (erule subst) | |
| 661 | apply (erule subst, simp_all) | |
| 662 | done | |
| 663 | ||
| 25942 | 664 | |
| 23164 | 665 | subsection{*More Algebraic Laws for div and mod*}
 | 
| 666 | ||
| 667 | text{*proving (a*b) div c = a * (b div c) + a * (b mod c) *}
 | |
| 668 | ||
| 669 | lemma zmult1_lemma: | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 670 | "[| divmod_rel b c (q, r); c \<noteq> 0 |] | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 671 | ==> divmod_rel (a * b) c (a*q + a*r div c, a*r mod c)" | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 672 | by (force simp add: split_ifs divmod_rel_def linorder_neq_iff right_distrib) | 
| 23164 | 673 | |
| 674 | lemma zdiv_zmult1_eq: "(a*b) div c = a*(b div c) + a*(b mod c) div (c::int)" | |
| 675 | apply (case_tac "c = 0", simp) | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 676 | apply (blast intro: divmod_rel_div_mod [THEN zmult1_lemma, THEN divmod_rel_div]) | 
| 23164 | 677 | done | 
| 678 | ||
| 679 | lemma zmod_zmult1_eq: "(a*b) mod c = a*(b mod c) mod (c::int)" | |
| 680 | apply (case_tac "c = 0", simp) | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 681 | apply (blast intro: divmod_rel_div_mod [THEN zmult1_lemma, THEN divmod_rel_mod]) | 
| 23164 | 682 | done | 
| 683 | ||
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29045diff
changeset | 684 | lemma zmod_zdiv_trivial: "(a mod b) div b = (0::int)" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 685 | apply (case_tac "b = 0", simp) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 686 | apply (auto simp add: linorder_neq_iff div_pos_pos_trivial div_neg_neg_trivial) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 687 | done | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 688 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 689 | text{*proving (a+b) div c = a div c + b div c + ((a mod c + b mod c) div c) *}
 | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 690 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 691 | lemma zadd1_lemma: | 
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 692 | "[| divmod_rel a c (aq, ar); divmod_rel b c (bq, br); c \<noteq> 0 |] | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 693 | ==> divmod_rel (a+b) c (aq + bq + (ar+br) div c, (ar+br) mod c)" | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 694 | by (force simp add: split_ifs divmod_rel_def linorder_neq_iff right_distrib) | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 695 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 696 | (*NOT suitable for rewriting: the RHS has an instance of the LHS*) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 697 | lemma zdiv_zadd1_eq: | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 698 | "(a+b) div (c::int) = a div c + b div c + ((a mod c + b mod c) div c)" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 699 | apply (case_tac "c = 0", simp) | 
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 700 | apply (blast intro: zadd1_lemma [OF divmod_rel_div_mod divmod_rel_div_mod] divmod_rel_div) | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 701 | done | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 702 | |
| 29405 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 703 | instance int :: ring_div | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 704 | proof | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 705 | fix a b c :: int | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 706 | assume not0: "b \<noteq> 0" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 707 | show "(a + c * b) div b = c + a div b" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 708 | unfolding zdiv_zadd1_eq [of a "c * b"] using not0 | 
| 30181 | 709 | by (simp add: zmod_zmult1_eq zmod_zdiv_trivial zdiv_zmult1_eq) | 
| 30930 | 710 | next | 
| 711 | fix a b c :: int | |
| 712 | assume "a \<noteq> 0" | |
| 713 | then show "(a * b) div (a * c) = b div c" | |
| 714 | proof (cases "b \<noteq> 0 \<and> c \<noteq> 0") | |
| 715 | case False then show ?thesis by auto | |
| 716 | next | |
| 717 | case True then have "b \<noteq> 0" and "c \<noteq> 0" by auto | |
| 718 | with `a \<noteq> 0` | |
| 719 | have "\<And>q r. divmod_rel b c (q, r) \<Longrightarrow> divmod_rel (a * b) (a * c) (q, a * r)" | |
| 720 | apply (auto simp add: divmod_rel_def) | |
| 721 | apply (auto simp add: algebra_simps) | |
| 722 | apply (auto simp add: zero_less_mult_iff zero_le_mult_iff mult_le_0_iff) | |
| 723 | done | |
| 724 | moreover with `c \<noteq> 0` divmod_rel_div_mod have "divmod_rel b c (b div c, b mod c)" by auto | |
| 725 | ultimately have "divmod_rel (a * b) (a * c) (b div c, a * (b mod c))" . | |
| 726 | moreover from `a \<noteq> 0` `c \<noteq> 0` have "a * c \<noteq> 0" by simp | |
| 727 | ultimately show ?thesis by (rule divmod_rel_div) | |
| 728 | qed | |
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 729 | qed auto | 
| 25942 | 730 | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 731 | lemma posDivAlg_div_mod: | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 732 | assumes "k \<ge> 0" | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 733 | and "l \<ge> 0" | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 734 | shows "posDivAlg k l = (k div l, k mod l)" | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 735 | proof (cases "l = 0") | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 736 | case True then show ?thesis by (simp add: posDivAlg.simps) | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 737 | next | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 738 | case False with assms posDivAlg_correct | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 739 | have "divmod_rel k l (fst (posDivAlg k l), snd (posDivAlg k l))" | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 740 | by simp | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 741 | from divmod_rel_div [OF this `l \<noteq> 0`] divmod_rel_mod [OF this `l \<noteq> 0`] | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 742 | show ?thesis by simp | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 743 | qed | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 744 | |
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 745 | lemma negDivAlg_div_mod: | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 746 | assumes "k < 0" | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 747 | and "l > 0" | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 748 | shows "negDivAlg k l = (k div l, k mod l)" | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 749 | proof - | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 750 | from assms have "l \<noteq> 0" by simp | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 751 | from assms negDivAlg_correct | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 752 | have "divmod_rel k l (fst (negDivAlg k l), snd (negDivAlg k l))" | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 753 | by simp | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 754 | from divmod_rel_div [OF this `l \<noteq> 0`] divmod_rel_mod [OF this `l \<noteq> 0`] | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 755 | show ?thesis by simp | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 756 | qed | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 757 | |
| 23164 | 758 | lemma zmod_eq_0_iff: "(m mod d = 0) = (EX q::int. m = d*q)" | 
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29045diff
changeset | 759 | by (simp add: dvd_eq_mod_eq_0 [symmetric] dvd_def) | 
| 23164 | 760 | |
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29045diff
changeset | 761 | (* REVISIT: should this be generalized to all semiring_div types? *) | 
| 23164 | 762 | lemmas zmod_eq_0D [dest!] = zmod_eq_0_iff [THEN iffD1] | 
| 763 | ||
| 23983 | 764 | |
| 23164 | 765 | subsection{*Proving  @{term "a div (b*c) = (a div b) div c"} *}
 | 
| 766 | ||
| 767 | (*The condition c>0 seems necessary. Consider that 7 div ~6 = ~2 but | |
| 768 | 7 div 2 div ~3 = 3 div ~3 = ~1. The subcase (a div b) mod c = 0 seems | |
| 769 | to cause particular problems.*) | |
| 770 | ||
| 771 | text{*first, four lemmas to bound the remainder for the cases b<0 and b>0 *}
 | |
| 772 | ||
| 773 | lemma zmult2_lemma_aux1: "[| (0::int) < c; b < r; r \<le> 0 |] ==> b*c < b*(q mod c) + r" | |
| 774 | apply (subgoal_tac "b * (c - q mod c) < r * 1") | |
| 29667 | 775 | apply (simp add: algebra_simps) | 
| 23164 | 776 | apply (rule order_le_less_trans) | 
| 29667 | 777 | apply (erule_tac [2] mult_strict_right_mono) | 
| 778 | apply (rule mult_left_mono_neg) | |
| 779 | using add1_zle_eq[of "q mod c"]apply(simp add: algebra_simps pos_mod_bound) | |
| 780 | apply (simp) | |
| 781 | apply (simp) | |
| 23164 | 782 | done | 
| 783 | ||
| 784 | lemma zmult2_lemma_aux2: | |
| 785 | "[| (0::int) < c; b < r; r \<le> 0 |] ==> b * (q mod c) + r \<le> 0" | |
| 786 | apply (subgoal_tac "b * (q mod c) \<le> 0") | |
| 787 | apply arith | |
| 788 | apply (simp add: mult_le_0_iff) | |
| 789 | done | |
| 790 | ||
| 791 | lemma zmult2_lemma_aux3: "[| (0::int) < c; 0 \<le> r; r < b |] ==> 0 \<le> b * (q mod c) + r" | |
| 792 | apply (subgoal_tac "0 \<le> b * (q mod c) ") | |
| 793 | apply arith | |
| 794 | apply (simp add: zero_le_mult_iff) | |
| 795 | done | |
| 796 | ||
| 797 | lemma zmult2_lemma_aux4: "[| (0::int) < c; 0 \<le> r; r < b |] ==> b * (q mod c) + r < b * c" | |
| 798 | apply (subgoal_tac "r * 1 < b * (c - q mod c) ") | |
| 29667 | 799 | apply (simp add: right_diff_distrib) | 
| 23164 | 800 | apply (rule order_less_le_trans) | 
| 29667 | 801 | apply (erule mult_strict_right_mono) | 
| 802 | apply (rule_tac [2] mult_left_mono) | |
| 803 | apply simp | |
| 804 | using add1_zle_eq[of "q mod c"] apply (simp add: algebra_simps pos_mod_bound) | |
| 805 | apply simp | |
| 23164 | 806 | done | 
| 807 | ||
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 808 | lemma zmult2_lemma: "[| divmod_rel a b (q, r); b \<noteq> 0; 0 < c |] | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 809 | ==> divmod_rel a (b * c) (q div c, b*(q mod c) + r)" | 
| 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 810 | by (auto simp add: mult_ac divmod_rel_def linorder_neq_iff | 
| 23164 | 811 | zero_less_mult_iff right_distrib [symmetric] | 
| 812 | zmult2_lemma_aux1 zmult2_lemma_aux2 zmult2_lemma_aux3 zmult2_lemma_aux4) | |
| 813 | ||
| 814 | lemma zdiv_zmult2_eq: "(0::int) < c ==> a div (b*c) = (a div b) div c" | |
| 815 | apply (case_tac "b = 0", simp) | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 816 | apply (force simp add: divmod_rel_div_mod [THEN zmult2_lemma, THEN divmod_rel_div]) | 
| 23164 | 817 | done | 
| 818 | ||
| 819 | lemma zmod_zmult2_eq: | |
| 820 | "(0::int) < c ==> a mod (b*c) = b*(a div b mod c) + a mod b" | |
| 821 | apply (case_tac "b = 0", simp) | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 822 | apply (force simp add: divmod_rel_div_mod [THEN zmult2_lemma, THEN divmod_rel_mod]) | 
| 23164 | 823 | done | 
| 824 | ||
| 825 | ||
| 826 | subsection {*Splitting Rules for div and mod*}
 | |
| 827 | ||
| 828 | text{*The proofs of the two lemmas below are essentially identical*}
 | |
| 829 | ||
| 830 | lemma split_pos_lemma: | |
| 831 | "0<k ==> | |
| 832 | P(n div k :: int)(n mod k) = (\<forall>i j. 0\<le>j & j<k & n = k*i + j --> P i j)" | |
| 833 | apply (rule iffI, clarify) | |
| 834 | apply (erule_tac P="P ?x ?y" in rev_mp) | |
| 29948 | 835 | apply (subst mod_add_eq) | 
| 23164 | 836 | apply (subst zdiv_zadd1_eq) | 
| 837 | apply (simp add: div_pos_pos_trivial mod_pos_pos_trivial) | |
| 838 | txt{*converse direction*}
 | |
| 839 | apply (drule_tac x = "n div k" in spec) | |
| 840 | apply (drule_tac x = "n mod k" in spec, simp) | |
| 841 | done | |
| 842 | ||
| 843 | lemma split_neg_lemma: | |
| 844 | "k<0 ==> | |
| 845 | P(n div k :: int)(n mod k) = (\<forall>i j. k<j & j\<le>0 & n = k*i + j --> P i j)" | |
| 846 | apply (rule iffI, clarify) | |
| 847 | apply (erule_tac P="P ?x ?y" in rev_mp) | |
| 29948 | 848 | apply (subst mod_add_eq) | 
| 23164 | 849 | apply (subst zdiv_zadd1_eq) | 
| 850 | apply (simp add: div_neg_neg_trivial mod_neg_neg_trivial) | |
| 851 | txt{*converse direction*}
 | |
| 852 | apply (drule_tac x = "n div k" in spec) | |
| 853 | apply (drule_tac x = "n mod k" in spec, simp) | |
| 854 | done | |
| 855 | ||
| 856 | lemma split_zdiv: | |
| 857 | "P(n div k :: int) = | |
| 858 | ((k = 0 --> P 0) & | |
| 859 | (0<k --> (\<forall>i j. 0\<le>j & j<k & n = k*i + j --> P i)) & | |
| 860 | (k<0 --> (\<forall>i j. k<j & j\<le>0 & n = k*i + j --> P i)))" | |
| 861 | apply (case_tac "k=0", simp) | |
| 862 | apply (simp only: linorder_neq_iff) | |
| 863 | apply (erule disjE) | |
| 864 | apply (simp_all add: split_pos_lemma [of concl: "%x y. P x"] | |
| 865 | split_neg_lemma [of concl: "%x y. P x"]) | |
| 866 | done | |
| 867 | ||
| 868 | lemma split_zmod: | |
| 869 | "P(n mod k :: int) = | |
| 870 | ((k = 0 --> P n) & | |
| 871 | (0<k --> (\<forall>i j. 0\<le>j & j<k & n = k*i + j --> P j)) & | |
| 872 | (k<0 --> (\<forall>i j. k<j & j\<le>0 & n = k*i + j --> P j)))" | |
| 873 | apply (case_tac "k=0", simp) | |
| 874 | apply (simp only: linorder_neq_iff) | |
| 875 | apply (erule disjE) | |
| 876 | apply (simp_all add: split_pos_lemma [of concl: "%x y. P y"] | |
| 877 | split_neg_lemma [of concl: "%x y. P y"]) | |
| 878 | done | |
| 879 | ||
| 880 | (* Enable arith to deal with div 2 and mod 2: *) | |
| 881 | declare split_zdiv [of _ _ "number_of k", simplified, standard, arith_split] | |
| 882 | declare split_zmod [of _ _ "number_of k", simplified, standard, arith_split] | |
| 883 | ||
| 884 | ||
| 885 | subsection{*Speeding up the Division Algorithm with Shifting*}
 | |
| 886 | ||
| 887 | text{*computing div by shifting *}
 | |
| 888 | ||
| 889 | lemma pos_zdiv_mult_2: "(0::int) \<le> a ==> (1 + 2*b) div (2*a) = b div a" | |
| 890 | proof cases | |
| 891 | assume "a=0" | |
| 892 | thus ?thesis by simp | |
| 893 | next | |
| 894 | assume "a\<noteq>0" and le_a: "0\<le>a" | |
| 895 | hence a_pos: "1 \<le> a" by arith | |
| 30652 
752329615264
distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
 haftmann parents: 
30517diff
changeset | 896 | hence one_less_a2: "1 < 2 * a" by arith | 
| 23164 | 897 | hence le_2a: "2 * (1 + b mod a) \<le> 2 * a" | 
| 30652 
752329615264
distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
 haftmann parents: 
30517diff
changeset | 898 | unfolding mult_le_cancel_left | 
| 
752329615264
distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
 haftmann parents: 
30517diff
changeset | 899 | by (simp add: add1_zle_eq add_commute [of 1]) | 
| 23164 | 900 | with a_pos have "0 \<le> b mod a" by simp | 
| 901 | hence le_addm: "0 \<le> 1 mod (2*a) + 2*(b mod a)" | |
| 902 | by (simp add: mod_pos_pos_trivial one_less_a2) | |
| 903 | with le_2a | |
| 904 | have "(1 mod (2*a) + 2*(b mod a)) div (2*a) = 0" | |
| 905 | by (simp add: div_pos_pos_trivial le_addm mod_pos_pos_trivial one_less_a2 | |
| 906 | right_distrib) | |
| 907 | thus ?thesis | |
| 908 | by (subst zdiv_zadd1_eq, | |
| 30930 | 909 | simp add: mod_mult_mult1 one_less_a2 | 
| 23164 | 910 | div_pos_pos_trivial) | 
| 911 | qed | |
| 912 | ||
| 913 | lemma neg_zdiv_mult_2: "a \<le> (0::int) ==> (1 + 2*b) div (2*a) = (b+1) div a" | |
| 914 | apply (subgoal_tac " (1 + 2* (-b - 1)) div (2 * (-a)) = (-b - 1) div (-a) ") | |
| 915 | apply (rule_tac [2] pos_zdiv_mult_2) | |
| 916 | apply (auto simp add: minus_mult_right [symmetric] right_diff_distrib) | |
| 917 | apply (subgoal_tac " (-1 - (2 * b)) = - (1 + (2 * b))") | |
| 918 | apply (simp only: zdiv_zminus_zminus diff_minus minus_add_distrib [symmetric], | |
| 919 | simp) | |
| 920 | done | |
| 921 | ||
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25961diff
changeset | 922 | lemma zdiv_number_of_Bit0 [simp]: | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25961diff
changeset | 923 | "number_of (Int.Bit0 v) div number_of (Int.Bit0 w) = | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25961diff
changeset | 924 | number_of v div (number_of w :: int)" | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25961diff
changeset | 925 | by (simp only: number_of_eq numeral_simps) simp | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25961diff
changeset | 926 | |
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25961diff
changeset | 927 | lemma zdiv_number_of_Bit1 [simp]: | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25961diff
changeset | 928 | "number_of (Int.Bit1 v) div number_of (Int.Bit0 w) = | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25961diff
changeset | 929 | (if (0::int) \<le> number_of w | 
| 23164 | 930 | then number_of v div (number_of w) | 
| 931 | else (number_of v + (1::int)) div (number_of w))" | |
| 932 | apply (simp only: number_of_eq numeral_simps UNIV_I split: split_if) | |
| 30930 | 933 | apply (simp add: pos_zdiv_mult_2 neg_zdiv_mult_2 add_ac) | 
| 23164 | 934 | done | 
| 935 | ||
| 936 | ||
| 937 | subsection{*Computing mod by Shifting (proofs resemble those for div)*}
 | |
| 938 | ||
| 939 | lemma pos_zmod_mult_2: | |
| 940 | "(0::int) \<le> a ==> (1 + 2*b) mod (2*a) = 1 + 2 * (b mod a)" | |
| 941 | apply (case_tac "a = 0", simp) | |
| 942 | apply (subgoal_tac "1 < a * 2") | |
| 943 | prefer 2 apply arith | |
| 944 | apply (subgoal_tac "2* (1 + b mod a) \<le> 2*a") | |
| 945 | apply (rule_tac [2] mult_left_mono) | |
| 946 | apply (auto simp add: add_commute [of 1] mult_commute add1_zle_eq | |
| 947 | pos_mod_bound) | |
| 29948 | 948 | apply (subst mod_add_eq) | 
| 30930 | 949 | apply (simp add: mod_mult_mult2 mod_pos_pos_trivial) | 
| 23164 | 950 | apply (rule mod_pos_pos_trivial) | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25961diff
changeset | 951 | apply (auto simp add: mod_pos_pos_trivial ring_distribs) | 
| 23164 | 952 | apply (subgoal_tac "0 \<le> b mod a", arith, simp) | 
| 953 | done | |
| 954 | ||
| 955 | lemma neg_zmod_mult_2: | |
| 956 | "a \<le> (0::int) ==> (1 + 2*b) mod (2*a) = 2 * ((b+1) mod a) - 1" | |
| 957 | apply (subgoal_tac "(1 + 2* (-b - 1)) mod (2* (-a)) = | |
| 958 | 1 + 2* ((-b - 1) mod (-a))") | |
| 959 | apply (rule_tac [2] pos_zmod_mult_2) | |
| 30042 | 960 | apply (auto simp add: right_diff_distrib) | 
| 23164 | 961 | apply (subgoal_tac " (-1 - (2 * b)) = - (1 + (2 * b))") | 
| 962 | prefer 2 apply simp | |
| 963 | apply (simp only: zmod_zminus_zminus diff_minus minus_add_distrib [symmetric]) | |
| 964 | done | |
| 965 | ||
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25961diff
changeset | 966 | lemma zmod_number_of_Bit0 [simp]: | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25961diff
changeset | 967 | "number_of (Int.Bit0 v) mod number_of (Int.Bit0 w) = | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25961diff
changeset | 968 | (2::int) * (number_of v mod number_of w)" | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25961diff
changeset | 969 | apply (simp only: number_of_eq numeral_simps) | 
| 30930 | 970 | apply (simp add: mod_mult_mult1 pos_zmod_mult_2 | 
| 29948 | 971 | neg_zmod_mult_2 add_ac) | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25961diff
changeset | 972 | done | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25961diff
changeset | 973 | |
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25961diff
changeset | 974 | lemma zmod_number_of_Bit1 [simp]: | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25961diff
changeset | 975 | "number_of (Int.Bit1 v) mod number_of (Int.Bit0 w) = | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25961diff
changeset | 976 | (if (0::int) \<le> number_of w | 
| 23164 | 977 | then 2 * (number_of v mod number_of w) + 1 | 
| 978 | else 2 * ((number_of v + (1::int)) mod number_of w) - 1)" | |
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25961diff
changeset | 979 | apply (simp only: number_of_eq numeral_simps) | 
| 30930 | 980 | apply (simp add: mod_mult_mult1 pos_zmod_mult_2 | 
| 29948 | 981 | neg_zmod_mult_2 add_ac) | 
| 23164 | 982 | done | 
| 983 | ||
| 984 | ||
| 985 | subsection{*Quotients of Signs*}
 | |
| 986 | ||
| 987 | lemma div_neg_pos_less0: "[| a < (0::int); 0 < b |] ==> a div b < 0" | |
| 988 | apply (subgoal_tac "a div b \<le> -1", force) | |
| 989 | apply (rule order_trans) | |
| 990 | apply (rule_tac a' = "-1" in zdiv_mono1) | |
| 29948 | 991 | apply (auto simp add: div_eq_minus1) | 
| 23164 | 992 | done | 
| 993 | ||
| 30323 | 994 | lemma div_nonneg_neg_le0: "[| (0::int) \<le> a; b < 0 |] ==> a div b \<le> 0" | 
| 23164 | 995 | by (drule zdiv_mono1_neg, auto) | 
| 996 | ||
| 30323 | 997 | lemma div_nonpos_pos_le0: "[| (a::int) \<le> 0; b > 0 |] ==> a div b \<le> 0" | 
| 998 | by (drule zdiv_mono1, auto) | |
| 999 | ||
| 23164 | 1000 | lemma pos_imp_zdiv_nonneg_iff: "(0::int) < b ==> (0 \<le> a div b) = (0 \<le> a)" | 
| 1001 | apply auto | |
| 1002 | apply (drule_tac [2] zdiv_mono1) | |
| 1003 | apply (auto simp add: linorder_neq_iff) | |
| 1004 | apply (simp (no_asm_use) add: linorder_not_less [symmetric]) | |
| 1005 | apply (blast intro: div_neg_pos_less0) | |
| 1006 | done | |
| 1007 | ||
| 1008 | lemma neg_imp_zdiv_nonneg_iff: | |
| 1009 | "b < (0::int) ==> (0 \<le> a div b) = (a \<le> (0::int))" | |
| 1010 | apply (subst zdiv_zminus_zminus [symmetric]) | |
| 1011 | apply (subst pos_imp_zdiv_nonneg_iff, auto) | |
| 1012 | done | |
| 1013 | ||
| 1014 | (*But not (a div b \<le> 0 iff a\<le>0); consider a=1, b=2 when a div b = 0.*) | |
| 1015 | lemma pos_imp_zdiv_neg_iff: "(0::int) < b ==> (a div b < 0) = (a < 0)" | |
| 1016 | by (simp add: linorder_not_le [symmetric] pos_imp_zdiv_nonneg_iff) | |
| 1017 | ||
| 1018 | (*Again the law fails for \<le>: consider a = -1, b = -2 when a div b = 0*) | |
| 1019 | lemma neg_imp_zdiv_neg_iff: "b < (0::int) ==> (a div b < 0) = (0 < a)" | |
| 1020 | by (simp add: linorder_not_le [symmetric] neg_imp_zdiv_nonneg_iff) | |
| 1021 | ||
| 1022 | ||
| 1023 | subsection {* The Divides Relation *}
 | |
| 1024 | ||
| 1025 | lemmas zdvd_iff_zmod_eq_0_number_of [simp] = | |
| 30042 | 1026 | dvd_eq_mod_eq_0 [of "number_of x::int" "number_of y::int", standard] | 
| 23164 | 1027 | |
| 1028 | lemma zdvd_anti_sym: | |
| 1029 | "0 < m ==> 0 < n ==> m dvd n ==> n dvd m ==> m = (n::int)" | |
| 1030 | apply (simp add: dvd_def, auto) | |
| 1031 | apply (simp add: mult_assoc zero_less_mult_iff zmult_eq_1_iff) | |
| 1032 | done | |
| 1033 | ||
| 30042 | 1034 | lemma zdvd_dvd_eq: assumes "a \<noteq> 0" and "(a::int) dvd b" and "b dvd a" | 
| 23164 | 1035 | shows "\<bar>a\<bar> = \<bar>b\<bar>" | 
| 1036 | proof- | |
| 30042 | 1037 | from `a dvd b` obtain k where k:"b = a*k" unfolding dvd_def by blast | 
| 1038 | from `b dvd a` obtain k' where k':"a = b*k'" unfolding dvd_def by blast | |
| 23164 | 1039 | from k k' have "a = a*k*k'" by simp | 
| 1040 | with mult_cancel_left1[where c="a" and b="k*k'"] | |
| 30042 | 1041 | have kk':"k*k' = 1" using `a\<noteq>0` by (simp add: mult_assoc) | 
| 23164 | 1042 | hence "k = 1 \<and> k' = 1 \<or> k = -1 \<and> k' = -1" by (simp add: zmult_eq_1_iff) | 
| 1043 | thus ?thesis using k k' by auto | |
| 1044 | qed | |
| 1045 | ||
| 1046 | lemma zdvd_zdiffD: "k dvd m - n ==> k dvd n ==> k dvd (m::int)" | |
| 1047 | apply (subgoal_tac "m = n + (m - n)") | |
| 1048 | apply (erule ssubst) | |
| 30042 | 1049 | apply (blast intro: dvd_add, simp) | 
| 23164 | 1050 | done | 
| 1051 | ||
| 1052 | lemma zdvd_reduce: "(k dvd n + k * m) = (k dvd (n::int))" | |
| 30042 | 1053 | apply (rule iffI) | 
| 1054 | apply (erule_tac [2] dvd_add) | |
| 1055 | apply (subgoal_tac "n = (n + k * m) - k * m") | |
| 1056 | apply (erule ssubst) | |
| 1057 | apply (erule dvd_diff) | |
| 1058 | apply(simp_all) | |
| 1059 | done | |
| 23164 | 1060 | |
| 1061 | lemma zdvd_zmod: "f dvd m ==> f dvd (n::int) ==> f dvd m mod n" | |
| 30930 | 1062 | by (auto elim!: dvdE simp add: mod_mult_mult1) | 
| 23164 | 1063 | |
| 1064 | lemma zdvd_zmod_imp_zdvd: "k dvd m mod n ==> k dvd n ==> k dvd (m::int)" | |
| 1065 | apply (subgoal_tac "k dvd n * (m div n) + m mod n") | |
| 1066 | apply (simp add: zmod_zdiv_equality [symmetric]) | |
| 30042 | 1067 | apply (simp only: dvd_add dvd_mult2) | 
| 23164 | 1068 | done | 
| 1069 | ||
| 1070 | lemma zdvd_not_zless: "0 < m ==> m < n ==> \<not> n dvd (m::int)" | |
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 1071 | apply (auto elim!: dvdE) | 
| 23164 | 1072 | apply (subgoal_tac "0 < n") | 
| 1073 | prefer 2 | |
| 1074 | apply (blast intro: order_less_trans) | |
| 1075 | apply (simp add: zero_less_mult_iff) | |
| 1076 | done | |
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 1077 | |
| 23164 | 1078 | lemma zmult_div_cancel: "(n::int) * (m div n) = m - (m mod n)" | 
| 1079 | using zmod_zdiv_equality[where a="m" and b="n"] | |
| 29667 | 1080 | by (simp add: algebra_simps) | 
| 23164 | 1081 | |
| 1082 | lemma zdvd_mult_div_cancel:"(n::int) dvd m \<Longrightarrow> n * (m div n) = m" | |
| 1083 | apply (subgoal_tac "m mod n = 0") | |
| 1084 | apply (simp add: zmult_div_cancel) | |
| 30042 | 1085 | apply (simp only: dvd_eq_mod_eq_0) | 
| 23164 | 1086 | done | 
| 1087 | ||
| 1088 | lemma zdvd_mult_cancel: assumes d:"k * m dvd k * n" and kz:"k \<noteq> (0::int)" | |
| 1089 | shows "m dvd n" | |
| 1090 | proof- | |
| 1091 | from d obtain h where h: "k*n = k*m * h" unfolding dvd_def by blast | |
| 1092 |   {assume "n \<noteq> m*h" hence "k* n \<noteq> k* (m*h)" using kz by simp
 | |
| 1093 | with h have False by (simp add: mult_assoc)} | |
| 1094 | hence "n = m * h" by blast | |
| 29410 | 1095 | thus ?thesis by simp | 
| 23164 | 1096 | qed | 
| 1097 | ||
| 23969 | 1098 | |
| 23164 | 1099 | theorem ex_nat: "(\<exists>x::nat. P x) = (\<exists>x::int. 0 <= x \<and> P (nat x))" | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1100 | apply (simp split add: split_nat) | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1101 | apply (rule iffI) | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1102 | apply (erule exE) | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1103 | apply (rule_tac x = "int x" in exI) | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1104 | apply simp | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1105 | apply (erule exE) | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1106 | apply (rule_tac x = "nat x" in exI) | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1107 | apply (erule conjE) | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1108 | apply (erule_tac x = "nat x" in allE) | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1109 | apply simp | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1110 | done | 
| 23164 | 1111 | |
| 23365 | 1112 | theorem zdvd_int: "(x dvd y) = (int x dvd int y)" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 1113 | proof - | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 1114 | have "\<And>k. int y = int x * k \<Longrightarrow> x dvd y" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 1115 | proof - | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 1116 | fix k | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 1117 | assume A: "int y = int x * k" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 1118 | then show "x dvd y" proof (cases k) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 1119 | case (1 n) with A have "y = x * n" by (simp add: zmult_int) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 1120 | then show ?thesis .. | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 1121 | next | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 1122 | case (2 n) with A have "int y = int x * (- int (Suc n))" by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 1123 | also have "\<dots> = - (int x * int (Suc n))" by (simp only: mult_minus_right) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 1124 | also have "\<dots> = - int (x * Suc n)" by (simp only: zmult_int) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 1125 | finally have "- int (x * Suc n) = int y" .. | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 1126 | then show ?thesis by (simp only: negative_eq_positive) auto | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 1127 | qed | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 1128 | qed | 
| 30042 | 1129 | then show ?thesis by (auto elim!: dvdE simp only: dvd_triv_left int_mult) | 
| 29410 | 1130 | qed | 
| 23164 | 1131 | |
| 1132 | lemma zdvd1_eq[simp]: "(x::int) dvd 1 = ( \<bar>x\<bar> = 1)" | |
| 1133 | proof | |
| 30042 | 1134 | assume d: "x dvd 1" hence "int (nat \<bar>x\<bar>) dvd int (nat 1)" by simp | 
| 23164 | 1135 | hence "nat \<bar>x\<bar> dvd 1" by (simp add: zdvd_int) | 
| 1136 | hence "nat \<bar>x\<bar> = 1" by simp | |
| 1137 | thus "\<bar>x\<bar> = 1" by (cases "x < 0", auto) | |
| 1138 | next | |
| 1139 | assume "\<bar>x\<bar>=1" thus "x dvd 1" | |
| 30042 | 1140 | by(cases "x < 0",simp_all add: minus_equation_iff dvd_eq_mod_eq_0) | 
| 23164 | 1141 | qed | 
| 1142 | lemma zdvd_mult_cancel1: | |
| 1143 | assumes mp:"m \<noteq>(0::int)" shows "(m * n dvd m) = (\<bar>n\<bar> = 1)" | |
| 1144 | proof | |
| 1145 | assume n1: "\<bar>n\<bar> = 1" thus "m * n dvd m" | |
| 30042 | 1146 | by (cases "n >0", auto simp add: minus_dvd_iff minus_equation_iff) | 
| 23164 | 1147 | next | 
| 1148 | assume H: "m * n dvd m" hence H2: "m * n dvd m * 1" by simp | |
| 1149 | from zdvd_mult_cancel[OF H2 mp] show "\<bar>n\<bar> = 1" by (simp only: zdvd1_eq) | |
| 1150 | qed | |
| 1151 | ||
| 23365 | 1152 | lemma int_dvd_iff: "(int m dvd z) = (m dvd nat (abs z))" | 
| 30042 | 1153 | unfolding zdvd_int by (cases "z \<ge> 0") simp_all | 
| 23306 
cdb027d0637e
add int_of_nat versions of lemmas about int::nat=>int
 huffman parents: 
23164diff
changeset | 1154 | |
| 23365 | 1155 | lemma dvd_int_iff: "(z dvd int m) = (nat (abs z) dvd m)" | 
| 30042 | 1156 | unfolding zdvd_int by (cases "z \<ge> 0") simp_all | 
| 23164 | 1157 | |
| 1158 | lemma nat_dvd_iff: "(nat z dvd m) = (if 0 \<le> z then (z dvd int m) else m = 0)" | |
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
26507diff
changeset | 1159 | by (auto simp add: dvd_int_iff) | 
| 23164 | 1160 | |
| 1161 | lemma zdvd_imp_le: "[| z dvd n; 0 < n |] ==> z \<le> (n::int)" | |
| 23365 | 1162 | apply (rule_tac z=n in int_cases) | 
| 1163 | apply (auto simp add: dvd_int_iff) | |
| 1164 | apply (rule_tac z=z in int_cases) | |
| 23307 
2fe3345035c7
modify proofs to avoid referring to int::nat=>int
 huffman parents: 
23306diff
changeset | 1165 | apply (auto simp add: dvd_imp_le) | 
| 23164 | 1166 | done | 
| 1167 | ||
| 1168 | lemma zpower_zmod: "((x::int) mod m)^y mod m = x^y mod m" | |
| 1169 | apply (induct "y", auto) | |
| 1170 | apply (rule zmod_zmult1_eq [THEN trans]) | |
| 1171 | apply (simp (no_asm_simp)) | |
| 29948 | 1172 | apply (rule mod_mult_eq [symmetric]) | 
| 23164 | 1173 | done | 
| 1174 | ||
| 23365 | 1175 | lemma zdiv_int: "int (a div b) = (int a) div (int b)" | 
| 23164 | 1176 | apply (subst split_div, auto) | 
| 1177 | apply (subst split_zdiv, auto) | |
| 23365 | 1178 | apply (rule_tac a="int (b * i) + int j" and b="int b" and r="int j" and r'=ja in IntDiv.unique_quotient) | 
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 1179 | apply (auto simp add: IntDiv.divmod_rel_def of_nat_mult) | 
| 23164 | 1180 | done | 
| 1181 | ||
| 1182 | lemma zmod_int: "int (a mod b) = (int a) mod (int b)" | |
| 23365 | 1183 | apply (subst split_mod, auto) | 
| 1184 | apply (subst split_zmod, auto) | |
| 1185 | apply (rule_tac a="int (b * i) + int j" and b="int b" and q="int i" and q'=ia | |
| 1186 | in unique_remainder) | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 1187 | apply (auto simp add: IntDiv.divmod_rel_def of_nat_mult) | 
| 23365 | 1188 | done | 
| 23164 | 1189 | |
| 30180 | 1190 | lemma abs_div: "(y::int) dvd x \<Longrightarrow> abs (x div y) = abs x div abs y" | 
| 1191 | by (unfold dvd_def, cases "y=0", auto simp add: abs_mult) | |
| 1192 | ||
| 23164 | 1193 | text{*Suggested by Matthias Daum*}
 | 
| 1194 | lemma int_power_div_base: | |
| 1195 | "\<lbrakk>0 < m; 0 < k\<rbrakk> \<Longrightarrow> k ^ m div k = (k::int) ^ (m - Suc 0)" | |
| 30079 
293b896b9c25
make proofs work whether or not One_nat_def is a simp rule; replace 1 with Suc 0 in the rhs of some simp rules
 huffman parents: 
30042diff
changeset | 1196 | apply (subgoal_tac "k ^ m = k ^ ((m - Suc 0) + Suc 0)") | 
| 23164 | 1197 | apply (erule ssubst) | 
| 1198 | apply (simp only: power_add) | |
| 1199 | apply simp_all | |
| 1200 | done | |
| 1201 | ||
| 23853 | 1202 | text {* by Brian Huffman *}
 | 
| 1203 | lemma zminus_zmod: "- ((x::int) mod m) mod m = - x mod m" | |
| 29405 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 1204 | by (rule mod_minus_eq [symmetric]) | 
| 23853 | 1205 | |
| 1206 | lemma zdiff_zmod_left: "(x mod m - y) mod m = (x - y) mod (m::int)" | |
| 29405 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 1207 | by (rule mod_diff_left_eq [symmetric]) | 
| 23853 | 1208 | |
| 1209 | lemma zdiff_zmod_right: "(x - y mod m) mod m = (x - y) mod (m::int)" | |
| 29405 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 1210 | by (rule mod_diff_right_eq [symmetric]) | 
| 23853 | 1211 | |
| 1212 | lemmas zmod_simps = | |
| 30034 | 1213 | mod_add_left_eq [symmetric] | 
| 1214 | mod_add_right_eq [symmetric] | |
| 30930 | 1215 | zmod_zmult1_eq [symmetric] | 
| 1216 | mod_mult_left_eq [symmetric] | |
| 1217 | zpower_zmod | |
| 23853 | 1218 | zminus_zmod zdiff_zmod_left zdiff_zmod_right | 
| 1219 | ||
| 29045 
3c8f48333731
move nat_{div,mod}_distrib from NatBin to IntDiv, simplified proofs
 huffman parents: 
28562diff
changeset | 1220 | text {* Distributive laws for function @{text nat}. *}
 | 
| 
3c8f48333731
move nat_{div,mod}_distrib from NatBin to IntDiv, simplified proofs
 huffman parents: 
28562diff
changeset | 1221 | |
| 
3c8f48333731
move nat_{div,mod}_distrib from NatBin to IntDiv, simplified proofs
 huffman parents: 
28562diff
changeset | 1222 | lemma nat_div_distrib: "0 \<le> x \<Longrightarrow> nat (x div y) = nat x div nat y" | 
| 
3c8f48333731
move nat_{div,mod}_distrib from NatBin to IntDiv, simplified proofs
 huffman parents: 
28562diff
changeset | 1223 | apply (rule linorder_cases [of y 0]) | 
| 
3c8f48333731
move nat_{div,mod}_distrib from NatBin to IntDiv, simplified proofs
 huffman parents: 
28562diff
changeset | 1224 | apply (simp add: div_nonneg_neg_le0) | 
| 
3c8f48333731
move nat_{div,mod}_distrib from NatBin to IntDiv, simplified proofs
 huffman parents: 
28562diff
changeset | 1225 | apply simp | 
| 
3c8f48333731
move nat_{div,mod}_distrib from NatBin to IntDiv, simplified proofs
 huffman parents: 
28562diff
changeset | 1226 | apply (simp add: nat_eq_iff pos_imp_zdiv_nonneg_iff zdiv_int) | 
| 
3c8f48333731
move nat_{div,mod}_distrib from NatBin to IntDiv, simplified proofs
 huffman parents: 
28562diff
changeset | 1227 | done | 
| 
3c8f48333731
move nat_{div,mod}_distrib from NatBin to IntDiv, simplified proofs
 huffman parents: 
28562diff
changeset | 1228 | |
| 
3c8f48333731
move nat_{div,mod}_distrib from NatBin to IntDiv, simplified proofs
 huffman parents: 
28562diff
changeset | 1229 | (*Fails if y<0: the LHS collapses to (nat z) but the RHS doesn't*) | 
| 
3c8f48333731
move nat_{div,mod}_distrib from NatBin to IntDiv, simplified proofs
 huffman parents: 
28562diff
changeset | 1230 | lemma nat_mod_distrib: | 
| 
3c8f48333731
move nat_{div,mod}_distrib from NatBin to IntDiv, simplified proofs
 huffman parents: 
28562diff
changeset | 1231 | "\<lbrakk>0 \<le> x; 0 \<le> y\<rbrakk> \<Longrightarrow> nat (x mod y) = nat x mod nat y" | 
| 
3c8f48333731
move nat_{div,mod}_distrib from NatBin to IntDiv, simplified proofs
 huffman parents: 
28562diff
changeset | 1232 | apply (case_tac "y = 0", simp add: DIVISION_BY_ZERO) | 
| 
3c8f48333731
move nat_{div,mod}_distrib from NatBin to IntDiv, simplified proofs
 huffman parents: 
28562diff
changeset | 1233 | apply (simp add: nat_eq_iff zmod_int) | 
| 
3c8f48333731
move nat_{div,mod}_distrib from NatBin to IntDiv, simplified proofs
 huffman parents: 
28562diff
changeset | 1234 | done | 
| 
3c8f48333731
move nat_{div,mod}_distrib from NatBin to IntDiv, simplified proofs
 huffman parents: 
28562diff
changeset | 1235 | |
| 
3c8f48333731
move nat_{div,mod}_distrib from NatBin to IntDiv, simplified proofs
 huffman parents: 
28562diff
changeset | 1236 | text{*Suggested by Matthias Daum*}
 | 
| 
3c8f48333731
move nat_{div,mod}_distrib from NatBin to IntDiv, simplified proofs
 huffman parents: 
28562diff
changeset | 1237 | lemma int_div_less_self: "\<lbrakk>0 < x; 1 < k\<rbrakk> \<Longrightarrow> x div k < (x::int)" | 
| 
3c8f48333731
move nat_{div,mod}_distrib from NatBin to IntDiv, simplified proofs
 huffman parents: 
28562diff
changeset | 1238 | apply (subgoal_tac "nat x div nat k < nat x") | 
| 
3c8f48333731
move nat_{div,mod}_distrib from NatBin to IntDiv, simplified proofs
 huffman parents: 
28562diff
changeset | 1239 | apply (simp (asm_lr) add: nat_div_distrib [symmetric]) | 
| 
3c8f48333731
move nat_{div,mod}_distrib from NatBin to IntDiv, simplified proofs
 huffman parents: 
28562diff
changeset | 1240 | apply (rule Divides.div_less_dividend, simp_all) | 
| 
3c8f48333731
move nat_{div,mod}_distrib from NatBin to IntDiv, simplified proofs
 huffman parents: 
28562diff
changeset | 1241 | done | 
| 
3c8f48333731
move nat_{div,mod}_distrib from NatBin to IntDiv, simplified proofs
 huffman parents: 
28562diff
changeset | 1242 | |
| 23853 | 1243 | text {* code generator setup *}
 | 
| 23164 | 1244 | |
| 26507 | 1245 | context ring_1 | 
| 1246 | begin | |
| 1247 | ||
| 28562 | 1248 | lemma of_int_num [code]: | 
| 26507 | 1249 | "of_int k = (if k = 0 then 0 else if k < 0 then | 
| 1250 | - of_int (- k) else let | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 1251 | (l, m) = divmod k 2; | 
| 26507 | 1252 | l' = of_int l | 
| 1253 | in if m = 0 then l' + l' else l' + l' + 1)" | |
| 1254 | proof - | |
| 1255 | have aux1: "k mod (2\<Colon>int) \<noteq> (0\<Colon>int) \<Longrightarrow> | |
| 1256 | of_int k = of_int (k div 2 * 2 + 1)" | |
| 1257 | proof - | |
| 1258 | have "k mod 2 < 2" by (auto intro: pos_mod_bound) | |
| 1259 | moreover have "0 \<le> k mod 2" by (auto intro: pos_mod_sign) | |
| 1260 | moreover assume "k mod 2 \<noteq> 0" | |
| 1261 | ultimately have "k mod 2 = 1" by arith | |
| 1262 | moreover have "of_int k = of_int (k div 2 * 2 + k mod 2)" by simp | |
| 1263 | ultimately show ?thesis by auto | |
| 1264 | qed | |
| 1265 | have aux2: "\<And>x. of_int 2 * x = x + x" | |
| 1266 | proof - | |
| 1267 | fix x | |
| 1268 | have int2: "(2::int) = 1 + 1" by arith | |
| 1269 | show "of_int 2 * x = x + x" | |
| 1270 | unfolding int2 of_int_add left_distrib by simp | |
| 1271 | qed | |
| 1272 | have aux3: "\<And>x. x * of_int 2 = x + x" | |
| 1273 | proof - | |
| 1274 | fix x | |
| 1275 | have int2: "(2::int) = 1 + 1" by arith | |
| 1276 | show "x * of_int 2 = x + x" | |
| 1277 | unfolding int2 of_int_add right_distrib by simp | |
| 1278 | qed | |
| 29651 
16a19466bf81
slightly adapted towards more uniformity with div/mod on nat
 haftmann parents: 
29410diff
changeset | 1279 | from aux1 show ?thesis by (auto simp add: divmod_mod_div Let_def aux2 aux3) | 
| 26507 | 1280 | qed | 
| 1281 | ||
| 1282 | end | |
| 1283 | ||
| 27667 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1284 | lemma zmod_eq_dvd_iff: "(x::int) mod n = y mod n \<longleftrightarrow> n dvd x - y" | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1285 | proof | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1286 | assume H: "x mod n = y mod n" | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1287 | hence "x mod n - y mod n = 0" by simp | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1288 | hence "(x mod n - y mod n) mod n = 0" by simp | 
| 30034 | 1289 | hence "(x - y) mod n = 0" by (simp add: mod_diff_eq[symmetric]) | 
| 30042 | 1290 | thus "n dvd x - y" by (simp add: dvd_eq_mod_eq_0) | 
| 27667 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1291 | next | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1292 | assume H: "n dvd x - y" | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1293 | then obtain k where k: "x-y = n*k" unfolding dvd_def by blast | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1294 | hence "x = n*k + y" by simp | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1295 | hence "x mod n = (n*k + y) mod n" by simp | 
| 30034 | 1296 | thus "x mod n = y mod n" by (simp add: mod_add_left_eq) | 
| 27667 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1297 | qed | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1298 | |
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1299 | lemma nat_mod_eq_lemma: assumes xyn: "(x::nat) mod n = y mod n" and xy:"y \<le> x" | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1300 | shows "\<exists>q. x = y + n * q" | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1301 | proof- | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1302 | from xy have th: "int x - int y = int (x - y)" by simp | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1303 | from xyn have "int x mod int n = int y mod int n" | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1304 | by (simp add: zmod_int[symmetric]) | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1305 | hence "int n dvd int x - int y" by (simp only: zmod_eq_dvd_iff[symmetric]) | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1306 | hence "n dvd x - y" by (simp add: th zdvd_int) | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1307 | then show ?thesis using xy unfolding dvd_def apply clarsimp apply (rule_tac x="k" in exI) by arith | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1308 | qed | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1309 | |
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1310 | lemma nat_mod_eq_iff: "(x::nat) mod n = y mod n \<longleftrightarrow> (\<exists>q1 q2. x + n * q1 = y + n * q2)" | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1311 | (is "?lhs = ?rhs") | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1312 | proof | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1313 | assume H: "x mod n = y mod n" | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1314 |   {assume xy: "x \<le> y"
 | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1315 | from H have th: "y mod n = x mod n" by simp | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1316 | from nat_mod_eq_lemma[OF th xy] have ?rhs | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1317 | apply clarify apply (rule_tac x="q" in exI) by (rule exI[where x="0"], simp)} | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1318 | moreover | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1319 |   {assume xy: "y \<le> x"
 | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1320 | from nat_mod_eq_lemma[OF H xy] have ?rhs | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1321 | apply clarify apply (rule_tac x="0" in exI) by (rule_tac x="q" in exI, simp)} | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1322 | ultimately show ?rhs using linear[of x y] by blast | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1323 | next | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1324 | assume ?rhs then obtain q1 q2 where q12: "x + n * q1 = y + n * q2" by blast | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1325 | hence "(x + n * q1) mod n = (y + n * q2) mod n" by simp | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1326 | thus ?lhs by simp | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1327 | qed | 
| 
62500b980749
Added theorems zmod_eq_dvd_iff and nat_mod_eq_iff previously in Pocklington.thy --- relevant for algebra
 chaieb parents: 
27651diff
changeset | 1328 | |
| 29936 | 1329 | |
| 1330 | subsection {* Code generation *}
 | |
| 1331 | ||
| 1332 | definition pdivmod :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where | |
| 1333 | "pdivmod k l = (\<bar>k\<bar> div \<bar>l\<bar>, \<bar>k\<bar> mod \<bar>l\<bar>)" | |
| 1334 | ||
| 1335 | lemma pdivmod_posDivAlg [code]: | |
| 1336 | "pdivmod k l = (if l = 0 then (0, \<bar>k\<bar>) else posDivAlg \<bar>k\<bar> \<bar>l\<bar>)" | |
| 1337 | by (subst posDivAlg_div_mod) (simp_all add: pdivmod_def) | |
| 1338 | ||
| 1339 | lemma divmod_pdivmod: "divmod k l = (if k = 0 then (0, 0) else if l = 0 then (0, k) else | |
| 1340 | apsnd ((op *) (sgn l)) (if 0 < l \<and> 0 \<le> k \<or> l < 0 \<and> k < 0 | |
| 1341 | then pdivmod k l | |
| 1342 | else (let (r, s) = pdivmod k l in | |
| 1343 | if s = 0 then (- r, 0) else (- r - 1, \<bar>l\<bar> - s))))" | |
| 1344 | proof - | |
| 1345 | have aux: "\<And>q::int. - k = l * q \<longleftrightarrow> k = l * - q" by auto | |
| 1346 | show ?thesis | |
| 1347 | by (simp add: divmod_mod_div pdivmod_def) | |
| 1348 | (auto simp add: aux not_less not_le zdiv_zminus1_eq_if | |
| 1349 | zmod_zminus1_eq_if zdiv_zminus2_eq_if zmod_zminus2_eq_if) | |
| 1350 | qed | |
| 1351 | ||
| 1352 | lemma divmod_code [code]: "divmod k l = (if k = 0 then (0, 0) else if l = 0 then (0, k) else | |
| 1353 | apsnd ((op *) (sgn l)) (if sgn k = sgn l | |
| 1354 | then pdivmod k l | |
| 1355 | else (let (r, s) = pdivmod k l in | |
| 1356 | if s = 0 then (- r, 0) else (- r - 1, \<bar>l\<bar> - s))))" | |
| 1357 | proof - | |
| 1358 | have "k \<noteq> 0 \<Longrightarrow> l \<noteq> 0 \<Longrightarrow> 0 < l \<and> 0 \<le> k \<or> l < 0 \<and> k < 0 \<longleftrightarrow> sgn k = sgn l" | |
| 1359 | by (auto simp add: not_less sgn_if) | |
| 1360 | then show ?thesis by (simp add: divmod_pdivmod) | |
| 1361 | qed | |
| 1362 | ||
| 23164 | 1363 | code_modulename SML | 
| 1364 | IntDiv Integer | |
| 1365 | ||
| 1366 | code_modulename OCaml | |
| 1367 | IntDiv Integer | |
| 1368 | ||
| 1369 | code_modulename Haskell | |
| 24195 | 1370 | IntDiv Integer | 
| 23164 | 1371 | |
| 1372 | end |