| author | wenzelm | 
| Wed, 06 Sep 2000 16:54:12 +0200 | |
| changeset 9876 | a069795f1060 | 
| parent 9747 | 043098ba5098 | 
| child 9945 | a0efbd7c88dc | 
| permissions | -rw-r--r-- | 
| 6917 | 1 | (* Title: HOL/IntDiv.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1999 University of Cambridge | |
| 5 | ||
| 6 | The division operators div, mod and the divides relation "dvd" | |
| 6992 | 7 | |
| 8 | Here is the division algorithm in ML: | |
| 9 | ||
| 10 | fun posDivAlg (a,b) = | |
| 11 | if a<b then (0,a) | |
| 12 | else let val (q,r) = posDivAlg(a, 2*b) | |
| 13 | in if 0<=r-b then (2*q+1, r-b) else (2*q, r) | |
| 14 | end; | |
| 15 | ||
| 16 | fun negDivAlg (a,b) = | |
| 17 | if 0<=a+b then (~1,a+b) | |
| 18 | else let val (q,r) = negDivAlg(a, 2*b) | |
| 19 | in if 0<=r-b then (2*q+1, r-b) else (2*q, r) | |
| 20 | end; | |
| 21 | ||
| 22 | fun negateSnd (q,r:int) = (q,~r); | |
| 23 | ||
| 24 | fun divAlg (a,b) = if 0<=a then | |
| 25 | if b>0 then posDivAlg (a,b) | |
| 26 | else if a=0 then (0,0) | |
| 27 | else negateSnd (negDivAlg (~a,~b)) | |
| 28 | else | |
| 29 | if 0<b then negDivAlg (a,b) | |
| 30 | else negateSnd (posDivAlg (~a,~b)); | |
| 6917 | 31 | *) | 
| 32 | ||
| 33 | Addsimps [zless_nat_conj]; | |
| 34 | ||
| 35 | (*** Uniqueness and monotonicity of quotients and remainders ***) | |
| 36 | ||
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 37 | Goal "[| b*q' + r' <= b*q + r; #0 <= r'; #0 < b; r < b |] \ | 
| 6917 | 38 | \ ==> q' <= (q::int)"; | 
| 39 | by (subgoal_tac "r' + b * (q'-q) <= r" 1); | |
| 8918 | 40 | by (simp_tac (simpset() addsimps [zdiff_zmult_distrib2]) 2); | 
| 6917 | 41 | by (subgoal_tac "#0 < b * (#1 + q - q')" 1); | 
| 42 | by (etac order_le_less_trans 2); | |
| 8918 | 43 | by (full_simp_tac (simpset() addsimps [zdiff_zmult_distrib2, | 
| 44 | zadd_zmult_distrib2]) 2); | |
| 6917 | 45 | by (subgoal_tac "b * q' < b * (#1 + q)" 1); | 
| 8918 | 46 | by (full_simp_tac (simpset() addsimps [zdiff_zmult_distrib2, | 
| 47 | zadd_zmult_distrib2]) 2); | |
| 9633 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9509diff
changeset | 48 | by (asm_full_simp_tac (simpset() addsimps [zmult_zless_cancel1]) 1); | 
| 6917 | 49 | qed "unique_quotient_lemma"; | 
| 50 | ||
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 51 | Goal "[| b*q' + r' <= b*q + r; r <= #0; b < #0; b < r' |] \ | 
| 6917 | 52 | \ ==> q <= (q'::int)"; | 
| 53 | by (res_inst_tac [("b", "-b"), ("r", "-r'"), ("r'", "-r")] 
 | |
| 54 | unique_quotient_lemma 1); | |
| 55 | by (auto_tac (claset(), | |
| 8918 | 56 | simpset() addsimps [zmult_zminus, zmult_zminus_right])); | 
| 6917 | 57 | qed "unique_quotient_lemma_neg"; | 
| 58 | ||
| 59 | ||
| 60 | Goal "[| quorem ((a,b), (q,r)); quorem ((a,b), (q',r')); b ~= #0 |] \ | |
| 61 | \ ==> q = q'"; | |
| 62 | by (asm_full_simp_tac | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 63 | (simpset() addsimps split_ifs@ | 
| 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 64 | [quorem_def, linorder_neq_iff]) 1); | 
| 6917 | 65 | by Safe_tac; | 
| 66 | by (ALLGOALS Asm_full_simp_tac); | |
| 67 | by (REPEAT | |
| 68 | (blast_tac (claset() addIs [order_antisym] | |
| 69 | addDs [order_eq_refl RS unique_quotient_lemma, | |
| 70 | order_eq_refl RS unique_quotient_lemma_neg, | |
| 71 | sym]) 1)); | |
| 72 | qed "unique_quotient"; | |
| 73 | ||
| 74 | ||
| 75 | Goal "[| quorem ((a,b), (q,r)); quorem ((a,b), (q',r')); b ~= #0 |] \ | |
| 76 | \ ==> r = r'"; | |
| 77 | by (subgoal_tac "q = q'" 1); | |
| 78 | by (blast_tac (claset() addIs [unique_quotient]) 2); | |
| 79 | by (asm_full_simp_tac (simpset() addsimps [quorem_def]) 1); | |
| 80 | qed "unique_remainder"; | |
| 81 | ||
| 82 | ||
| 83 | (*** Correctness of posDivAlg, the division algorithm for a>=0 and b>0 ***) | |
| 84 | ||
| 6943 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 85 | |
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 86 | Goal "adjust a b (q,r) = (let diff = r-b in \ | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 87 | \ if #0 <= diff then (#2*q + #1, diff) \ | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 88 | \ else (#2*q, r))"; | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 89 | by (simp_tac (simpset() addsimps [Let_def,adjust_def]) 1); | 
| 6917 | 90 | qed "adjust_eq"; | 
| 91 | Addsimps [adjust_eq]; | |
| 92 | ||
| 93 | (*Proving posDivAlg's termination condition*) | |
| 94 | val [tc] = posDivAlg.tcs; | |
| 9367 | 95 | goalw_cterm [] (cterm_of (sign_of (the_context ())) (HOLogic.mk_Trueprop tc)); | 
| 8787 | 96 | by Auto_tac; | 
| 6917 | 97 | val lemma = result(); | 
| 98 | ||
| 99 | (* removing the termination condition from the generated theorems *) | |
| 100 | ||
| 8624 | 101 | bind_thm ("posDivAlg_raw_eqn", lemma RS hd posDivAlg.simps);
 | 
| 6917 | 102 | |
| 103 | (**use with simproc to avoid re-proving the premise*) | |
| 104 | Goal "#0 < b ==> \ | |
| 105 | \ posDivAlg (a,b) = \ | |
| 106 | \ (if a<b then (#0,a) else adjust a b (posDivAlg(a, #2*b)))"; | |
| 107 | by (rtac (posDivAlg_raw_eqn RS trans) 1); | |
| 108 | by (Asm_simp_tac 1); | |
| 109 | qed "posDivAlg_eqn"; | |
| 110 | ||
| 9108 | 111 | bind_thm ("posDivAlg_induct", lemma RS posDivAlg.induct);
 | 
| 6917 | 112 | |
| 113 | ||
| 114 | (*Correctness of posDivAlg: it computes quotients correctly*) | |
| 115 | Goal "#0 <= a --> #0 < b --> quorem ((a, b), posDivAlg (a, b))"; | |
| 9747 | 116 | by (induct_thm_tac posDivAlg_induct "a b" 1); | 
| 6917 | 117 | by Auto_tac; | 
| 9063 | 118 | by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [quorem_def]))); | 
| 6917 | 119 | (*base case: a<b*) | 
| 120 | by (asm_full_simp_tac (simpset() addsimps [posDivAlg_eqn]) 1); | |
| 121 | (*main argument*) | |
| 122 | by (stac posDivAlg_eqn 1); | |
| 123 | by (ALLGOALS Asm_simp_tac); | |
| 124 | by (etac splitE 1); | |
| 8918 | 125 | by (auto_tac (claset(), simpset() addsimps [zadd_zmult_distrib2, Let_def])); | 
| 6917 | 126 | qed_spec_mp "posDivAlg_correct"; | 
| 127 | ||
| 128 | ||
| 129 | (*** Correctness of negDivAlg, the division algorithm for a<0 and b>0 ***) | |
| 130 | ||
| 131 | (*Proving negDivAlg's termination condition*) | |
| 132 | val [tc] = negDivAlg.tcs; | |
| 9422 | 133 | goalw_cterm [] (cterm_of (sign_of (the_context ())) (HOLogic.mk_Trueprop tc)); | 
| 8787 | 134 | by Auto_tac; | 
| 6917 | 135 | val lemma = result(); | 
| 136 | ||
| 137 | (* removing the termination condition from the generated theorems *) | |
| 138 | ||
| 8624 | 139 | bind_thm ("negDivAlg_raw_eqn", lemma RS hd negDivAlg.simps);
 | 
| 6917 | 140 | |
| 141 | (**use with simproc to avoid re-proving the premise*) | |
| 142 | Goal "#0 < b ==> \ | |
| 143 | \ negDivAlg (a,b) = \ | |
| 144 | \ (if #0<=a+b then (#-1,a+b) else adjust a b (negDivAlg(a, #2*b)))"; | |
| 145 | by (rtac (negDivAlg_raw_eqn RS trans) 1); | |
| 146 | by (Asm_simp_tac 1); | |
| 147 | qed "negDivAlg_eqn"; | |
| 148 | ||
| 9108 | 149 | bind_thm ("negDivAlg_induct", lemma RS negDivAlg.induct);
 | 
| 6917 | 150 | |
| 151 | ||
| 152 | (*Correctness of negDivAlg: it computes quotients correctly | |
| 153 | It doesn't work if a=0 because the 0/b=0 rather than -1*) | |
| 154 | Goal "a < #0 --> #0 < b --> quorem ((a, b), negDivAlg (a, b))"; | |
| 9747 | 155 | by (induct_thm_tac negDivAlg_induct "a b" 1); | 
| 6917 | 156 | by Auto_tac; | 
| 9063 | 157 | by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [quorem_def]))); | 
| 6917 | 158 | (*base case: 0<=a+b*) | 
| 159 | by (asm_full_simp_tac (simpset() addsimps [negDivAlg_eqn]) 1); | |
| 160 | (*main argument*) | |
| 161 | by (stac negDivAlg_eqn 1); | |
| 162 | by (ALLGOALS Asm_simp_tac); | |
| 163 | by (etac splitE 1); | |
| 8918 | 164 | by (auto_tac (claset(), simpset() addsimps [zadd_zmult_distrib2, Let_def])); | 
| 6917 | 165 | qed_spec_mp "negDivAlg_correct"; | 
| 166 | ||
| 167 | ||
| 168 | (*** Existence shown by proving the division algorithm to be correct ***) | |
| 169 | ||
| 170 | (*the case a=0*) | |
| 171 | Goal "b ~= #0 ==> quorem ((#0,b), (#0,#0))"; | |
| 172 | by (auto_tac (claset(), | |
| 173 | simpset() addsimps [quorem_def, linorder_neq_iff])); | |
| 174 | qed "quorem_0"; | |
| 175 | ||
| 176 | Goal "posDivAlg (#0, b) = (#0, #0)"; | |
| 177 | by (stac posDivAlg_raw_eqn 1); | |
| 178 | by Auto_tac; | |
| 179 | qed "posDivAlg_0"; | |
| 180 | Addsimps [posDivAlg_0]; | |
| 181 | ||
| 182 | Goal "negDivAlg (#-1, b) = (#-1, b-#1)"; | |
| 183 | by (stac negDivAlg_raw_eqn 1); | |
| 184 | by Auto_tac; | |
| 185 | qed "negDivAlg_minus1"; | |
| 186 | Addsimps [negDivAlg_minus1]; | |
| 187 | ||
| 188 | Goalw [negateSnd_def] "negateSnd(q,r) = (q,-r)"; | |
| 189 | by Auto_tac; | |
| 190 | qed "negateSnd_eq"; | |
| 191 | Addsimps [negateSnd_eq]; | |
| 192 | ||
| 193 | Goal "quorem ((-a,-b), qr) ==> quorem ((a,b), negateSnd qr)"; | |
| 8785 | 194 | by (auto_tac (claset(), simpset() addsimps split_ifs@[quorem_def])); | 
| 6917 | 195 | qed "quorem_neg"; | 
| 196 | ||
| 197 | Goal "b ~= #0 ==> quorem ((a,b), divAlg(a,b))"; | |
| 198 | by (auto_tac (claset(), | |
| 199 | simpset() addsimps [quorem_0, divAlg_def])); | |
| 200 | by (REPEAT_FIRST (resolve_tac [quorem_neg, posDivAlg_correct, | |
| 201 | negDivAlg_correct])); | |
| 202 | by (auto_tac (claset(), | |
| 203 | simpset() addsimps [quorem_def, linorder_neq_iff])); | |
| 204 | qed "divAlg_correct"; | |
| 205 | ||
| 6992 | 206 | (** Aribtrary definitions for division by zero. Useful to simplify | 
| 207 | certain equations **) | |
| 208 | ||
| 209 | Goal "a div (#0::int) = #0"; | |
| 210 | by (simp_tac (simpset() addsimps [div_def, divAlg_def, posDivAlg_raw_eqn]) 1); | |
| 7035 | 211 | qed "DIVISION_BY_ZERO_ZDIV"; (*NOT for adding to default simpset*) | 
| 6992 | 212 | |
| 213 | Goal "a mod (#0::int) = a"; | |
| 214 | by (simp_tac (simpset() addsimps [mod_def, divAlg_def, posDivAlg_raw_eqn]) 1); | |
| 7035 | 215 | qed "DIVISION_BY_ZERO_ZMOD"; (*NOT for adding to default simpset*) | 
| 6992 | 216 | |
| 7035 | 217 | fun zdiv_undefined_case_tac s i = | 
| 6992 | 218 | case_tac s i THEN | 
| 7035 | 219 | asm_simp_tac (simpset() addsimps [DIVISION_BY_ZERO_ZDIV, | 
| 220 | DIVISION_BY_ZERO_ZMOD]) i; | |
| 6992 | 221 | |
| 222 | (** Basic laws about division and remainder **) | |
| 223 | ||
| 224 | Goal "(a::int) = b * (a div b) + (a mod b)"; | |
| 7035 | 225 | by (zdiv_undefined_case_tac "b = #0" 1); | 
| 6917 | 226 | by (cut_inst_tac [("a","a"),("b","b")] divAlg_correct 1);
 | 
| 227 | by (auto_tac (claset(), | |
| 228 | simpset() addsimps [quorem_def, div_def, mod_def])); | |
| 229 | qed "zmod_zdiv_equality"; | |
| 230 | ||
| 231 | Goal "(#0::int) < b ==> #0 <= a mod b & a mod b < b"; | |
| 232 | by (cut_inst_tac [("a","a"),("b","b")] divAlg_correct 1);
 | |
| 233 | by (auto_tac (claset(), | |
| 234 | simpset() addsimps [quorem_def, mod_def])); | |
| 235 | bind_thm ("pos_mod_sign", result() RS conjunct1);
 | |
| 236 | bind_thm ("pos_mod_bound", result() RS conjunct2);
 | |
| 237 | ||
| 238 | Goal "b < (#0::int) ==> a mod b <= #0 & b < a mod b"; | |
| 239 | by (cut_inst_tac [("a","a"),("b","b")] divAlg_correct 1);
 | |
| 240 | by (auto_tac (claset(), | |
| 241 | simpset() addsimps [quorem_def, div_def, mod_def])); | |
| 242 | bind_thm ("neg_mod_sign", result() RS conjunct1);
 | |
| 243 | bind_thm ("neg_mod_bound", result() RS conjunct2);
 | |
| 244 | ||
| 245 | ||
| 6992 | 246 | (** proving general properties of div and mod **) | 
| 247 | ||
| 248 | Goal "b ~= #0 ==> quorem ((a, b), (a div b, a mod b))"; | |
| 249 | by (cut_inst_tac [("a","a"),("b","b")] zmod_zdiv_equality 1);
 | |
| 250 | by (auto_tac | |
| 251 | (claset(), | |
| 252 | simpset() addsimps [quorem_def, linorder_neq_iff, | |
| 253 | pos_mod_sign,pos_mod_bound, | |
| 254 | neg_mod_sign, neg_mod_bound])); | |
| 255 | qed "quorem_div_mod"; | |
| 256 | ||
| 257 | Goal "[| quorem((a,b),(q,r)); b ~= #0 |] ==> a div b = q"; | |
| 258 | by (asm_simp_tac (simpset() addsimps [quorem_div_mod RS unique_quotient]) 1); | |
| 259 | qed "quorem_div"; | |
| 260 | ||
| 261 | Goal "[| quorem((a,b),(q,r)); b ~= #0 |] ==> a mod b = r"; | |
| 262 | by (asm_simp_tac (simpset() addsimps [quorem_div_mod RS unique_remainder]) 1); | |
| 263 | qed "quorem_mod"; | |
| 264 | ||
| 265 | Goal "[| (#0::int) <= a; a < b |] ==> a div b = #0"; | |
| 266 | by (rtac quorem_div 1); | |
| 267 | by (auto_tac (claset(), simpset() addsimps [quorem_def])); | |
| 6999 | 268 | qed "div_pos_pos_trivial"; | 
| 6992 | 269 | |
| 270 | Goal "[| a <= (#0::int); b < a |] ==> a div b = #0"; | |
| 271 | by (rtac quorem_div 1); | |
| 272 | by (auto_tac (claset(), simpset() addsimps [quorem_def])); | |
| 6999 | 273 | qed "div_neg_neg_trivial"; | 
| 274 | ||
| 275 | Goal "[| (#0::int) < a; a+b <= #0 |] ==> a div b = #-1"; | |
| 276 | by (rtac quorem_div 1); | |
| 277 | by (auto_tac (claset(), simpset() addsimps [quorem_def])); | |
| 278 | qed "div_pos_neg_trivial"; | |
| 279 | ||
| 280 | (*There is no div_neg_pos_trivial because #0 div b = #0 would supersede it*) | |
| 6992 | 281 | |
| 282 | Goal "[| (#0::int) <= a; a < b |] ==> a mod b = a"; | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 283 | by (res_inst_tac [("q","#0")] quorem_mod 1);
 | 
| 6992 | 284 | by (auto_tac (claset(), simpset() addsimps [quorem_def])); | 
| 6999 | 285 | qed "mod_pos_pos_trivial"; | 
| 6992 | 286 | |
| 287 | Goal "[| a <= (#0::int); b < a |] ==> a mod b = a"; | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 288 | by (res_inst_tac [("q","#0")] quorem_mod 1);
 | 
| 6992 | 289 | by (auto_tac (claset(), simpset() addsimps [quorem_def])); | 
| 6999 | 290 | qed "mod_neg_neg_trivial"; | 
| 291 | ||
| 292 | Goal "[| (#0::int) < a; a+b <= #0 |] ==> a mod b = a+b"; | |
| 293 | by (res_inst_tac [("q","#-1")] quorem_mod 1);
 | |
| 294 | by (auto_tac (claset(), simpset() addsimps [quorem_def])); | |
| 295 | qed "mod_pos_neg_trivial"; | |
| 296 | ||
| 297 | (*There is no mod_neg_pos_trivial...*) | |
| 6992 | 298 | |
| 299 | ||
| 300 | (*Simpler laws such as -a div b = -(a div b) FAIL*) | |
| 301 | Goal "(-a) div (-b) = a div (b::int)"; | |
| 7035 | 302 | by (zdiv_undefined_case_tac "b = #0" 1); | 
| 6992 | 303 | by (stac ((simplify(simpset()) (quorem_div_mod RS quorem_neg)) | 
| 304 | RS quorem_div) 1); | |
| 305 | by Auto_tac; | |
| 306 | qed "zdiv_zminus_zminus"; | |
| 307 | Addsimps [zdiv_zminus_zminus]; | |
| 308 | ||
| 309 | (*Simpler laws such as -a mod b = -(a mod b) FAIL*) | |
| 310 | Goal "(-a) mod (-b) = - (a mod (b::int))"; | |
| 7035 | 311 | by (zdiv_undefined_case_tac "b = #0" 1); | 
| 6992 | 312 | by (stac ((simplify(simpset()) (quorem_div_mod RS quorem_neg)) | 
| 313 | RS quorem_mod) 1); | |
| 314 | by Auto_tac; | |
| 315 | qed "zmod_zminus_zminus"; | |
| 316 | Addsimps [zmod_zminus_zminus]; | |
| 317 | ||
| 318 | ||
| 319 | (*** division of a number by itself ***) | |
| 320 | ||
| 321 | Goal "[| (#0::int) < a; a = r + a*q; r < a |] ==> #1 <= q"; | |
| 322 | by (subgoal_tac "#0 < a*q" 1); | |
| 323 | by (arith_tac 2); | |
| 9063 | 324 | by (asm_full_simp_tac (simpset() addsimps [int_0_less_mult_iff]) 1); | 
| 6992 | 325 | val lemma1 = result(); | 
| 326 | ||
| 327 | Goal "[| (#0::int) < a; a = r + a*q; #0 <= r |] ==> q <= #1"; | |
| 328 | by (subgoal_tac "#0 <= a*(#1-q)" 1); | |
| 329 | by (asm_simp_tac (simpset() addsimps [zdiff_zmult_distrib2]) 2); | |
| 9063 | 330 | by (asm_full_simp_tac (simpset() addsimps [int_0_le_mult_iff]) 1); | 
| 6992 | 331 | val lemma2 = result(); | 
| 332 | ||
| 333 | Goal "[| quorem((a,a),(q,r)); a ~= (#0::int) |] ==> q = #1"; | |
| 334 | by (asm_full_simp_tac | |
| 335 | (simpset() addsimps split_ifs@[quorem_def, linorder_neq_iff]) 1); | |
| 336 | by (rtac order_antisym 1); | |
| 337 | by Safe_tac; | |
| 338 | by Auto_tac; | |
| 339 | by (res_inst_tac [("a", "-a"),("r", "-r")] lemma1 3);
 | |
| 340 | by (res_inst_tac [("a", "-a"),("r", "-r")] lemma2 1);
 | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 341 | by (REPEAT (force_tac (claset() addIs [lemma1,lemma2], | 
| 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 342 | simpset() addsimps [zadd_commute, zmult_zminus]) 1)); | 
| 6992 | 343 | qed "self_quotient"; | 
| 344 | ||
| 345 | Goal "[| quorem((a,a),(q,r)); a ~= (#0::int) |] ==> r = #0"; | |
| 7499 | 346 | by (ftac self_quotient 1); | 
| 6992 | 347 | by (assume_tac 1); | 
| 348 | by (asm_full_simp_tac (simpset() addsimps [quorem_def]) 1); | |
| 349 | qed "self_remainder"; | |
| 350 | ||
| 351 | Goal "a ~= #0 ==> a div a = (#1::int)"; | |
| 352 | by (asm_simp_tac (simpset() addsimps [quorem_div_mod RS self_quotient]) 1); | |
| 353 | qed "zdiv_self"; | |
| 354 | Addsimps [zdiv_self]; | |
| 355 | ||
| 356 | (*Here we have 0 mod 0 = 0, also assumed by Knuth (who puts m mod 0 = 0) *) | |
| 357 | Goal "a mod a = (#0::int)"; | |
| 7035 | 358 | by (zdiv_undefined_case_tac "a = #0" 1); | 
| 6992 | 359 | by (asm_simp_tac (simpset() addsimps [quorem_div_mod RS self_remainder]) 1); | 
| 360 | qed "zmod_self"; | |
| 361 | Addsimps [zmod_self]; | |
| 362 | ||
| 363 | ||
| 6917 | 364 | (*** Computation of division and remainder ***) | 
| 365 | ||
| 366 | Goal "(#0::int) div b = #0"; | |
| 367 | by (simp_tac (simpset() addsimps [div_def, divAlg_def]) 1); | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 368 | qed "zdiv_zero"; | 
| 6917 | 369 | |
| 7035 | 370 | Goal "(#0::int) < b ==> #-1 div b = #-1"; | 
| 371 | by (asm_simp_tac (simpset() addsimps [div_def, divAlg_def]) 1); | |
| 372 | qed "div_eq_minus1"; | |
| 373 | ||
| 6917 | 374 | Goal "(#0::int) mod b = #0"; | 
| 375 | by (simp_tac (simpset() addsimps [mod_def, divAlg_def]) 1); | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 376 | qed "zmod_zero"; | 
| 6917 | 377 | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 378 | Addsimps [zdiv_zero, zmod_zero]; | 
| 6917 | 379 | |
| 7035 | 380 | Goal "(#0::int) < b ==> #-1 div b = #-1"; | 
| 381 | by (asm_simp_tac (simpset() addsimps [div_def, divAlg_def]) 1); | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 382 | qed "zdiv_minus1"; | 
| 7035 | 383 | |
| 384 | Goal "(#0::int) < b ==> #-1 mod b = b-#1"; | |
| 385 | by (asm_simp_tac (simpset() addsimps [mod_def, divAlg_def]) 1); | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 386 | qed "zmod_minus1"; | 
| 7035 | 387 | |
| 6917 | 388 | (** a positive, b positive **) | 
| 389 | ||
| 6992 | 390 | Goal "[| #0 < a; #0 <= b |] ==> a div b = fst (posDivAlg(a,b))"; | 
| 6917 | 391 | by (asm_simp_tac (simpset() addsimps [div_def, divAlg_def]) 1); | 
| 392 | qed "div_pos_pos"; | |
| 393 | ||
| 6992 | 394 | Goal "[| #0 < a; #0 <= b |] ==> a mod b = snd (posDivAlg(a,b))"; | 
| 6917 | 395 | by (asm_simp_tac (simpset() addsimps [mod_def, divAlg_def]) 1); | 
| 396 | qed "mod_pos_pos"; | |
| 397 | ||
| 398 | (** a negative, b positive **) | |
| 399 | ||
| 400 | Goal "[| a < #0; #0 < b |] ==> a div b = fst (negDivAlg(a,b))"; | |
| 401 | by (asm_simp_tac (simpset() addsimps [div_def, divAlg_def]) 1); | |
| 402 | qed "div_neg_pos"; | |
| 403 | ||
| 404 | Goal "[| a < #0; #0 < b |] ==> a mod b = snd (negDivAlg(a,b))"; | |
| 405 | by (asm_simp_tac (simpset() addsimps [mod_def, divAlg_def]) 1); | |
| 406 | qed "mod_neg_pos"; | |
| 407 | ||
| 408 | (** a positive, b negative **) | |
| 409 | ||
| 410 | Goal "[| #0 < a; b < #0 |] ==> a div b = fst (negateSnd(negDivAlg(-a,-b)))"; | |
| 411 | by (asm_simp_tac (simpset() addsimps [div_def, divAlg_def]) 1); | |
| 412 | qed "div_pos_neg"; | |
| 413 | ||
| 414 | Goal "[| #0 < a; b < #0 |] ==> a mod b = snd (negateSnd(negDivAlg(-a,-b)))"; | |
| 415 | by (asm_simp_tac (simpset() addsimps [mod_def, divAlg_def]) 1); | |
| 416 | qed "mod_pos_neg"; | |
| 417 | ||
| 418 | (** a negative, b negative **) | |
| 419 | ||
| 6992 | 420 | Goal "[| a < #0; b <= #0 |] ==> a div b = fst (negateSnd(posDivAlg(-a,-b)))"; | 
| 6917 | 421 | by (asm_simp_tac (simpset() addsimps [div_def, divAlg_def]) 1); | 
| 422 | qed "div_neg_neg"; | |
| 423 | ||
| 6992 | 424 | Goal "[| a < #0; b <= #0 |] ==> a mod b = snd (negateSnd(posDivAlg(-a,-b)))"; | 
| 6917 | 425 | by (asm_simp_tac (simpset() addsimps [mod_def, divAlg_def]) 1); | 
| 426 | qed "mod_neg_neg"; | |
| 427 | ||
| 9422 | 428 | Addsimps (map (read_instantiate_sg (sign_of (the_context ())) | 
| 6917 | 429 | 	       [("a", "number_of ?v"), ("b", "number_of ?w")])
 | 
| 430 | [div_pos_pos, div_neg_pos, div_pos_neg, div_neg_neg, | |
| 431 | mod_pos_pos, mod_neg_pos, mod_pos_neg, mod_neg_neg, | |
| 432 | posDivAlg_eqn, negDivAlg_eqn]); | |
| 6992 | 433 | |
| 6917 | 434 | |
| 435 | (** Special-case simplification **) | |
| 436 | ||
| 437 | Goal "a mod (#1::int) = #0"; | |
| 438 | by (cut_inst_tac [("a","a"),("b","#1")] pos_mod_sign 1);
 | |
| 439 | by (cut_inst_tac [("a","a"),("b","#1")] pos_mod_bound 2);
 | |
| 440 | by Auto_tac; | |
| 441 | qed "zmod_1"; | |
| 442 | Addsimps [zmod_1]; | |
| 443 | ||
| 444 | Goal "a div (#1::int) = a"; | |
| 445 | by (cut_inst_tac [("a","a"),("b","#1")] zmod_zdiv_equality 1);
 | |
| 446 | by Auto_tac; | |
| 447 | qed "zdiv_1"; | |
| 448 | Addsimps [zdiv_1]; | |
| 449 | ||
| 450 | Goal "a mod (#-1::int) = #0"; | |
| 451 | by (cut_inst_tac [("a","a"),("b","#-1")] neg_mod_sign 1);
 | |
| 452 | by (cut_inst_tac [("a","a"),("b","#-1")] neg_mod_bound 2);
 | |
| 453 | by Auto_tac; | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 454 | qed "zmod_minus1_right"; | 
| 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 455 | Addsimps [zmod_minus1_right]; | 
| 6917 | 456 | |
| 457 | Goal "a div (#-1::int) = -a"; | |
| 458 | by (cut_inst_tac [("a","a"),("b","#-1")] zmod_zdiv_equality 1);
 | |
| 459 | by Auto_tac; | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 460 | qed "zdiv_minus1_right"; | 
| 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 461 | Addsimps [zdiv_minus1_right]; | 
| 6917 | 462 | |
| 463 | ||
| 6943 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 464 | (*** Monotonicity in the first argument (divisor) ***) | 
| 6917 | 465 | |
| 466 | Goal "[| a <= a'; #0 < (b::int) |] ==> a div b <= a' div b"; | |
| 467 | by (cut_inst_tac [("a","a"),("b","b")] zmod_zdiv_equality 1);
 | |
| 6992 | 468 | by (cut_inst_tac [("a","a'"),("b","b")] zmod_zdiv_equality 1);
 | 
| 6917 | 469 | by (rtac unique_quotient_lemma 1); | 
| 470 | by (etac subst 1); | |
| 471 | by (etac subst 1); | |
| 472 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [pos_mod_sign,pos_mod_bound]))); | |
| 473 | qed "zdiv_mono1"; | |
| 474 | ||
| 475 | Goal "[| a <= a'; (b::int) < #0 |] ==> a' div b <= a div b"; | |
| 476 | by (cut_inst_tac [("a","a"),("b","b")] zmod_zdiv_equality 1);
 | |
| 6992 | 477 | by (cut_inst_tac [("a","a'"),("b","b")] zmod_zdiv_equality 1);
 | 
| 6917 | 478 | by (rtac unique_quotient_lemma_neg 1); | 
| 479 | by (etac subst 1); | |
| 480 | by (etac subst 1); | |
| 481 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [neg_mod_sign,neg_mod_bound]))); | |
| 482 | qed "zdiv_mono1_neg"; | |
| 483 | ||
| 484 | ||
| 6943 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 485 | (*** Monotonicity in the second argument (dividend) ***) | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 486 | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 487 | Goal "[| b*q + r = b'*q' + r'; #0 <= b'*q' + r'; \ | 
| 6943 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 488 | \ r' < b'; #0 <= r; #0 < b'; b' <= b |] \ | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 489 | \ ==> q <= (q'::int)"; | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 490 | by (subgoal_tac "#0 <= q'" 1); | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 491 | by (subgoal_tac "#0 < b'*(q' + #1)" 2); | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 492 | by (asm_simp_tac (simpset() addsimps [zadd_zmult_distrib2]) 3); | 
| 9063 | 493 | by (asm_full_simp_tac (simpset() addsimps [int_0_less_mult_iff]) 2); | 
| 6943 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 494 | by (subgoal_tac "b*q < b*(q' + #1)" 1); | 
| 9633 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9509diff
changeset | 495 | by (asm_full_simp_tac (simpset() addsimps [zmult_zless_cancel1]) 1); | 
| 6943 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 496 | by (subgoal_tac "b*q = r' - r + b'*q'" 1); | 
| 8918 | 497 | by (Simp_tac 2); | 
| 6943 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 498 | by (asm_simp_tac (simpset() addsimps [zadd_zmult_distrib2]) 1); | 
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 499 | by (stac zadd_commute 1 THEN rtac zadd_zless_mono 1 THEN arith_tac 1); | 
| 6943 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 500 | by (rtac zmult_zle_mono1 1); | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 501 | by Auto_tac; | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 502 | qed "zdiv_mono2_lemma"; | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 503 | |
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 504 | Goal "[| (#0::int) <= a; #0 < b'; b' <= b |] \ | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 505 | \ ==> a div b <= a div b'"; | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 506 | by (subgoal_tac "b ~= #0" 1); | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 507 | by (arith_tac 2); | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 508 | by (cut_inst_tac [("a","a"),("b","b")] zmod_zdiv_equality 1);
 | 
| 6992 | 509 | by (cut_inst_tac [("a","a"),("b","b'")] zmod_zdiv_equality 1);
 | 
| 6943 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 510 | by (rtac zdiv_mono2_lemma 1); | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 511 | by (etac subst 1); | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 512 | by (etac subst 1); | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 513 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [pos_mod_sign,pos_mod_bound]))); | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 514 | qed "zdiv_mono2"; | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 515 | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 516 | Goal "[| b*q + r = b'*q' + r'; b'*q' + r' < #0; \ | 
| 6943 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 517 | \ r < b; #0 <= r'; #0 < b'; b' <= b |] \ | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 518 | \ ==> q' <= (q::int)"; | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 519 | by (subgoal_tac "q' < #0" 1); | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 520 | by (subgoal_tac "b'*q' < #0" 2); | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 521 | by (arith_tac 3); | 
| 9063 | 522 | by (asm_full_simp_tac (simpset() addsimps [zmult_less_0_iff]) 2); | 
| 6943 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 523 | by (subgoal_tac "b*q' < b*(q + #1)" 1); | 
| 9633 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9509diff
changeset | 524 | by (asm_full_simp_tac (simpset() addsimps [zmult_zless_cancel1]) 1); | 
| 6943 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 525 | by (asm_simp_tac (simpset() addsimps [zadd_zmult_distrib2]) 1); | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 526 | by (subgoal_tac "b*q' <= b'*q'" 1); | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 527 | by (asm_simp_tac (simpset() addsimps [zmult_zle_mono1_neg]) 2); | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 528 | by (subgoal_tac "b'*q' < b + b*q" 1); | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 529 | by (Asm_simp_tac 2); | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 530 | by (arith_tac 1); | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 531 | qed "zdiv_mono2_neg_lemma"; | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 532 | |
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 533 | Goal "[| a < (#0::int); #0 < b'; b' <= b |] \ | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 534 | \ ==> a div b' <= a div b"; | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 535 | by (cut_inst_tac [("a","a"),("b","b")] zmod_zdiv_equality 1);
 | 
| 6992 | 536 | by (cut_inst_tac [("a","a"),("b","b'")] zmod_zdiv_equality 1);
 | 
| 6943 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 537 | by (rtac zdiv_mono2_neg_lemma 1); | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 538 | by (etac subst 1); | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 539 | by (etac subst 1); | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 540 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [pos_mod_sign,pos_mod_bound]))); | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 541 | qed "zdiv_mono2_neg"; | 
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 542 | |
| 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 543 | |
| 6992 | 544 | (*** More algebraic laws for div and mod ***) | 
| 6943 
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
 paulson parents: 
6917diff
changeset | 545 | |
| 6992 | 546 | (** proving (a*b) div c = a * (b div c) + a * (b mod c) **) | 
| 547 | ||
| 548 | Goal "[| quorem((b,c),(q,r)); c ~= #0 |] \ | |
| 549 | \ ==> quorem ((a*b, c), (a*q + a*r div c, a*r mod c))"; | |
| 550 | by (auto_tac | |
| 551 | (claset(), | |
| 8918 | 552 | simpset() addsimps split_ifs@ | 
| 6992 | 553 | [quorem_def, linorder_neq_iff, | 
| 554 | zadd_zmult_distrib2, | |
| 555 | pos_mod_sign,pos_mod_bound, | |
| 556 | neg_mod_sign, neg_mod_bound])); | |
| 7549 | 557 | by (ALLGOALS(rtac zmod_zdiv_equality)); | 
| 6992 | 558 | val lemma = result(); | 
| 559 | ||
| 560 | Goal "(a*b) div c = a*(b div c) + a*(b mod c) div (c::int)"; | |
| 7035 | 561 | by (zdiv_undefined_case_tac "c = #0" 1); | 
| 6992 | 562 | by (blast_tac (claset() addIs [quorem_div_mod RS lemma RS quorem_div]) 1); | 
| 563 | qed "zdiv_zmult1_eq"; | |
| 564 | ||
| 565 | Goal "(a*b) mod c = a*(b mod c) mod (c::int)"; | |
| 7035 | 566 | by (zdiv_undefined_case_tac "c = #0" 1); | 
| 6992 | 567 | by (blast_tac (claset() addIs [quorem_div_mod RS lemma RS quorem_mod]) 1); | 
| 568 | qed "zmod_zmult1_eq"; | |
| 569 | ||
| 9509 | 570 | Goal "(a*b) mod (c::int) = ((a mod c) * b) mod c"; | 
| 571 | by (rtac trans 1); | |
| 572 | by (res_inst_tac [("s","b*a mod c")] trans 1);
 | |
| 573 | by (rtac zmod_zmult1_eq 2); | |
| 574 | by (ALLGOALS (simp_tac (simpset() addsimps [zmult_commute]))); | |
| 575 | qed "zmod_zmult1_eq'"; | |
| 576 | ||
| 577 | Goal "(a*b) mod (c::int) = ((a mod c) * (b mod c)) mod c"; | |
| 578 | by (rtac (zmod_zmult1_eq' RS trans) 1); | |
| 579 | by (rtac zmod_zmult1_eq 1); | |
| 580 | qed "zmod_zmult_distrib"; | |
| 581 | ||
| 6992 | 582 | Goal "b ~= (#0::int) ==> (a*b) div b = a"; | 
| 583 | by (asm_simp_tac (simpset() addsimps [zdiv_zmult1_eq]) 1); | |
| 584 | qed "zdiv_zmult_self1"; | |
| 585 | ||
| 586 | Goal "b ~= (#0::int) ==> (b*a) div b = a"; | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 587 | by (stac zmult_commute 1 THEN etac zdiv_zmult_self1 1); | 
| 6992 | 588 | qed "zdiv_zmult_self2"; | 
| 589 | ||
| 590 | Addsimps [zdiv_zmult_self1, zdiv_zmult_self2]; | |
| 591 | ||
| 592 | Goal "(a*b) mod b = (#0::int)"; | |
| 593 | by (simp_tac (simpset() addsimps [zmod_zmult1_eq]) 1); | |
| 594 | qed "zmod_zmult_self1"; | |
| 595 | ||
| 596 | Goal "(b*a) mod b = (#0::int)"; | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 597 | by (simp_tac (simpset() addsimps [zmult_commute, zmod_zmult1_eq]) 1); | 
| 6992 | 598 | qed "zmod_zmult_self2"; | 
| 599 | ||
| 600 | Addsimps [zmod_zmult_self1, zmod_zmult_self2]; | |
| 601 | ||
| 602 | ||
| 603 | (** proving (a+b) div c = a div c + b div c + ((a mod c + b mod c) div c) **) | |
| 6917 | 604 | |
| 6992 | 605 | Goal "[| quorem((a,c),(aq,ar)); quorem((b,c),(bq,br)); c ~= #0 |] \ | 
| 606 | \ ==> quorem ((a+b, c), (aq + bq + (ar+br) div c, (ar+br) mod c))"; | |
| 607 | by (auto_tac | |
| 608 | (claset(), | |
| 8918 | 609 | simpset() addsimps split_ifs@ | 
| 6992 | 610 | [quorem_def, linorder_neq_iff, | 
| 611 | zadd_zmult_distrib2, | |
| 612 | pos_mod_sign,pos_mod_bound, | |
| 613 | neg_mod_sign, neg_mod_bound])); | |
| 7549 | 614 | by (ALLGOALS(rtac zmod_zdiv_equality)); | 
| 6992 | 615 | val lemma = result(); | 
| 616 | ||
| 6999 | 617 | (*NOT suitable for rewriting: the RHS has an instance of the LHS*) | 
| 6992 | 618 | Goal "(a+b) div (c::int) = a div c + b div c + ((a mod c + b mod c) div c)"; | 
| 7035 | 619 | by (zdiv_undefined_case_tac "c = #0" 1); | 
| 6992 | 620 | by (blast_tac (claset() addIs [[quorem_div_mod,quorem_div_mod] | 
| 621 | MRS lemma RS quorem_div]) 1); | |
| 622 | qed "zdiv_zadd1_eq"; | |
| 623 | ||
| 624 | Goal "(a+b) mod (c::int) = (a mod c + b mod c) mod c"; | |
| 7035 | 625 | by (zdiv_undefined_case_tac "c = #0" 1); | 
| 6992 | 626 | by (blast_tac (claset() addIs [[quorem_div_mod,quorem_div_mod] | 
| 627 | MRS lemma RS quorem_mod]) 1); | |
| 628 | qed "zmod_zadd1_eq"; | |
| 629 | ||
| 630 | Goal "(a mod b) div b = (#0::int)"; | |
| 7035 | 631 | by (zdiv_undefined_case_tac "b = #0" 1); | 
| 6992 | 632 | by (auto_tac (claset(), | 
| 6999 | 633 | simpset() addsimps [linorder_neq_iff, | 
| 634 | pos_mod_sign, pos_mod_bound, div_pos_pos_trivial, | |
| 635 | neg_mod_sign, neg_mod_bound, div_neg_neg_trivial])); | |
| 6992 | 636 | qed "mod_div_trivial"; | 
| 637 | Addsimps [mod_div_trivial]; | |
| 638 | ||
| 639 | Goal "(a mod b) mod b = a mod (b::int)"; | |
| 7035 | 640 | by (zdiv_undefined_case_tac "b = #0" 1); | 
| 6992 | 641 | by (auto_tac (claset(), | 
| 6999 | 642 | simpset() addsimps [linorder_neq_iff, | 
| 643 | pos_mod_sign, pos_mod_bound, mod_pos_pos_trivial, | |
| 644 | neg_mod_sign, neg_mod_bound, mod_neg_neg_trivial])); | |
| 6992 | 645 | qed "mod_mod_trivial"; | 
| 646 | Addsimps [mod_mod_trivial]; | |
| 647 | ||
| 9509 | 648 | Goal "(a+b) mod (c::int) = ((a mod c) + b) mod c"; | 
| 649 | by (rtac (trans RS sym) 1); | |
| 650 | by (rtac zmod_zadd1_eq 1); | |
| 651 | by (Simp_tac 1); | |
| 652 | by (rtac (zmod_zadd1_eq RS sym) 1); | |
| 653 | qed "zmod_zadd_left_eq"; | |
| 654 | ||
| 655 | Goal "(a+b) mod (c::int) = (a + (b mod c)) mod c"; | |
| 656 | by (rtac (trans RS sym) 1); | |
| 657 | by (rtac zmod_zadd1_eq 1); | |
| 658 | by (Simp_tac 1); | |
| 659 | by (rtac (zmod_zadd1_eq RS sym) 1); | |
| 660 | qed "zmod_zadd_right_eq"; | |
| 661 | ||
| 6992 | 662 | |
| 663 | Goal "a ~= (#0::int) ==> (a+b) div a = b div a + #1"; | |
| 664 | by (asm_simp_tac (simpset() addsimps [zdiv_zadd1_eq]) 1); | |
| 665 | qed "zdiv_zadd_self1"; | |
| 666 | ||
| 667 | Goal "a ~= (#0::int) ==> (b+a) div a = b div a + #1"; | |
| 668 | by (asm_simp_tac (simpset() addsimps [zdiv_zadd1_eq]) 1); | |
| 669 | qed "zdiv_zadd_self2"; | |
| 670 | Addsimps [zdiv_zadd_self1, zdiv_zadd_self2]; | |
| 671 | ||
| 672 | Goal "(a+b) mod a = b mod (a::int)"; | |
| 7035 | 673 | by (zdiv_undefined_case_tac "a = #0" 1); | 
| 6992 | 674 | by (asm_simp_tac (simpset() addsimps [zmod_zadd1_eq]) 1); | 
| 675 | qed "zmod_zadd_self1"; | |
| 676 | ||
| 677 | Goal "(b+a) mod a = b mod (a::int)"; | |
| 7035 | 678 | by (zdiv_undefined_case_tac "a = #0" 1); | 
| 6992 | 679 | by (asm_simp_tac (simpset() addsimps [zmod_zadd1_eq]) 1); | 
| 680 | qed "zmod_zadd_self2"; | |
| 681 | Addsimps [zmod_zadd_self1, zmod_zadd_self2]; | |
| 682 | ||
| 683 | ||
| 684 | (*** proving a div (b*c) = (a div b) div c ***) | |
| 685 | ||
| 686 | (*The condition c>0 seems necessary. Consider that 7 div ~6 = ~2 but | |
| 687 | 7 div 2 div ~3 = 3 div ~3 = ~1. The subcase (a div b) mod c = 0 seems | |
| 688 | to cause particular problems.*) | |
| 689 | ||
| 690 | (** first, four lemmas to bound the remainder for the cases b<0 and b>0 **) | |
| 6917 | 691 | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 692 | Goal "[| (#0::int) < c; b < r; r <= #0 |] ==> b*c < b*(q mod c) + r"; | 
| 6992 | 693 | by (subgoal_tac "b * (c - q mod c) < r * #1" 1); | 
| 694 | by (asm_full_simp_tac (simpset() addsimps [zdiff_zmult_distrib2]) 1); | |
| 695 | by (rtac order_le_less_trans 1); | |
| 696 | by (etac zmult_zless_mono1 2); | |
| 697 | by (rtac zmult_zle_mono2_neg 1); | |
| 698 | by (auto_tac | |
| 699 | (claset(), | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 700 | simpset() addsimps zcompare_rls@ | 
| 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 701 | [zadd_commute, add1_zle_eq, pos_mod_bound])); | 
| 6992 | 702 | val lemma1 = result(); | 
| 703 | ||
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 704 | Goal "[| (#0::int) < c; b < r; r <= #0 |] ==> b * (q mod c) + r <= #0"; | 
| 6992 | 705 | by (subgoal_tac "b * (q mod c) <= #0" 1); | 
| 706 | by (arith_tac 1); | |
| 9063 | 707 | by (asm_simp_tac (simpset() addsimps [zmult_le_0_iff, pos_mod_sign]) 1); | 
| 6992 | 708 | val lemma2 = result(); | 
| 709 | ||
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 710 | Goal "[| (#0::int) < c; #0 <= r; r < b |] ==> #0 <= b * (q mod c) + r"; | 
| 6992 | 711 | by (subgoal_tac "#0 <= b * (q mod c)" 1); | 
| 712 | by (arith_tac 1); | |
| 9063 | 713 | by (asm_simp_tac (simpset() addsimps [int_0_le_mult_iff, pos_mod_sign]) 1); | 
| 6992 | 714 | val lemma3 = result(); | 
| 715 | ||
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 716 | Goal "[| (#0::int) < c; #0 <= r; r < b |] ==> b * (q mod c) + r < b * c"; | 
| 6992 | 717 | by (subgoal_tac "r * #1 < b * (c - q mod c)" 1); | 
| 718 | by (asm_full_simp_tac (simpset() addsimps [zdiff_zmult_distrib2]) 1); | |
| 719 | by (rtac order_less_le_trans 1); | |
| 720 | by (etac zmult_zless_mono1 1); | |
| 721 | by (rtac zmult_zle_mono2 2); | |
| 722 | by (auto_tac | |
| 723 | (claset(), | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 724 | simpset() addsimps zcompare_rls@ | 
| 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 725 | [zadd_commute, add1_zle_eq, pos_mod_bound])); | 
| 6992 | 726 | val lemma4 = result(); | 
| 727 | ||
| 728 | Goal "[| quorem ((a,b), (q,r)); b ~= #0; #0 < c |] \ | |
| 729 | \ ==> quorem ((a, b*c), (q div c, b*(q mod c) + r))"; | |
| 8918 | 730 | by (auto_tac | 
| 6992 | 731 | (claset(), | 
| 8918 | 732 | simpset() addsimps zmult_ac@ | 
| 733 | [zmod_zdiv_equality, quorem_def, linorder_neq_iff, | |
| 9063 | 734 | int_0_less_mult_iff, | 
| 6992 | 735 | zadd_zmult_distrib2 RS sym, | 
| 736 | lemma1, lemma2, lemma3, lemma4])); | |
| 737 | val lemma = result(); | |
| 738 | ||
| 739 | Goal "(#0::int) < c ==> a div (b*c) = (a div b) div c"; | |
| 7035 | 740 | by (zdiv_undefined_case_tac "b = #0" 1); | 
| 6992 | 741 | by (force_tac (claset(), | 
| 742 | simpset() addsimps [quorem_div_mod RS lemma RS quorem_div, | |
| 743 | zmult_eq_0_iff]) 1); | |
| 744 | qed "zdiv_zmult2_eq"; | |
| 6917 | 745 | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 746 | Goal "(#0::int) < c ==> a mod (b*c) = b*(a div b mod c) + a mod b"; | 
| 7035 | 747 | by (zdiv_undefined_case_tac "b = #0" 1); | 
| 6992 | 748 | by (force_tac (claset(), | 
| 749 | simpset() addsimps [quorem_div_mod RS lemma RS quorem_mod, | |
| 750 | zmult_eq_0_iff]) 1); | |
| 751 | qed "zmod_zmult2_eq"; | |
| 752 | ||
| 753 | ||
| 754 | (*** Cancellation of common factors in "div" ***) | |
| 755 | ||
| 756 | Goal "[| (#0::int) < b; c ~= #0 |] ==> (c*a) div (c*b) = a div b"; | |
| 757 | by (stac zdiv_zmult2_eq 1); | |
| 758 | by Auto_tac; | |
| 759 | val lemma1 = result(); | |
| 760 | ||
| 761 | Goal "[| b < (#0::int); c ~= #0 |] ==> (c*a) div (c*b) = a div b"; | |
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7086diff
changeset | 762 | by (subgoal_tac "(c * (-a)) div (c * (-b)) = (-a) div (-b)" 1); | 
| 6992 | 763 | by (rtac lemma1 2); | 
| 8918 | 764 | by Auto_tac; | 
| 6992 | 765 | val lemma2 = result(); | 
| 766 | ||
| 767 | Goal "c ~= (#0::int) ==> (c*a) div (c*b) = a div b"; | |
| 7035 | 768 | by (zdiv_undefined_case_tac "b = #0" 1); | 
| 6992 | 769 | by (auto_tac | 
| 770 | (claset(), | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 771 |      simpset() addsimps [read_instantiate [("x", "b")] linorder_neq_iff, 
 | 
| 6992 | 772 | lemma1, lemma2])); | 
| 773 | qed "zdiv_zmult_zmult1"; | |
| 6917 | 774 | |
| 6992 | 775 | Goal "c ~= (#0::int) ==> (a*c) div (b*c) = a div b"; | 
| 776 | by (dtac zdiv_zmult_zmult1 1); | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 777 | by (auto_tac (claset(), simpset() addsimps [zmult_commute])); | 
| 6992 | 778 | qed "zdiv_zmult_zmult2"; | 
| 779 | ||
| 780 | ||
| 781 | ||
| 782 | (*** Distribution of factors over "mod" ***) | |
| 783 | ||
| 784 | Goal "[| (#0::int) < b; c ~= #0 |] ==> (c*a) mod (c*b) = c * (a mod b)"; | |
| 785 | by (stac zmod_zmult2_eq 1); | |
| 786 | by Auto_tac; | |
| 787 | val lemma1 = result(); | |
| 788 | ||
| 789 | Goal "[| b < (#0::int); c ~= #0 |] ==> (c*a) mod (c*b) = c * (a mod b)"; | |
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7086diff
changeset | 790 | by (subgoal_tac "(c * (-a)) mod (c * (-b)) = c * ((-a) mod (-b))" 1); | 
| 6992 | 791 | by (rtac lemma1 2); | 
| 8918 | 792 | by Auto_tac; | 
| 6992 | 793 | val lemma2 = result(); | 
| 794 | ||
| 6999 | 795 | Goal "(c*a) mod (c*b) = (c::int) * (a mod b)"; | 
| 7035 | 796 | by (zdiv_undefined_case_tac "b = #0" 1); | 
| 797 | by (zdiv_undefined_case_tac "c = #0" 1); | |
| 6992 | 798 | by (auto_tac | 
| 799 | (claset(), | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 800 |      simpset() addsimps [read_instantiate [("x", "b")] linorder_neq_iff, 
 | 
| 6992 | 801 | lemma1, lemma2])); | 
| 802 | qed "zmod_zmult_zmult1"; | |
| 803 | ||
| 6999 | 804 | Goal "(a*c) mod (b*c) = (a mod b) * (c::int)"; | 
| 805 | by (cut_inst_tac [("c","c")] zmod_zmult_zmult1 1);
 | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 806 | by (auto_tac (claset(), simpset() addsimps [zmult_commute])); | 
| 6992 | 807 | qed "zmod_zmult_zmult2"; | 
| 808 | ||
| 809 | ||
| 6999 | 810 | (*** Speeding up the division algorithm with shifting ***) | 
| 6992 | 811 | |
| 7035 | 812 | (** computing "div" by shifting **) | 
| 6999 | 813 | |
| 814 | Goal "(#0::int) <= a ==> (#1 + #2*b) div (#2*a) = b div a"; | |
| 7035 | 815 | by (zdiv_undefined_case_tac "a = #0" 1); | 
| 6999 | 816 | by (subgoal_tac "#1 <= a" 1); | 
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 817 | by (arith_tac 2); | 
| 6999 | 818 | by (subgoal_tac "#1 < a * #2" 1); | 
| 8785 | 819 | by (arith_tac 2); | 
| 6999 | 820 | by (subgoal_tac "#2*(#1 + b mod a) <= #2*a" 1); | 
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 821 | by (rtac zmult_zle_mono2 2); | 
| 6999 | 822 | by (auto_tac (claset(), | 
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 823 | simpset() addsimps [zadd_commute, zmult_commute, | 
| 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 824 | add1_zle_eq, pos_mod_bound])); | 
| 6999 | 825 | by (stac zdiv_zadd1_eq 1); | 
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 826 | by (asm_simp_tac (simpset() addsimps [zdiv_zmult_zmult2, zmod_zmult_zmult2, | 
| 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 827 | div_pos_pos_trivial]) 1); | 
| 6999 | 828 | by (stac div_pos_pos_trivial 1); | 
| 8765 
1bc30ff5fc54
[Int_CC.sum_conv, Int_CC.rel_conv] no longer exist
 paulson parents: 
8624diff
changeset | 829 | by (asm_simp_tac (simpset() | 
| 8918 | 830 | addsimps [mod_pos_pos_trivial, | 
| 7549 | 831 | pos_mod_sign RS zadd_zle_mono1 RSN (2,order_trans)]) 1); | 
| 6999 | 832 | by (auto_tac (claset(), | 
| 833 | simpset() addsimps [mod_pos_pos_trivial])); | |
| 8785 | 834 | by (subgoal_tac "#0 <= b mod a" 1); | 
| 835 | by (asm_simp_tac (simpset() addsimps [pos_mod_sign]) 2); | |
| 836 | by (arith_tac 1); | |
| 9063 | 837 | qed "pos_zdiv_mult_2"; | 
| 6999 | 838 | |
| 839 | ||
| 840 | Goal "a <= (#0::int) ==> (#1 + #2*b) div (#2*a) = (b+#1) div a"; | |
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7086diff
changeset | 841 | by (subgoal_tac "(#1 + #2*(-b-#1)) div (#2 * (-a)) = (-b-#1) div (-a)" 1); | 
| 9063 | 842 | by (rtac pos_zdiv_mult_2 2); | 
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 843 | by (auto_tac (claset(), | 
| 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 844 | simpset() addsimps [zmult_zminus_right])); | 
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7086diff
changeset | 845 | by (subgoal_tac "(#-1 - (#2 * b)) = - (#1 + (#2 * b))" 1); | 
| 6999 | 846 | by (Simp_tac 2); | 
| 847 | by (asm_full_simp_tac (HOL_ss | |
| 848 | addsimps [zdiv_zminus_zminus, zdiff_def, | |
| 849 | zminus_zadd_distrib RS sym]) 1); | |
| 9063 | 850 | qed "neg_zdiv_mult_2"; | 
| 6999 | 851 | |
| 852 | ||
| 853 | (*Not clear why this must be proved separately; probably number_of causes | |
| 854 | simplification problems*) | |
| 855 | Goal "~ #0 <= x ==> x <= (#0::int)"; | |
| 7035 | 856 | by Auto_tac; | 
| 6999 | 857 | val lemma = result(); | 
| 858 | ||
| 859 | Goal "number_of (v BIT b) div number_of (w BIT False) = \ | |
| 860 | \ (if ~b | (#0::int) <= number_of w \ | |
| 861 | \ then number_of v div (number_of w) \ | |
| 862 | \ else (number_of v + (#1::int)) div (number_of w))"; | |
| 9422 | 863 | by (simp_tac (simpset_of Int.thy addsimps [zadd_assoc, number_of_BIT]) 1); | 
| 7035 | 864 | by (asm_simp_tac (simpset() | 
| 7549 | 865 | delsimps bin_arith_extra_simps@bin_rel_simps | 
| 8787 | 866 | addsimps [zdiv_zmult_zmult1, | 
| 9063 | 867 | pos_zdiv_mult_2, lemma, neg_zdiv_mult_2]) 1); | 
| 6999 | 868 | qed "zdiv_number_of_BIT"; | 
| 869 | ||
| 870 | Addsimps [zdiv_number_of_BIT]; | |
| 871 | ||
| 872 | ||
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 873 | (** computing "mod" by shifting (proofs resemble those for "div") **) | 
| 7035 | 874 | |
| 875 | Goal "(#0::int) <= a ==> (#1 + #2*b) mod (#2*a) = #1 + #2 * (b mod a)"; | |
| 876 | by (zdiv_undefined_case_tac "a = #0" 1); | |
| 877 | by (subgoal_tac "#1 <= a" 1); | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 878 | by (arith_tac 2); | 
| 7035 | 879 | by (subgoal_tac "#1 < a * #2" 1); | 
| 8785 | 880 | by (arith_tac 2); | 
| 7035 | 881 | by (subgoal_tac "#2*(#1 + b mod a) <= #2*a" 1); | 
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 882 | by (rtac zmult_zle_mono2 2); | 
| 7035 | 883 | by (auto_tac (claset(), | 
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 884 | simpset() addsimps [zadd_commute, zmult_commute, | 
| 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 885 | add1_zle_eq, pos_mod_bound])); | 
| 7035 | 886 | by (stac zmod_zadd1_eq 1); | 
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 887 | by (asm_simp_tac (simpset() addsimps [zmod_zmult_zmult2, | 
| 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 888 | mod_pos_pos_trivial]) 1); | 
| 7035 | 889 | by (rtac mod_pos_pos_trivial 1); | 
| 8765 
1bc30ff5fc54
[Int_CC.sum_conv, Int_CC.rel_conv] no longer exist
 paulson parents: 
8624diff
changeset | 890 | by (asm_simp_tac (simpset() | 
| 8918 | 891 | addsimps [mod_pos_pos_trivial, | 
| 7549 | 892 | pos_mod_sign RS zadd_zle_mono1 RSN (2,order_trans)]) 1); | 
| 7035 | 893 | by (auto_tac (claset(), | 
| 894 | simpset() addsimps [mod_pos_pos_trivial])); | |
| 8785 | 895 | by (subgoal_tac "#0 <= b mod a" 1); | 
| 896 | by (asm_simp_tac (simpset() addsimps [pos_mod_sign]) 2); | |
| 897 | by (arith_tac 1); | |
| 9063 | 898 | qed "pos_zmod_mult_2"; | 
| 7035 | 899 | |
| 900 | ||
| 901 | Goal "a <= (#0::int) ==> (#1 + #2*b) mod (#2*a) = #2 * ((b+#1) mod a) - #1"; | |
| 902 | by (subgoal_tac | |
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7086diff
changeset | 903 | "(#1 + #2*(-b-#1)) mod (#2*(-a)) = #1 + #2*((-b-#1) mod (-a))" 1); | 
| 9063 | 904 | by (rtac pos_zmod_mult_2 2); | 
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 905 | by (auto_tac (claset(), | 
| 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 906 | simpset() addsimps [zmult_zminus_right])); | 
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7086diff
changeset | 907 | by (subgoal_tac "(#-1 - (#2 * b)) = - (#1 + (#2 * b))" 1); | 
| 7035 | 908 | by (Simp_tac 2); | 
| 909 | by (asm_full_simp_tac (HOL_ss | |
| 910 | addsimps [zmod_zminus_zminus, zdiff_def, | |
| 911 | zminus_zadd_distrib RS sym]) 1); | |
| 7086 | 912 | by (dtac (zminus_equation RS iffD1 RS sym) 1); | 
| 8918 | 913 | by Auto_tac; | 
| 9063 | 914 | qed "neg_zmod_mult_2"; | 
| 7035 | 915 | |
| 916 | Goal "number_of (v BIT b) mod number_of (w BIT False) = \ | |
| 917 | \ (if b then \ | |
| 918 | \ if (#0::int) <= number_of w \ | |
| 919 | \ then #2 * (number_of v mod number_of w) + #1 \ | |
| 920 | \ else #2 * ((number_of v + (#1::int)) mod number_of w) - #1 \ | |
| 921 | \ else #2 * (number_of v mod number_of w))"; | |
| 9422 | 922 | by (simp_tac (simpset_of Int.thy addsimps [zadd_assoc, number_of_BIT]) 1); | 
| 7035 | 923 | by (asm_simp_tac (simpset() | 
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 924 | delsimps bin_arith_extra_simps@bin_rel_simps | 
| 8787 | 925 | addsimps [zmod_zmult_zmult1, | 
| 9063 | 926 | pos_zmod_mult_2, lemma, neg_zmod_mult_2]) 1); | 
| 7035 | 927 | qed "zmod_number_of_BIT"; | 
| 928 | ||
| 929 | Addsimps [zmod_number_of_BIT]; | |
| 930 | ||
| 931 | ||
| 932 | (** Quotients of signs **) | |
| 933 | ||
| 934 | Goal "[| a < (#0::int); #0 < b |] ==> a div b < #0"; | |
| 935 | by (subgoal_tac "a div b <= #-1" 1); | |
| 936 | by (Force_tac 1); | |
| 937 | by (rtac order_trans 1); | |
| 938 | by (res_inst_tac [("a'","#-1")]  zdiv_mono1 1);
 | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7035diff
changeset | 939 | by (auto_tac (claset(), simpset() addsimps [zdiv_minus1])); | 
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7086diff
changeset | 940 | qed "div_neg_pos_less0"; | 
| 7035 | 941 | |
| 942 | Goal "[| (#0::int) <= a; b < #0 |] ==> a div b <= #0"; | |
| 943 | by (dtac zdiv_mono1_neg 1); | |
| 944 | by Auto_tac; | |
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7086diff
changeset | 945 | qed "div_nonneg_neg_le0"; | 
| 7035 | 946 | |
| 947 | Goal "(#0::int) < b ==> (#0 <= a div b) = (#0 <= a)"; | |
| 948 | by Auto_tac; | |
| 949 | by (dtac zdiv_mono1 2); | |
| 950 | by (auto_tac (claset(), simpset() addsimps [linorder_neq_iff])); | |
| 951 | by (full_simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1); | |
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7086diff
changeset | 952 | by (blast_tac (claset() addIs [div_neg_pos_less0]) 1); | 
| 7035 | 953 | qed "pos_imp_zdiv_nonneg_iff"; | 
| 954 | ||
| 955 | Goal "b < (#0::int) ==> (#0 <= a div b) = (a <= (#0::int))"; | |
| 956 | by (stac (zdiv_zminus_zminus RS sym) 1); | |
| 957 | by (stac pos_imp_zdiv_nonneg_iff 1); | |
| 958 | by Auto_tac; | |
| 959 | qed "neg_imp_zdiv_nonneg_iff"; | |
| 960 | ||
| 961 | (*But not (a div b <= 0 iff a<=0); consider a=1, b=2 when a div b = 0.*) | |
| 962 | Goal "(#0::int) < b ==> (a div b < #0) = (a < #0)"; | |
| 963 | by (asm_simp_tac (simpset() addsimps [linorder_not_le RS sym, | |
| 964 | pos_imp_zdiv_nonneg_iff]) 1); | |
| 965 | qed "pos_imp_zdiv_neg_iff"; | |
| 966 | ||
| 967 | (*Again the law fails for <=: consider a = -1, b = -2 when a div b = 0*) | |
| 968 | Goal "b < (#0::int) ==> (a div b < #0) = (#0 < a)"; | |
| 969 | by (asm_simp_tac (simpset() addsimps [linorder_not_le RS sym, | |
| 970 | neg_imp_zdiv_nonneg_iff]) 1); | |
| 971 | qed "neg_imp_zdiv_neg_iff"; |