| author | paulson | 
| Fri, 25 Sep 1998 14:05:13 +0200 | |
| changeset 5565 | 301a3a4d3dc7 | 
| parent 5562 | 02261e6880d1 | 
| child 5582 | a356fb49e69e | 
| permissions | -rw-r--r-- | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 1 | (* Title: HOL/Integ/Bin.ML | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 2 | Authors: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 3 | David Spelt, University of Twente | 
| 1632 | 4 | Copyright 1994 University of Cambridge | 
| 5 | Copyright 1996 University of Twente | |
| 6 | ||
| 7 | Arithmetic on binary integers. | |
| 8 | *) | |
| 9 | ||
| 10 | (** extra rules for bin_succ, bin_pred, bin_add, bin_mult **) | |
| 11 | ||
| 5512 | 12 | qed_goal "NCons_Pls_0" Bin.thy | 
| 13 | "NCons Pls False = Pls" | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 14 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 15 | |
| 5512 | 16 | qed_goal "NCons_Pls_1" Bin.thy | 
| 17 | "NCons Pls True = Pls BIT True" | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 18 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 19 | |
| 5512 | 20 | qed_goal "NCons_Min_0" Bin.thy | 
| 21 | "NCons Min False = Min BIT False" | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 22 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 23 | |
| 5512 | 24 | qed_goal "NCons_Min_1" Bin.thy | 
| 25 | "NCons Min True = Min" | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 26 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 27 | |
| 5512 | 28 | qed_goal "bin_succ_1" Bin.thy | 
| 29 | "bin_succ(w BIT True) = (bin_succ w) BIT False" | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 30 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 31 | |
| 5512 | 32 | qed_goal "bin_succ_0" Bin.thy | 
| 33 | "bin_succ(w BIT False) = NCons w True" | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 34 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 35 | |
| 5512 | 36 | qed_goal "bin_pred_1" Bin.thy | 
| 37 | "bin_pred(w BIT True) = NCons w False" | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 38 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 39 | |
| 5512 | 40 | qed_goal "bin_pred_0" Bin.thy | 
| 41 | "bin_pred(w BIT False) = (bin_pred w) BIT True" | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 42 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 43 | |
| 5512 | 44 | qed_goal "bin_minus_1" Bin.thy | 
| 45 | "bin_minus(w BIT True) = bin_pred (NCons (bin_minus w) False)" | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 46 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 47 | |
| 5512 | 48 | qed_goal "bin_minus_0" Bin.thy | 
| 49 | "bin_minus(w BIT False) = (bin_minus w) BIT False" | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 50 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 51 | |
| 5491 | 52 | |
| 1632 | 53 | (*** bin_add: binary addition ***) | 
| 54 | ||
| 5512 | 55 | qed_goal "bin_add_BIT_11" Bin.thy | 
| 56 | "bin_add (v BIT True) (w BIT True) = \ | |
| 57 | \ NCons (bin_add v (bin_succ w)) False" | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 58 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 59 | |
| 5512 | 60 | qed_goal "bin_add_BIT_10" Bin.thy | 
| 61 | "bin_add (v BIT True) (w BIT False) = NCons (bin_add v w) True" | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 62 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 63 | |
| 5512 | 64 | qed_goal "bin_add_BIT_0" Bin.thy | 
| 65 | "bin_add (v BIT False) (w BIT y) = NCons (bin_add v w) y" | |
| 5491 | 66 | (fn _ => [Auto_tac]); | 
| 1632 | 67 | |
| 5551 | 68 | Goal "bin_add w Pls = w"; | 
| 69 | by (induct_tac "w" 1); | |
| 70 | by Auto_tac; | |
| 71 | qed "bin_add_Pls_right"; | |
| 1632 | 72 | |
| 5512 | 73 | qed_goal "bin_add_BIT_Min" Bin.thy | 
| 74 | "bin_add (v BIT x) Min = bin_pred (v BIT x)" | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 75 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 76 | |
| 5512 | 77 | qed_goal "bin_add_BIT_BIT" Bin.thy | 
| 78 | "bin_add (v BIT x) (w BIT y) = \ | |
| 79 | \ NCons(bin_add v (if x & y then (bin_succ w) else w)) (x~= y)" | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 80 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 81 | |
| 82 | ||
| 83 | (*** bin_add: binary multiplication ***) | |
| 84 | ||
| 5512 | 85 | qed_goal "bin_mult_1" Bin.thy | 
| 86 | "bin_mult (v BIT True) w = bin_add (NCons (bin_mult v w) False) w" | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 87 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 88 | |
| 5512 | 89 | qed_goal "bin_mult_0" Bin.thy | 
| 90 | "bin_mult (v BIT False) w = NCons (bin_mult v w) False" | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 91 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 92 | |
| 93 | ||
| 94 | (**** The carry/borrow functions, bin_succ and bin_pred ****) | |
| 95 | ||
| 96 | ||
| 5491 | 97 | (**** integ_of ****) | 
| 1632 | 98 | |
| 5512 | 99 | qed_goal "integ_of_NCons" Bin.thy | 
| 100 | "integ_of(NCons w b) = integ_of(w BIT b)" | |
| 5184 | 101 | (fn _ =>[(induct_tac "w" 1), | 
| 5510 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 102 | (ALLGOALS Asm_simp_tac) ]); | 
| 1632 | 103 | |
| 5512 | 104 | Addsimps [integ_of_NCons]; | 
| 1632 | 105 | |
| 5491 | 106 | qed_goal "integ_of_succ" Bin.thy | 
| 107 | "integ_of(bin_succ w) = $#1 + integ_of w" | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 108 | (fn _ =>[(rtac bin.induct 1), | 
| 5510 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 109 | (ALLGOALS(asm_simp_tac (simpset() addsimps zadd_ac))) ]); | 
| 5491 | 110 | |
| 111 | qed_goal "integ_of_pred" Bin.thy | |
| 112 | "integ_of(bin_pred w) = - ($#1) + integ_of w" | |
| 113 | (fn _ =>[(rtac bin.induct 1), | |
| 5510 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 114 | (ALLGOALS(asm_simp_tac (simpset() addsimps zadd_ac))) ]); | 
| 1632 | 115 | |
| 5491 | 116 | Goal "integ_of(bin_minus w) = - (integ_of w)"; | 
| 117 | by (rtac bin.induct 1); | |
| 118 | by (Simp_tac 1); | |
| 119 | by (Simp_tac 1); | |
| 120 | by (asm_simp_tac (simpset() | |
| 5551 | 121 | delsimps [bin_pred_Pls, bin_pred_Min, bin_pred_BIT] | 
| 5491 | 122 | addsimps [integ_of_succ,integ_of_pred, | 
| 123 | zadd_assoc]) 1); | |
| 124 | qed "integ_of_minus"; | |
| 1632 | 125 | |
| 126 | ||
| 5512 | 127 | val bin_add_simps = [bin_add_BIT_BIT, | 
| 5510 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 128 | integ_of_succ, integ_of_pred]; | 
| 1632 | 129 | |
| 5491 | 130 | Goal "! w. integ_of(bin_add v w) = integ_of v + integ_of w"; | 
| 5184 | 131 | by (induct_tac "v" 1); | 
| 4686 | 132 | by (simp_tac (simpset() addsimps bin_add_simps) 1); | 
| 133 | by (simp_tac (simpset() addsimps bin_add_simps) 1); | |
| 1632 | 134 | by (rtac allI 1); | 
| 5184 | 135 | by (induct_tac "w" 1); | 
| 5540 | 136 | by (ALLGOALS (asm_simp_tac (simpset() addsimps bin_add_simps @ zadd_ac))); | 
| 5491 | 137 | qed_spec_mp "integ_of_add"; | 
| 1632 | 138 | |
| 5512 | 139 | val bin_mult_simps = [zmult_zminus, integ_of_minus, integ_of_add]; | 
| 1632 | 140 | |
| 5491 | 141 | |
| 142 | Goal "integ_of(bin_mult v w) = integ_of v * integ_of w"; | |
| 5184 | 143 | by (induct_tac "v" 1); | 
| 4686 | 144 | by (simp_tac (simpset() addsimps bin_mult_simps) 1); | 
| 145 | by (simp_tac (simpset() addsimps bin_mult_simps) 1); | |
| 5491 | 146 | by (asm_simp_tac | 
| 5540 | 147 | (simpset() addsimps bin_mult_simps @ [zadd_zmult_distrib] @ zadd_ac) 1); | 
| 5491 | 148 | qed "integ_of_mult"; | 
| 149 | ||
| 1632 | 150 | |
| 5491 | 151 | (** Simplification rules with integer constants **) | 
| 152 | ||
| 153 | Goal "#0 + z = z"; | |
| 154 | by (Simp_tac 1); | |
| 155 | qed "zadd_0"; | |
| 156 | ||
| 157 | Goal "z + #0 = z"; | |
| 158 | by (Simp_tac 1); | |
| 159 | qed "zadd_0_right"; | |
| 160 | ||
| 161 | Goal "z + (- z) = #0"; | |
| 5510 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 162 | by (Simp_tac 1); | 
| 5491 | 163 | qed "zadd_zminus_inverse"; | 
| 164 | ||
| 165 | Goal "(- z) + z = #0"; | |
| 5510 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 166 | by (Simp_tac 1); | 
| 5491 | 167 | qed "zadd_zminus_inverse2"; | 
| 168 | ||
| 5510 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 169 | (*These rewrite to $# 0. Henceforth we should rewrite to #0 *) | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 170 | Delsimps [zadd_zminus_inverse_nat, zadd_zminus_inverse_nat2]; | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 171 | |
| 5491 | 172 | Addsimps [zadd_0, zadd_0_right, zadd_zminus_inverse, zadd_zminus_inverse2]; | 
| 173 | ||
| 174 | Goal "- (#0) = #0"; | |
| 175 | by (Simp_tac 1); | |
| 176 | qed "zminus_0"; | |
| 177 | ||
| 178 | Addsimps [zminus_0]; | |
| 179 | ||
| 180 | Goal "#0 * z = #0"; | |
| 181 | by (Simp_tac 1); | |
| 182 | qed "zmult_0"; | |
| 183 | ||
| 184 | Goal "#1 * z = z"; | |
| 185 | by (Simp_tac 1); | |
| 186 | qed "zmult_1"; | |
| 187 | ||
| 188 | Goal "#2 * z = z+z"; | |
| 189 | by (simp_tac (simpset() addsimps [zadd_zmult_distrib]) 1); | |
| 190 | qed "zmult_2"; | |
| 191 | ||
| 192 | Goal "z * #0 = #0"; | |
| 193 | by (Simp_tac 1); | |
| 194 | qed "zmult_0_right"; | |
| 195 | ||
| 196 | Goal "z * #1 = z"; | |
| 197 | by (Simp_tac 1); | |
| 198 | qed "zmult_1_right"; | |
| 199 | ||
| 200 | Goal "z * #2 = z+z"; | |
| 201 | by (simp_tac (simpset() addsimps [zadd_zmult_distrib2]) 1); | |
| 202 | qed "zmult_2_right"; | |
| 203 | ||
| 204 | Addsimps [zmult_0, zmult_0_right, | |
| 205 | zmult_1, zmult_1_right, | |
| 206 | zmult_2, zmult_2_right]; | |
| 207 | ||
| 208 | Goal "(w < z + #1) = (w<z | w=z)"; | |
| 209 | by (simp_tac (simpset() addsimps [zless_add_nat1_eq]) 1); | |
| 210 | qed "zless_add1_eq"; | |
| 211 | ||
| 212 | Goal "(w + #1 <= z) = (w<z)"; | |
| 213 | by (simp_tac (simpset() addsimps [add_nat1_zle_eq]) 1); | |
| 214 | qed "add1_zle_eq"; | |
| 215 | Addsimps [add1_zle_eq]; | |
| 216 | ||
| 5540 | 217 | Goal "neg x = (x < #0)"; | 
| 218 | by (simp_tac (simpset() addsimps [neg_eq_less_nat0]) 1); | |
| 219 | qed "neg_eq_less_0"; | |
| 5510 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 220 | |
| 5540 | 221 | Goal "(~neg x) = ($# 0 <= x)"; | 
| 222 | by (simp_tac (simpset() addsimps [not_neg_eq_ge_nat0]) 1); | |
| 223 | qed "not_neg_eq_ge_0"; | |
| 5510 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 224 | |
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 225 | |
| 5491 | 226 | |
| 227 | (** Simplification rules for comparison of binary numbers (Norbert Voelker) **) | |
| 228 | ||
| 229 | (** Equals (=) **) | |
| 1632 | 230 | |
| 5491 | 231 | Goalw [iszero_def] | 
| 232 | "(integ_of x = integ_of y) \ | |
| 233 | \ = iszero(integ_of (bin_add x (bin_minus y)))"; | |
| 234 | by (simp_tac (simpset() addsimps | |
| 235 | (zcompare_rls @ [integ_of_add, integ_of_minus])) 1); | |
| 236 | qed "eq_integ_of_eq"; | |
| 237 | ||
| 5512 | 238 | Goalw [iszero_def] "iszero (integ_of Pls)"; | 
| 5491 | 239 | by (Simp_tac 1); | 
| 5512 | 240 | qed "iszero_integ_of_Pls"; | 
| 5491 | 241 | |
| 5512 | 242 | Goalw [iszero_def] "~ iszero(integ_of Min)"; | 
| 5491 | 243 | by (Simp_tac 1); | 
| 5512 | 244 | qed "nonzero_integ_of_Min"; | 
| 5491 | 245 | |
| 246 | Goalw [iszero_def] | |
| 5512 | 247 | "iszero (integ_of (w BIT x)) = (~x & iszero (integ_of w))"; | 
| 5491 | 248 | by (Simp_tac 1); | 
| 249 | by (int_case_tac "integ_of w" 1); | |
| 250 | by (ALLGOALS (asm_simp_tac | |
| 5540 | 251 | (simpset() addsimps zcompare_rls @ | 
| 252 | [zminus_zadd_distrib RS sym, | |
| 253 | add_nat]))); | |
| 5512 | 254 | qed "iszero_integ_of_BIT"; | 
| 5491 | 255 | |
| 256 | ||
| 257 | (** Less-than (<) **) | |
| 258 | ||
| 259 | Goalw [zless_def,zdiff_def] | |
| 260 | "integ_of x < integ_of y \ | |
| 5540 | 261 | \ = neg (integ_of (bin_add x (bin_minus y)))"; | 
| 5491 | 262 | by (simp_tac (simpset() addsimps bin_mult_simps) 1); | 
| 5540 | 263 | qed "less_integ_of_eq_neg"; | 
| 5491 | 264 | |
| 5540 | 265 | Goal "~ neg (integ_of Pls)"; | 
| 5491 | 266 | by (Simp_tac 1); | 
| 5512 | 267 | qed "not_neg_integ_of_Pls"; | 
| 5491 | 268 | |
| 5540 | 269 | Goal "neg (integ_of Min)"; | 
| 5491 | 270 | by (Simp_tac 1); | 
| 5512 | 271 | qed "neg_integ_of_Min"; | 
| 5491 | 272 | |
| 5540 | 273 | Goal "neg (integ_of (w BIT x)) = neg (integ_of w)"; | 
| 5491 | 274 | by (Asm_simp_tac 1); | 
| 275 | by (int_case_tac "integ_of w" 1); | |
| 276 | by (ALLGOALS (asm_simp_tac | |
| 5540 | 277 | (simpset() addsimps [add_nat, neg_eq_less_nat0, | 
| 278 | symmetric zdiff_def] @ zcompare_rls))); | |
| 5512 | 279 | qed "neg_integ_of_BIT"; | 
| 5491 | 280 | |
| 281 | ||
| 282 | (** Less-than-or-equals (<=) **) | |
| 283 | ||
| 284 | Goal "(integ_of x <= integ_of y) = (~ integ_of y < integ_of x)"; | |
| 285 | by (simp_tac (simpset() addsimps [zle_def]) 1); | |
| 286 | qed "le_integ_of_eq_not_less"; | |
| 287 | ||
| 5540 | 288 | (*Delete the original rewrites, with their clumsy conditional expressions*) | 
| 5551 | 289 | Delsimps [bin_succ_BIT, bin_pred_BIT, bin_minus_BIT, | 
| 290 | NCons_Pls, NCons_Min, bin_add_BIT, bin_mult_BIT]; | |
| 5491 | 291 | |
| 292 | (*Hide the binary representation of integer constants*) | |
| 5540 | 293 | Delsimps [integ_of_Pls, integ_of_Min, integ_of_BIT]; | 
| 5491 | 294 | |
| 295 | (*Add simplification of arithmetic operations on integer constants*) | |
| 296 | Addsimps [integ_of_add RS sym, | |
| 297 | integ_of_minus RS sym, | |
| 298 | integ_of_mult RS sym, | |
| 5512 | 299 | bin_succ_1, bin_succ_0, | 
| 300 | bin_pred_1, bin_pred_0, | |
| 301 | bin_minus_1, bin_minus_0, | |
| 5551 | 302 | bin_add_Pls_right, bin_add_BIT_Min, | 
| 5512 | 303 | bin_add_BIT_0, bin_add_BIT_10, bin_add_BIT_11, | 
| 304 | bin_mult_1, bin_mult_0, | |
| 305 | NCons_Pls_0, NCons_Pls_1, | |
| 306 | NCons_Min_0, NCons_Min_1, | |
| 307 | NCons_BIT]; | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 308 | |
| 5491 | 309 | (*... and simplification of relational operations*) | 
| 5512 | 310 | Addsimps [eq_integ_of_eq, iszero_integ_of_Pls, nonzero_integ_of_Min, | 
| 311 | iszero_integ_of_BIT, | |
| 5540 | 312 | less_integ_of_eq_neg, | 
| 5512 | 313 | not_neg_integ_of_Pls, neg_integ_of_Min, neg_integ_of_BIT, | 
| 5491 | 314 | le_integ_of_eq_not_less]; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 315 | |
| 5491 | 316 | Goalw [zdiff_def] | 
| 317 | "integ_of v - integ_of w = integ_of(bin_add v (bin_minus w))"; | |
| 318 | by (Simp_tac 1); | |
| 319 | qed "diff_integ_of_eq"; | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 320 | |
| 5491 | 321 | (*... and finally subtraction*) | 
| 322 | Addsimps [diff_integ_of_eq]; | |
| 5510 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 323 | |
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 324 | |
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 325 | (** Simplification of inequalities involving numerical constants **) | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 326 | |
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 327 | Goal "(w <= z + #1) = (w<=z | w = z + #1)"; | 
| 5540 | 328 | by (simp_tac (simpset() addsimps [integ_le_less, zless_add1_eq]) 1); | 
| 5510 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 329 | qed "zle_add1_eq"; | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 330 | |
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 331 | Goal "(w <= z - #1) = (w<z)"; | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 332 | by (simp_tac (simpset() addsimps zcompare_rls) 1); | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 333 | qed "zle_diff1_eq"; | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 334 | Addsimps [zle_diff1_eq]; | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 335 | |
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 336 | (*2nd premise can be proved automatically if v is a literal*) | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 337 | Goal "[| w <= z; #0 <= v |] ==> w <= z + v"; | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 338 | by (dtac zadd_zle_mono 1); | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 339 | by (assume_tac 1); | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 340 | by (Full_simp_tac 1); | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 341 | qed "zle_imp_zle_zadd"; | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 342 | |
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 343 | Goal "w <= z ==> w <= z + #1"; | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 344 | by (asm_simp_tac (simpset() addsimps [zle_imp_zle_zadd]) 1); | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 345 | qed "zle_imp_zle_zadd1"; | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 346 | |
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 347 | (*2nd premise can be proved automatically if v is a literal*) | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 348 | Goal "[| w < z; #0 <= v |] ==> w < z + v"; | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 349 | by (dtac zadd_zless_mono 1); | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 350 | by (assume_tac 1); | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 351 | by (Full_simp_tac 1); | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 352 | qed "zless_imp_zless_zadd"; | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 353 | |
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 354 | Goal "w < z ==> w < z + #1"; | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 355 | by (asm_simp_tac (simpset() addsimps [zless_imp_zless_zadd]) 1); | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 356 | qed "zless_imp_zless_zadd1"; | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 357 | |
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 358 | Goal "(w < z + #1) = (w<=z)"; | 
| 5540 | 359 | by (simp_tac (simpset() addsimps [zless_add1_eq, integ_le_less]) 1); | 
| 5510 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 360 | qed "zle_add1_eq_le"; | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 361 | Addsimps [zle_add1_eq_le]; | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 362 | |
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 363 | Goal "(z = z + w) = (w = #0)"; | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 364 | by (rtac trans 1); | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 365 | by (rtac zadd_left_cancel 2); | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 366 | by (simp_tac (simpset() addsimps [eq_sym_conv]) 1); | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 367 | qed "zadd_left_cancel0"; | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 368 | Addsimps [zadd_left_cancel0]; | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 369 | |
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 370 | (*LOOPS as a simprule!*) | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 371 | Goal "[| w + v < z; #0 <= v |] ==> w < z"; | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 372 | by (dtac zadd_zless_mono 1); | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 373 | by (assume_tac 1); | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 374 | by (full_simp_tac (simpset() addsimps zadd_ac) 1); | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 375 | qed "zless_zadd_imp_zless"; | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 376 | |
| 5540 | 377 | (*LOOPS as a simprule! Analogous to Suc_lessD*) | 
| 5510 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 378 | Goal "w + #1 < z ==> w < z"; | 
| 5540 | 379 | by (dtac zless_zadd_imp_zless 1); | 
| 380 | by (assume_tac 2); | |
| 5510 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 381 | by (Simp_tac 1); | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 382 | qed "zless_zadd1_imp_zless"; | 
| 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 383 | |
| 5551 | 384 | Goal "w + #-1 = w - #1"; | 
| 385 | by (simp_tac (simpset() addsimps zadd_ac@zcompare_0_rls) 1); | |
| 386 | qed "zplus_minus1_conv"; | |
| 5510 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 387 | |
| 5551 | 388 | (*Eliminates neg from the subgoal, introduced e.g. by zcompare_0_rls*) | 
| 389 | val no_neg_ss = | |
| 390 | simpset() | |
| 391 | delsimps [less_integ_of_eq_neg] (*loops: it introduces neg*) | |
| 392 | addsimps [zadd_assoc RS sym, zplus_minus1_conv, | |
| 393 | neg_eq_less_0, iszero_def] @ zcompare_rls; | |
| 394 | ||
| 395 | ||
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 396 | (*** nat ***) | 
| 5551 | 397 | |
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 398 | Goal "#0 <= z ==> $# (nat z) = z"; | 
| 5551 | 399 | by (asm_full_simp_tac | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 400 | (simpset() addsimps [neg_eq_less_0, zle_def, not_neg_nat]) 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 401 | qed "nat_0_le"; | 
| 5551 | 402 | |
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 403 | Goal "z < #0 ==> nat z = 0"; | 
| 5551 | 404 | by (asm_full_simp_tac | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 405 | (simpset() addsimps [neg_eq_less_0, zle_def, neg_nat]) 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 406 | qed "nat_less_0"; | 
| 5551 | 407 | |
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 408 | Addsimps [nat_0_le, nat_less_0]; | 
| 5551 | 409 | |
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 410 | Goal "#0 <= w ==> (nat w = m) = (w = $# m)"; | 
| 5551 | 411 | by Auto_tac; | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 412 | qed "nat_eq_iff"; | 
| 5551 | 413 | |
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 414 | Goal "#0 <= w ==> (nat w < m) = (w < $# m)"; | 
| 5551 | 415 | by (rtac iffI 1); | 
| 416 | by (asm_full_simp_tac | |
| 417 | (simpset() delsimps [zless_eq_less] addsimps [zless_eq_less RS sym]) 2); | |
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 418 | by (etac (nat_0_le RS subst) 1); | 
| 5551 | 419 | by (Simp_tac 1); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 420 | qed "nat_less_iff"; | 
| 5551 | 421 | |
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 422 | Goal "#0 <= w ==> (nat w < nat z) = (w<z)"; | 
| 5551 | 423 | by (case_tac "neg z" 1); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 424 | by (auto_tac (claset(), simpset() addsimps [nat_less_iff])); | 
| 5551 | 425 | by (auto_tac (claset() addIs [zless_trans], | 
| 426 | simpset() addsimps [neg_eq_less_0, integ_of_Pls, zle_def])); | |
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 427 | qed "nat_less_eq_zless"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 428 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 429 | Goal "#0 <= w ==> (nat w < nat z) = (w<z)"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 430 | by (stac nat_less_iff 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 431 | ba 1; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 432 | by (case_tac "neg z" 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 433 | by (auto_tac (claset(), simpset() addsimps [not_neg_nat, neg_nat])); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 434 | by (auto_tac (claset(), | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 435 | simpset() addsimps [neg_eq_less_0, integ_of_Pls, zle_def])); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 436 | by (blast_tac (claset() addIs [zless_trans]) 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 437 | qed "nat_less_eq_zless"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 438 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 439 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 440 | (**Can these be generalized without evaluating large numbers?**) | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 441 | Goal "($# k = #0) = (k=0)"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 442 | by (simp_tac (simpset() addsimps [integ_of_Pls]) 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 443 | qed "nat_eq_0_conv"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 444 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 445 | Goal "(#0 = $# k) = (k=0)"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 446 | by (auto_tac (claset(), simpset() addsimps [integ_of_Pls])); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 447 | qed "nat_eq_0_conv'"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 448 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 449 | Addsimps [nat_eq_0_conv, nat_eq_0_conv']; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 450 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5551diff
changeset | 451 |