| author | wenzelm | 
| Wed, 27 Oct 1999 18:12:40 +0200 | |
| changeset 7956 | edaca60a54cd | 
| parent 7037 | 77d596a5ffae | 
| child 8423 | 3c19160b6432 | 
| permissions | -rw-r--r-- | 
| 5545 | 1 | (* Title: HOL/ex/BinEx.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1998 University of Cambridge | |
| 5199 | 5 | |
| 5545 | 6 | Examples of performing binary arithmetic by simplification | 
| 7 | ||
| 8 | Also a proof that binary arithmetic on normalized operands | |
| 9 | yields normalized results. | |
| 10 | *) | |
| 11 | ||
| 12 | set proof_timing; | |
| 5199 | 13 | |
| 7037 | 14 | (**** The Integers ****) | 
| 15 | ||
| 6920 | 16 | (*** Addition ***) | 
| 5491 | 17 | |
| 6909 | 18 | Goal "(#13::int) + #19 = #32"; | 
| 5199 | 19 | by (Simp_tac 1); | 
| 20 | result(); | |
| 21 | ||
| 6909 | 22 | Goal "(#1234::int) + #5678 = #6912"; | 
| 5199 | 23 | by (Simp_tac 1); | 
| 24 | result(); | |
| 25 | ||
| 6909 | 26 | Goal "(#1359::int) + #-2468 = #-1109"; | 
| 5199 | 27 | by (Simp_tac 1); | 
| 28 | result(); | |
| 29 | ||
| 6909 | 30 | Goal "(#93746::int) + #-46375 = #47371"; | 
| 5199 | 31 | by (Simp_tac 1); | 
| 32 | result(); | |
| 33 | ||
| 6920 | 34 | (*** Negation ***) | 
| 5491 | 35 | |
| 6909 | 36 | Goal "- (#65745::int) = #-65745"; | 
| 5199 | 37 | by (Simp_tac 1); | 
| 38 | result(); | |
| 39 | ||
| 6909 | 40 | Goal "- (#-54321::int) = #54321"; | 
| 5199 | 41 | by (Simp_tac 1); | 
| 42 | result(); | |
| 43 | ||
| 5491 | 44 | |
| 6920 | 45 | (*** Multiplication ***) | 
| 5491 | 46 | |
| 6909 | 47 | Goal "(#13::int) * #19 = #247"; | 
| 5199 | 48 | by (Simp_tac 1); | 
| 49 | result(); | |
| 50 | ||
| 6909 | 51 | Goal "(#-84::int) * #51 = #-4284"; | 
| 5199 | 52 | by (Simp_tac 1); | 
| 53 | result(); | |
| 54 | ||
| 6909 | 55 | Goal "(#255::int) * #255 = #65025"; | 
| 5199 | 56 | by (Simp_tac 1); | 
| 57 | result(); | |
| 58 | ||
| 6909 | 59 | Goal "(#1359::int) * #-2468 = #-3354012"; | 
| 5199 | 60 | by (Simp_tac 1); | 
| 61 | result(); | |
| 62 | ||
| 6909 | 63 | Goal "(#89::int) * #10 ~= #889"; | 
| 5199 | 64 | by (Simp_tac 1); | 
| 65 | result(); | |
| 66 | ||
| 6909 | 67 | Goal "(#13::int) < #18 - #4"; | 
| 5199 | 68 | by (Simp_tac 1); | 
| 69 | result(); | |
| 70 | ||
| 6909 | 71 | Goal "(#-345::int) < #-242 + #-100"; | 
| 5199 | 72 | by (Simp_tac 1); | 
| 73 | result(); | |
| 74 | ||
| 6909 | 75 | Goal "(#13557456::int) < #18678654"; | 
| 5199 | 76 | by (Simp_tac 1); | 
| 77 | result(); | |
| 78 | ||
| 6909 | 79 | Goal "(#999999::int) <= (#1000001 + #1)-#2"; | 
| 5199 | 80 | by (Simp_tac 1); | 
| 81 | result(); | |
| 82 | ||
| 6909 | 83 | Goal "(#1234567::int) <= #1234567"; | 
| 5199 | 84 | by (Simp_tac 1); | 
| 85 | result(); | |
| 5545 | 86 | |
| 87 | ||
| 7037 | 88 | (*** Quotient and Remainder ***) | 
| 6920 | 89 | |
| 90 | Goal "(#10::int) div #3 = #3"; | |
| 91 | by (Simp_tac 1); | |
| 92 | result(); | |
| 93 | ||
| 94 | Goal "(#10::int) mod #3 = #1"; | |
| 95 | by (Simp_tac 1); | |
| 96 | result(); | |
| 97 | ||
| 98 | (** A negative divisor **) | |
| 99 | ||
| 100 | Goal "(#10::int) div #-3 = #-4"; | |
| 101 | by (Simp_tac 1); | |
| 102 | result(); | |
| 103 | ||
| 104 | Goal "(#10::int) mod #-3 = #-2"; | |
| 105 | by (Simp_tac 1); | |
| 106 | result(); | |
| 107 | ||
| 108 | (** A negative dividend | |
| 109 | [ The definition agrees with mathematical convention but not with | |
| 110 | the hardware of most computers ] | |
| 111 | **) | |
| 112 | ||
| 113 | Goal "(#-10::int) div #3 = #-4"; | |
| 114 | by (Simp_tac 1); | |
| 115 | result(); | |
| 116 | ||
| 117 | Goal "(#-10::int) mod #3 = #2"; | |
| 118 | by (Simp_tac 1); | |
| 119 | result(); | |
| 120 | ||
| 121 | (** A negative dividend AND divisor **) | |
| 122 | ||
| 123 | Goal "(#-10::int) div #-3 = #3"; | |
| 124 | by (Simp_tac 1); | |
| 125 | result(); | |
| 126 | ||
| 127 | Goal "(#-10::int) mod #-3 = #-1"; | |
| 128 | by (Simp_tac 1); | |
| 129 | result(); | |
| 130 | ||
| 131 | (** A few bigger examples **) | |
| 132 | ||
| 133 | Goal "(#8452::int) mod #3 = #1"; | |
| 134 | by (Simp_tac 1); | |
| 135 | result(); | |
| 136 | ||
| 137 | Goal "(#59485::int) div #434 = #137"; | |
| 138 | by (Simp_tac 1); | |
| 139 | result(); | |
| 140 | ||
| 141 | Goal "(#1000006::int) mod #10 = #6"; | |
| 142 | by (Simp_tac 1); | |
| 143 | result(); | |
| 144 | ||
| 145 | ||
| 7037 | 146 | (** division by shifting **) | 
| 147 | ||
| 148 | Goal "#10000000 div #2 = (#5000000::int)"; | |
| 149 | by (Simp_tac 1); | |
| 150 | result(); | |
| 151 | ||
| 152 | Goal "#10000001 mod #2 = (#1::int)"; | |
| 153 | by (Simp_tac 1); | |
| 154 | result(); | |
| 155 | ||
| 156 | Goal "#10000055 div #32 = (#312501::int)"; | |
| 157 | by (Simp_tac 1); | |
| 158 | ||
| 159 | Goal "#10000055 mod #32 = (#23::int)"; | |
| 160 | by (Simp_tac 1); | |
| 161 | ||
| 162 | Goal "#100094 div #144 = (#695::int)"; | |
| 163 | by (Simp_tac 1); | |
| 164 | result(); | |
| 165 | ||
| 166 | Goal "#100094 mod #144 = (#14::int)"; | |
| 167 | by (Simp_tac 1); | |
| 168 | result(); | |
| 169 | ||
| 170 | ||
| 171 | ||
| 172 | (**** The Natural Numbers ****) | |
| 173 | ||
| 174 | (** Successor **) | |
| 175 | ||
| 176 | Goal "Suc #99999 = #100000"; | |
| 177 | by (asm_simp_tac (simpset() addsimps [Suc_nat_number_of]) 1); | |
| 178 | (*not a default rewrite since sometimes we want to have Suc(#nnn)*) | |
| 179 | result(); | |
| 180 | ||
| 181 | ||
| 182 | (** Addition **) | |
| 183 | ||
| 184 | Goal "(#13::nat) + #19 = #32"; | |
| 185 | by (Simp_tac 1); | |
| 186 | result(); | |
| 187 | ||
| 188 | Goal "(#1234::nat) + #5678 = #6912"; | |
| 189 | by (Simp_tac 1); | |
| 190 | result(); | |
| 191 | ||
| 192 | Goal "(#973646::nat) + #6475 = #980121"; | |
| 193 | by (Simp_tac 1); | |
| 194 | result(); | |
| 195 | ||
| 196 | ||
| 197 | (** Subtraction **) | |
| 198 | ||
| 199 | Goal "(#32::nat) - #14 = #18"; | |
| 200 | by (Simp_tac 1); | |
| 201 | result(); | |
| 202 | ||
| 203 | Goal "(#14::nat) - #15 = #0"; | |
| 204 | by (Simp_tac 1); | |
| 205 | result(); | |
| 206 | ||
| 207 | Goal "(#14::nat) - #1576644 = #0"; | |
| 208 | by (Simp_tac 1); | |
| 209 | result(); | |
| 210 | ||
| 211 | Goal "(#48273776::nat) - #3873737 = #44400039"; | |
| 212 | by (Simp_tac 1); | |
| 213 | result(); | |
| 214 | ||
| 215 | ||
| 216 | (** Multiplication **) | |
| 217 | ||
| 218 | Goal "(#12::nat) * #11 = #132"; | |
| 219 | by (Simp_tac 1); | |
| 220 | result(); | |
| 221 | ||
| 222 | Goal "(#647::nat) * #3643 = #2357021"; | |
| 223 | by (Simp_tac 1); | |
| 224 | result(); | |
| 225 | ||
| 226 | ||
| 227 | (** Quotient and Remainder **) | |
| 228 | ||
| 229 | Goal "(#10::nat) div #3 = #3"; | |
| 230 | by (Simp_tac 1); | |
| 231 | result(); | |
| 232 | ||
| 233 | Goal "(#10::nat) mod #3 = #1"; | |
| 234 | by (Simp_tac 1); | |
| 235 | result(); | |
| 236 | ||
| 237 | Goal "(#10000::nat) div #9 = #1111"; | |
| 238 | by (Simp_tac 1); | |
| 239 | result(); | |
| 240 | ||
| 241 | Goal "(#10000::nat) mod #9 = #1"; | |
| 242 | by (Simp_tac 1); | |
| 243 | result(); | |
| 244 | ||
| 245 | Goal "(#10000::nat) div #16 = #625"; | |
| 246 | by (Simp_tac 1); | |
| 247 | result(); | |
| 248 | ||
| 249 | Goal "(#10000::nat) mod #16 = #0"; | |
| 250 | by (Simp_tac 1); | |
| 251 | result(); | |
| 252 | ||
| 253 | ||
| 6920 | 254 | (*** Testing the cancellation of complementary terms ***) | 
| 5601 | 255 | |
| 6909 | 256 | Goal "y + (x + -x) = (#0::int) + y"; | 
| 5601 | 257 | by (Simp_tac 1); | 
| 258 | result(); | |
| 259 | ||
| 6909 | 260 | Goal "y + (-x + (- y + x)) = (#0::int)"; | 
| 5601 | 261 | by (Simp_tac 1); | 
| 262 | result(); | |
| 263 | ||
| 6909 | 264 | Goal "-x + (y + (- y + x)) = (#0::int)"; | 
| 5601 | 265 | by (Simp_tac 1); | 
| 266 | result(); | |
| 267 | ||
| 6909 | 268 | Goal "x + (x + (- x + (- x + (- y + - z)))) = (#0::int) - y - z"; | 
| 5601 | 269 | by (Simp_tac 1); | 
| 270 | result(); | |
| 271 | ||
| 6909 | 272 | Goal "x + x - x - x - y - z = (#0::int) - y - z"; | 
| 5601 | 273 | by (Simp_tac 1); | 
| 274 | result(); | |
| 275 | ||
| 6909 | 276 | Goal "x + y + z - (x + z) = y - (#0::int)"; | 
| 5601 | 277 | by (Simp_tac 1); | 
| 278 | result(); | |
| 279 | ||
| 6909 | 280 | Goal "x+(y+(y+(y+ (-x + -x)))) = (#0::int) + y - x + y + y"; | 
| 5601 | 281 | by (Simp_tac 1); | 
| 282 | result(); | |
| 283 | ||
| 6909 | 284 | Goal "x+(y+(y+(y+ (-y + -x)))) = y + (#0::int) + y"; | 
| 5601 | 285 | by (Simp_tac 1); | 
| 286 | result(); | |
| 287 | ||
| 6909 | 288 | Goal "x + y - x + z - x - y - z + x < (#1::int)"; | 
| 5601 | 289 | by (Simp_tac 1); | 
| 290 | result(); | |
| 291 | ||
| 292 | ||
| 5545 | 293 | Addsimps normal.intrs; | 
| 294 | ||
| 295 | Goal "(w BIT b): normal ==> (w BIT b BIT c): normal"; | |
| 296 | by (case_tac "c" 1); | |
| 297 | by Auto_tac; | |
| 298 | qed "normal_BIT_I"; | |
| 299 | ||
| 300 | Addsimps [normal_BIT_I]; | |
| 301 | ||
| 302 | Goal "w BIT b: normal ==> w: normal"; | |
| 303 | by (etac normal.elim 1); | |
| 304 | by Auto_tac; | |
| 305 | qed "normal_BIT_D"; | |
| 306 | ||
| 307 | Goal "w : normal --> NCons w b : normal"; | |
| 308 | by (induct_tac "w" 1); | |
| 309 | by (auto_tac (claset(), simpset() addsimps [NCons_Pls, NCons_Min])); | |
| 310 | qed_spec_mp "NCons_normal"; | |
| 311 | ||
| 312 | Addsimps [NCons_normal]; | |
| 313 | ||
| 314 | Goal "NCons w True ~= Pls"; | |
| 315 | by (induct_tac "w" 1); | |
| 316 | by Auto_tac; | |
| 317 | qed "NCons_True"; | |
| 318 | ||
| 319 | Goal "NCons w False ~= Min"; | |
| 320 | by (induct_tac "w" 1); | |
| 321 | by Auto_tac; | |
| 322 | qed "NCons_False"; | |
| 323 | ||
| 324 | Goal "w: normal ==> bin_succ w : normal"; | |
| 325 | by (etac normal.induct 1); | |
| 326 | by (exhaust_tac "w" 4); | |
| 327 | by (auto_tac (claset(), simpset() addsimps [NCons_True, bin_succ_BIT])); | |
| 328 | qed "bin_succ_normal"; | |
| 329 | ||
| 330 | Goal "w: normal ==> bin_pred w : normal"; | |
| 331 | by (etac normal.induct 1); | |
| 332 | by (exhaust_tac "w" 3); | |
| 333 | by (auto_tac (claset(), simpset() addsimps [NCons_False, bin_pred_BIT])); | |
| 334 | qed "bin_pred_normal"; | |
| 335 | ||
| 336 | Addsimps [bin_succ_normal, bin_pred_normal]; | |
| 337 | ||
| 338 | Goal "w : normal --> (ALL z. z: normal --> bin_add w z : normal)"; | |
| 339 | by (induct_tac "w" 1); | |
| 340 | by (Simp_tac 1); | |
| 341 | by (Simp_tac 1); | |
| 342 | by (rtac impI 1); | |
| 343 | by (rtac allI 1); | |
| 344 | by (induct_tac "z" 1); | |
| 345 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [bin_add_BIT]))); | |
| 346 | by (safe_tac (claset() addSDs [normal_BIT_D])); | |
| 347 | by (ALLGOALS Asm_simp_tac); | |
| 348 | qed_spec_mp "bin_add_normal"; | |
| 349 | ||
| 6909 | 350 | Goal "w: normal ==> (w = Pls) = (number_of w = (#0::int))"; | 
| 5545 | 351 | by (etac normal.induct 1); | 
| 352 | by Auto_tac; | |
| 353 | qed "normal_Pls_eq_0"; | |
| 354 | ||
| 355 | Goal "w : normal ==> bin_minus w : normal"; | |
| 356 | by (etac normal.induct 1); | |
| 357 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [bin_minus_BIT]))); | |
| 358 | by (resolve_tac normal.intrs 1); | |
| 359 | by (assume_tac 1); | |
| 360 | by (asm_full_simp_tac (simpset() addsimps [normal_Pls_eq_0]) 1); | |
| 361 | by (asm_full_simp_tac | |
| 6920 | 362 | (simpset_of Int.thy | 
| 363 | addsimps [number_of_minus, iszero_def, | |
| 364 | 	       read_instantiate [("y","int 0")] zminus_equation]) 1);
 | |
| 365 | by (etac not_sym 1); | |
| 5545 | 366 | qed "bin_minus_normal"; | 
| 367 | ||
| 368 | Goal "w : normal ==> z: normal --> bin_mult w z : normal"; | |
| 369 | by (etac normal.induct 1); | |
| 5569 
8c7e1190e789
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5545diff
changeset | 370 | by (ALLGOALS | 
| 
8c7e1190e789
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5545diff
changeset | 371 | (asm_simp_tac (simpset() addsimps [bin_minus_normal, bin_mult_BIT]))); | 
| 5545 | 372 | by (safe_tac (claset() addSDs [normal_BIT_D])); | 
| 373 | by (asm_simp_tac (simpset() addsimps [bin_add_normal]) 1); | |
| 374 | qed_spec_mp "bin_mult_normal"; | |
| 375 |