| author | lcp | 
| Thu, 12 Jan 1995 03:02:34 +0100 | |
| changeset 854 | 2e3ca37dfa14 | 
| parent 782 | 200a16083201 | 
| child 906 | 6cd9c397f36a | 
| permissions | -rw-r--r-- | 
| 515 | 1 | (* Title: ZF/ex/Bin.ML | 
| 0 | 2 | ID: $Id$ | 
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 515 | 4 | Copyright 1994 University of Cambridge | 
| 0 | 5 | |
| 515 | 6 | For Bin.thy. Arithmetic on binary integers. | 
| 0 | 7 | *) | 
| 8 | ||
| 515 | 9 | open Bin; | 
| 0 | 10 | |
| 11 | (*Perform induction on l, then prove the major premise using prems. *) | |
| 12 | fun bin_ind_tac a prems i = | |
| 515 | 13 |     EVERY [res_inst_tac [("x",a)] bin.induct i,
 | 
| 0 | 14 | rename_last_tac a ["1"] (i+3), | 
| 15 | ares_tac prems i]; | |
| 16 | ||
| 515 | 17 | |
| 18 | (** bin_rec -- by Vset recursion **) | |
| 19 | ||
| 20 | goal Bin.thy "bin_rec(Plus,a,b,h) = a"; | |
| 21 | by (rtac (bin_rec_def RS def_Vrec RS trans) 1); | |
| 22 | by (rewrite_goals_tac bin.con_defs); | |
| 23 | by (simp_tac rank_ss 1); | |
| 760 | 24 | qed "bin_rec_Plus"; | 
| 515 | 25 | |
| 26 | goal Bin.thy "bin_rec(Minus,a,b,h) = b"; | |
| 27 | by (rtac (bin_rec_def RS def_Vrec RS trans) 1); | |
| 28 | by (rewrite_goals_tac bin.con_defs); | |
| 29 | by (simp_tac rank_ss 1); | |
| 760 | 30 | qed "bin_rec_Minus"; | 
| 515 | 31 | |
| 32 | goal Bin.thy "bin_rec(w$$x,a,b,h) = h(w, x, bin_rec(w,a,b,h))"; | |
| 33 | by (rtac (bin_rec_def RS def_Vrec RS trans) 1); | |
| 34 | by (rewrite_goals_tac bin.con_defs); | |
| 35 | by (simp_tac rank_ss 1); | |
| 760 | 36 | qed "bin_rec_Bcons"; | 
| 515 | 37 | |
| 38 | (*Type checking*) | |
| 39 | val prems = goal Bin.thy | |
| 40 | "[| w: bin; \ | |
| 41 | \ a: C(Plus); b: C(Minus); \ | |
| 42 | \ !!w x r. [| w: bin; x: bool; r: C(w) |] ==> h(w,x,r): C(w$$x) \ | |
| 43 | \ |] ==> bin_rec(w,a,b,h) : C(w)"; | |
| 44 | by (bin_ind_tac "w" prems 1); | |
| 45 | by (ALLGOALS | |
| 46 | (asm_simp_tac (ZF_ss addsimps (prems@[bin_rec_Plus,bin_rec_Minus, | |
| 47 | bin_rec_Bcons])))); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 48 | qed "bin_rec_type"; | 
| 515 | 49 | |
| 50 | (** Versions for use with definitions **) | |
| 51 | ||
| 52 | val [rew] = goal Bin.thy | |
| 53 | "[| !!w. j(w)==bin_rec(w,a,b,h) |] ==> j(Plus) = a"; | |
| 54 | by (rewtac rew); | |
| 55 | by (rtac bin_rec_Plus 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 56 | qed "def_bin_rec_Plus"; | 
| 515 | 57 | |
| 58 | val [rew] = goal Bin.thy | |
| 59 | "[| !!w. j(w)==bin_rec(w,a,b,h) |] ==> j(Minus) = b"; | |
| 60 | by (rewtac rew); | |
| 61 | by (rtac bin_rec_Minus 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 62 | qed "def_bin_rec_Minus"; | 
| 515 | 63 | |
| 64 | val [rew] = goal Bin.thy | |
| 65 | "[| !!w. j(w)==bin_rec(w,a,b,h) |] ==> j(w$$x) = h(w,x,j(w))"; | |
| 66 | by (rewtac rew); | |
| 67 | by (rtac bin_rec_Bcons 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 68 | qed "def_bin_rec_Bcons"; | 
| 515 | 69 | |
| 70 | fun bin_recs def = map standard | |
| 71 | ([def] RL [def_bin_rec_Plus, def_bin_rec_Minus, def_bin_rec_Bcons]); | |
| 72 | ||
| 73 | (** Type checking **) | |
| 74 | ||
| 75 | val bin_typechecks0 = bin_rec_type :: bin.intrs; | |
| 76 | ||
| 77 | goalw Bin.thy [integ_of_bin_def] | |
| 78 | "!!w. w: bin ==> integ_of_bin(w) : integ"; | |
| 79 | by (typechk_tac (bin_typechecks0@integ_typechecks@ | |
| 80 | nat_typechecks@[bool_into_nat])); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 81 | qed "integ_of_bin_type"; | 
| 515 | 82 | |
| 83 | goalw Bin.thy [bin_succ_def] | |
| 84 | "!!w. w: bin ==> bin_succ(w) : bin"; | |
| 85 | by (typechk_tac (bin_typechecks0@bool_typechecks)); | |
| 760 | 86 | qed "bin_succ_type"; | 
| 515 | 87 | |
| 88 | goalw Bin.thy [bin_pred_def] | |
| 89 | "!!w. w: bin ==> bin_pred(w) : bin"; | |
| 90 | by (typechk_tac (bin_typechecks0@bool_typechecks)); | |
| 760 | 91 | qed "bin_pred_type"; | 
| 515 | 92 | |
| 93 | goalw Bin.thy [bin_minus_def] | |
| 94 | "!!w. w: bin ==> bin_minus(w) : bin"; | |
| 95 | by (typechk_tac ([bin_pred_type]@bin_typechecks0@bool_typechecks)); | |
| 760 | 96 | qed "bin_minus_type"; | 
| 515 | 97 | |
| 98 | goalw Bin.thy [bin_add_def] | |
| 99 | "!!v w. [| v: bin; w: bin |] ==> bin_add(v,w) : bin"; | |
| 100 | by (typechk_tac ([bin_succ_type,bin_pred_type]@bin_typechecks0@ | |
| 101 | bool_typechecks@ZF_typechecks)); | |
| 760 | 102 | qed "bin_add_type"; | 
| 515 | 103 | |
| 104 | goalw Bin.thy [bin_mult_def] | |
| 105 | "!!v w. [| v: bin; w: bin |] ==> bin_mult(v,w) : bin"; | |
| 106 | by (typechk_tac ([bin_minus_type,bin_add_type]@bin_typechecks0@ | |
| 107 | bool_typechecks)); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 108 | qed "bin_mult_type"; | 
| 515 | 109 | |
| 110 | val bin_typechecks = bin_typechecks0 @ | |
| 111 | [integ_of_bin_type, bin_succ_type, bin_pred_type, | |
| 112 | bin_minus_type, bin_add_type, bin_mult_type]; | |
| 113 | ||
| 114 | val bin_ss = integ_ss | |
| 115 | addsimps([bool_1I, bool_0I, | |
| 116 | bin_rec_Plus, bin_rec_Minus, bin_rec_Bcons] @ | |
| 117 | bin_recs integ_of_bin_def @ bool_simps @ bin_typechecks); | |
| 118 | ||
| 119 | val typechecks = bin_typechecks @ integ_typechecks @ nat_typechecks @ | |
| 120 | [bool_subset_nat RS subsetD]; | |
| 121 | ||
| 122 | (**** The carry/borrow functions, bin_succ and bin_pred ****) | |
| 123 | ||
| 124 | (** Lemmas **) | |
| 125 | ||
| 126 | goal Integ.thy | |
| 127 | "!!z v. [| z $+ v = z' $+ v'; \ | |
| 128 | \ z: integ; z': integ; v: integ; v': integ; w: integ |] \ | |
| 129 | \ ==> z $+ (v $+ w) = z' $+ (v' $+ w)"; | |
| 130 | by (asm_simp_tac (integ_ss addsimps ([zadd_assoc RS sym])) 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 131 | qed "zadd_assoc_cong"; | 
| 515 | 132 | |
| 133 | goal Integ.thy | |
| 134 | "!!z v w. [| z: integ; v: integ; w: integ |] \ | |
| 135 | \ ==> z $+ (v $+ w) = v $+ (z $+ w)"; | |
| 136 | by (REPEAT (ares_tac [zadd_commute RS zadd_assoc_cong] 1)); | |
| 760 | 137 | qed "zadd_assoc_swap"; | 
| 515 | 138 | |
| 139 | val zadd_cong = | |
| 140 |     read_instantiate_sg (sign_of Integ.thy) [("t","op $+")] subst_context2;
 | |
| 141 | ||
| 142 | val zadd_kill = (refl RS zadd_cong); | |
| 143 | val zadd_assoc_swap_kill = zadd_kill RSN (4, zadd_assoc_swap RS trans); | |
| 144 | ||
| 145 | (*Pushes 'constants' of the form $#m to the right -- LOOPS if two!*) | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 146 | bind_thm ("zadd_assoc_znat", (znat_type RS zadd_assoc_swap));
 | 
| 515 | 147 | |
| 148 | goal Integ.thy | |
| 149 | "!!z w. [| z: integ; w: integ |] \ | |
| 150 | \ ==> w $+ (z $+ (w $+ z)) = w $+ (w $+ (z $+ z))"; | |
| 151 | by (REPEAT (ares_tac [zadd_kill, zadd_assoc_swap] 1)); | |
| 760 | 152 | qed "zadd_swap_pairs"; | 
| 515 | 153 | |
| 154 | ||
| 155 | val carry_ss = bin_ss addsimps | |
| 156 | (bin_recs bin_succ_def @ bin_recs bin_pred_def); | |
| 157 | ||
| 158 | goal Bin.thy | |
| 159 | "!!w. w: bin ==> integ_of_bin(bin_succ(w)) = $#1 $+ integ_of_bin(w)"; | |
| 160 | by (etac bin.induct 1); | |
| 161 | by (simp_tac (carry_ss addsimps [zadd_0_right]) 1); | |
| 162 | by (simp_tac (carry_ss addsimps [zadd_zminus_inverse]) 1); | |
| 163 | by (etac boolE 1); | |
| 164 | by (ALLGOALS (asm_simp_tac (carry_ss addsimps [zadd_assoc]))); | |
| 165 | by (REPEAT (ares_tac (zadd_swap_pairs::typechecks) 1)); | |
| 760 | 166 | qed "integ_of_bin_succ"; | 
| 515 | 167 | |
| 168 | goal Bin.thy | |
| 169 | "!!w. w: bin ==> integ_of_bin(bin_pred(w)) = $~ ($#1) $+ integ_of_bin(w)"; | |
| 170 | by (etac bin.induct 1); | |
| 171 | by (simp_tac (carry_ss addsimps [zadd_0_right]) 1); | |
| 172 | by (simp_tac (carry_ss addsimps [zadd_zminus_inverse]) 1); | |
| 173 | by (etac boolE 1); | |
| 174 | by (ALLGOALS | |
| 175 | (asm_simp_tac | |
| 176 | (carry_ss addsimps [zadd_assoc RS sym, | |
| 177 | zadd_zminus_inverse, zadd_zminus_inverse2]))); | |
| 178 | by (REPEAT (ares_tac ([zadd_commute, zadd_cong, refl]@typechecks) 1)); | |
| 760 | 179 | qed "integ_of_bin_pred"; | 
| 515 | 180 | |
| 181 | (*These two results replace the definitions of bin_succ and bin_pred*) | |
| 182 | ||
| 183 | ||
| 184 | (*** bin_minus: (unary!) negation of binary integers ***) | |
| 185 | ||
| 186 | val bin_minus_ss = | |
| 187 | bin_ss addsimps (bin_recs bin_minus_def @ | |
| 188 | [integ_of_bin_succ, integ_of_bin_pred]); | |
| 189 | ||
| 190 | goal Bin.thy | |
| 191 | "!!w. w: bin ==> integ_of_bin(bin_minus(w)) = $~ integ_of_bin(w)"; | |
| 192 | by (etac bin.induct 1); | |
| 193 | by (simp_tac (bin_minus_ss addsimps [zminus_0]) 1); | |
| 194 | by (simp_tac (bin_minus_ss addsimps [zadd_0_right]) 1); | |
| 195 | by (etac boolE 1); | |
| 196 | by (ALLGOALS | |
| 197 | (asm_simp_tac (bin_minus_ss addsimps [zminus_zadd_distrib, zadd_assoc]))); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 198 | qed "integ_of_bin_minus"; | 
| 515 | 199 | |
| 200 | ||
| 201 | (*** bin_add: binary addition ***) | |
| 202 | ||
| 203 | goalw Bin.thy [bin_add_def] "!!w. w: bin ==> bin_add(Plus,w) = w"; | |
| 204 | by (asm_simp_tac bin_ss 1); | |
| 760 | 205 | qed "bin_add_Plus"; | 
| 515 | 206 | |
| 207 | goalw Bin.thy [bin_add_def] "!!w. w: bin ==> bin_add(Minus,w) = bin_pred(w)"; | |
| 208 | by (asm_simp_tac bin_ss 1); | |
| 760 | 209 | qed "bin_add_Minus"; | 
| 515 | 210 | |
| 211 | goalw Bin.thy [bin_add_def] "bin_add(v$$x,Plus) = v$$x"; | |
| 212 | by (simp_tac bin_ss 1); | |
| 760 | 213 | qed "bin_add_Bcons_Plus"; | 
| 515 | 214 | |
| 215 | goalw Bin.thy [bin_add_def] "bin_add(v$$x,Minus) = bin_pred(v$$x)"; | |
| 216 | by (simp_tac bin_ss 1); | |
| 760 | 217 | qed "bin_add_Bcons_Minus"; | 
| 515 | 218 | |
| 219 | goalw Bin.thy [bin_add_def] | |
| 220 | "!!w y. [| w: bin; y: bool |] ==> \ | |
| 221 | \ bin_add(v$$x, w$$y) = \ | |
| 222 | \ bin_add(v, cond(x and y, bin_succ(w), w)) $$ (x xor y)"; | |
| 223 | by (asm_simp_tac bin_ss 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 224 | qed "bin_add_Bcons_Bcons"; | 
| 515 | 225 | |
| 226 | val bin_add_rews = [bin_add_Plus, bin_add_Minus, bin_add_Bcons_Plus, | |
| 227 | bin_add_Bcons_Minus, bin_add_Bcons_Bcons, | |
| 228 | integ_of_bin_succ, integ_of_bin_pred]; | |
| 229 | ||
| 230 | val bin_add_ss = bin_ss addsimps ([bool_subset_nat RS subsetD] @ bin_add_rews); | |
| 231 | ||
| 232 | goal Bin.thy | |
| 233 | "!!v. v: bin ==> \ | |
| 234 | \ ALL w: bin. integ_of_bin(bin_add(v,w)) = \ | |
| 235 | \ integ_of_bin(v) $+ integ_of_bin(w)"; | |
| 236 | by (etac bin.induct 1); | |
| 237 | by (simp_tac bin_add_ss 1); | |
| 238 | by (simp_tac bin_add_ss 1); | |
| 239 | by (rtac ballI 1); | |
| 240 | by (bin_ind_tac "wa" [] 1); | |
| 241 | by (asm_simp_tac (bin_add_ss addsimps [zadd_0_right]) 1); | |
| 242 | by (asm_simp_tac bin_add_ss 1); | |
| 243 | by (REPEAT (ares_tac (zadd_commute::typechecks) 1)); | |
| 244 | by (etac boolE 1); | |
| 245 | by (asm_simp_tac (bin_add_ss addsimps [zadd_assoc, zadd_swap_pairs]) 2); | |
| 246 | by (REPEAT (ares_tac ([refl, zadd_kill, zadd_assoc_swap_kill]@typechecks) 2)); | |
| 247 | by (etac boolE 1); | |
| 248 | by (ALLGOALS (asm_simp_tac (bin_add_ss addsimps [zadd_assoc,zadd_swap_pairs]))); | |
| 249 | by (REPEAT (ares_tac ([refl, zadd_kill, zadd_assoc_swap_kill RS sym]@ | |
| 250 | typechecks) 1)); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 251 | qed "integ_of_bin_add_lemma"; | 
| 515 | 252 | |
| 253 | val integ_of_bin_add = integ_of_bin_add_lemma RS bspec; | |
| 254 | ||
| 255 | ||
| 256 | (*** bin_add: binary multiplication ***) | |
| 257 | ||
| 258 | val bin_mult_ss = | |
| 259 | bin_ss addsimps (bin_recs bin_mult_def @ | |
| 260 | [integ_of_bin_minus, integ_of_bin_add]); | |
| 261 | ||
| 262 | ||
| 263 | val major::prems = goal Bin.thy | |
| 264 | "[| v: bin; w: bin |] ==> \ | |
| 265 | \ integ_of_bin(bin_mult(v,w)) = \ | |
| 266 | \ integ_of_bin(v) $* integ_of_bin(w)"; | |
| 267 | by (cut_facts_tac prems 1); | |
| 268 | by (bin_ind_tac "v" [major] 1); | |
| 269 | by (asm_simp_tac (bin_mult_ss addsimps [zmult_0]) 1); | |
| 270 | by (asm_simp_tac (bin_mult_ss addsimps [zmult_1,zmult_zminus]) 1); | |
| 271 | by (etac boolE 1); | |
| 272 | by (asm_simp_tac (bin_mult_ss addsimps [zadd_zmult_distrib]) 2); | |
| 273 | by (asm_simp_tac | |
| 274 | (bin_mult_ss addsimps [zadd_zmult_distrib, zmult_1, zadd_assoc]) 1); | |
| 275 | by (REPEAT (ares_tac ([zadd_commute, zadd_assoc_swap_kill RS sym]@ | |
| 276 | typechecks) 1)); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 277 | qed "integ_of_bin_mult"; | 
| 515 | 278 | |
| 279 | (**** Computations ****) | |
| 280 | ||
| 281 | (** extra rules for bin_succ, bin_pred **) | |
| 282 | ||
| 283 | val [bin_succ_Plus, bin_succ_Minus, _] = bin_recs bin_succ_def; | |
| 284 | val [bin_pred_Plus, bin_pred_Minus, _] = bin_recs bin_pred_def; | |
| 285 | ||
| 286 | goal Bin.thy "bin_succ(w$$1) = bin_succ(w) $$ 0"; | |
| 287 | by (simp_tac carry_ss 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 288 | qed "bin_succ_Bcons1"; | 
| 515 | 289 | |
| 290 | goal Bin.thy "bin_succ(w$$0) = w$$1"; | |
| 291 | by (simp_tac carry_ss 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 292 | qed "bin_succ_Bcons0"; | 
| 515 | 293 | |
| 294 | goal Bin.thy "bin_pred(w$$1) = w$$0"; | |
| 295 | by (simp_tac carry_ss 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 296 | qed "bin_pred_Bcons1"; | 
| 515 | 297 | |
| 298 | goal Bin.thy "bin_pred(w$$0) = bin_pred(w) $$ 1"; | |
| 299 | by (simp_tac carry_ss 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 300 | qed "bin_pred_Bcons0"; | 
| 515 | 301 | |
| 302 | (** extra rules for bin_minus **) | |
| 303 | ||
| 304 | val [bin_minus_Plus, bin_minus_Minus, _] = bin_recs bin_minus_def; | |
| 305 | ||
| 306 | goal Bin.thy "bin_minus(w$$1) = bin_pred(bin_minus(w) $$ 0)"; | |
| 307 | by (simp_tac bin_minus_ss 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 308 | qed "bin_minus_Bcons1"; | 
| 515 | 309 | |
| 310 | goal Bin.thy "bin_minus(w$$0) = bin_minus(w) $$ 0"; | |
| 311 | by (simp_tac bin_minus_ss 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 312 | qed "bin_minus_Bcons0"; | 
| 515 | 313 | |
| 314 | (** extra rules for bin_add **) | |
| 315 | ||
| 316 | goal Bin.thy | |
| 317 | "!!w. w: bin ==> bin_add(v$$1, w$$1) = bin_add(v, bin_succ(w)) $$ 0"; | |
| 318 | by (asm_simp_tac bin_add_ss 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 319 | qed "bin_add_Bcons_Bcons11"; | 
| 515 | 320 | |
| 321 | goal Bin.thy | |
| 322 | "!!w. w: bin ==> bin_add(v$$1, w$$0) = bin_add(v,w) $$ 1"; | |
| 323 | by (asm_simp_tac bin_add_ss 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 324 | qed "bin_add_Bcons_Bcons10"; | 
| 515 | 325 | |
| 326 | goal Bin.thy | |
| 327 | "!!w y.[| w: bin; y: bool |] ==> bin_add(v$$0, w$$y) = bin_add(v,w) $$ y"; | |
| 328 | by (asm_simp_tac bin_add_ss 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 329 | qed "bin_add_Bcons_Bcons0"; | 
| 515 | 330 | |
| 331 | (** extra rules for bin_mult **) | |
| 332 | ||
| 333 | val [bin_mult_Plus, bin_mult_Minus, _] = bin_recs bin_mult_def; | |
| 334 | ||
| 335 | goal Bin.thy "bin_mult(v$$1, w) = bin_add(bin_mult(v,w)$$0, w)"; | |
| 336 | by (simp_tac bin_mult_ss 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 337 | qed "bin_mult_Bcons1"; | 
| 515 | 338 | |
| 339 | goal Bin.thy "bin_mult(v$$0, w) = bin_mult(v,w)$$0"; | |
| 340 | by (simp_tac bin_mult_ss 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 341 | qed "bin_mult_Bcons0"; | 
| 515 | 342 | |
| 343 | ||
| 344 | (*** The computation simpset ***) | |
| 345 | ||
| 346 | val bin_comp_ss = integ_ss | |
| 347 | addsimps [bin_succ_Plus, bin_succ_Minus, | |
| 348 | bin_succ_Bcons1, bin_succ_Bcons0, | |
| 349 | bin_pred_Plus, bin_pred_Minus, | |
| 350 | bin_pred_Bcons1, bin_pred_Bcons0, | |
| 351 | bin_minus_Plus, bin_minus_Minus, | |
| 352 | bin_minus_Bcons1, bin_minus_Bcons0, | |
| 353 | bin_add_Plus, bin_add_Minus, bin_add_Bcons_Plus, | |
| 354 | bin_add_Bcons_Minus, bin_add_Bcons_Bcons0, | |
| 355 | bin_add_Bcons_Bcons10, bin_add_Bcons_Bcons11, | |
| 356 | bin_mult_Plus, bin_mult_Minus, | |
| 357 | bin_mult_Bcons1, bin_mult_Bcons0] | |
| 358 | setsolver (type_auto_tac ([bool_1I, bool_0I] @ bin_typechecks0)); | |
| 359 | ||
| 360 | (*** Examples of performing binary arithmetic by simplification ***) | |
| 361 | ||
| 362 | proof_timing := true; | |
| 363 | (*All runtimes below are on a SPARCserver 10*) | |
| 364 | ||
| 365 | (* 13+19 = 32 *) | |
| 366 | goal Bin.thy | |
| 367 | "bin_add(Plus$$1$$1$$0$$1, Plus$$1$$0$$0$$1$$1) = Plus$$1$$0$$0$$0$$0$$0"; | |
| 368 | by (simp_tac bin_comp_ss 1); (*0.6 secs*) | |
| 369 | result(); | |
| 370 | ||
| 371 | bin_add(binary_of_int 13, binary_of_int 19); | |
| 372 | ||
| 373 | (* 1234+5678 = 6912 *) | |
| 374 | goal Bin.thy | |
| 375 | "bin_add(Plus$$1$$0$$0$$1$$1$$0$$1$$0$$0$$1$$0, \ | |
| 376 | \ Plus$$1$$0$$1$$1$$0$$0$$0$$1$$0$$1$$1$$1$$0) = \ | |
| 377 | \ Plus$$1$$1$$0$$1$$1$$0$$0$$0$$0$$0$$0$$0$$0"; | |
| 378 | by (simp_tac bin_comp_ss 1); (*2.6 secs*) | |
| 379 | result(); | |
| 380 | ||
| 381 | bin_add(binary_of_int 1234, binary_of_int 5678); | |
| 382 | ||
| 383 | (* 1359-2468 = ~1109 *) | |
| 384 | goal Bin.thy | |
| 385 | "bin_add(Plus$$1$$0$$1$$0$$1$$0$$0$$1$$1$$1$$1, \ | |
| 386 | \ Minus$$0$$1$$1$$0$$0$$1$$0$$1$$1$$1$$0$$0) = \ | |
| 387 | \ Minus$$1$$0$$1$$1$$1$$0$$1$$0$$1$$0$$1$$1"; | |
| 388 | by (simp_tac bin_comp_ss 1); (*2.3 secs*) | |
| 389 | result(); | |
| 390 | ||
| 391 | bin_add(binary_of_int 1359, binary_of_int ~2468); | |
| 392 | ||
| 393 | (* 93746-46375 = 47371 *) | |
| 394 | goal Bin.thy | |
| 395 | "bin_add(Plus$$1$$0$$1$$1$$0$$1$$1$$1$$0$$0$$0$$1$$1$$0$$0$$1$$0, \ | |
| 396 | \ Minus$$0$$1$$0$$0$$1$$0$$1$$0$$1$$1$$0$$1$$1$$0$$0$$1) = \ | |
| 397 | \ Plus$$0$$1$$0$$1$$1$$1$$0$$0$$1$$0$$0$$0$$0$$1$$0$$1$$1"; | |
| 398 | by (simp_tac bin_comp_ss 1); (*3.9 secs*) | |
| 399 | result(); | |
| 400 | ||
| 401 | bin_add(binary_of_int 93746, binary_of_int ~46375); | |
| 402 | ||
| 403 | (* negation of 65745 *) | |
| 404 | goal Bin.thy | |
| 405 | "bin_minus(Plus$$1$$0$$0$$0$$0$$0$$0$$0$$0$$1$$1$$0$$1$$0$$0$$0$$1) = \ | |
| 406 | \ Minus$$0$$1$$1$$1$$1$$1$$1$$1$$1$$0$$0$$1$$0$$1$$1$$1$$1"; | |
| 407 | by (simp_tac bin_comp_ss 1); (*0.6 secs*) | |
| 408 | result(); | |
| 409 | ||
| 410 | bin_minus(binary_of_int 65745); | |
| 411 | ||
| 412 | (* negation of ~54321 *) | |
| 413 | goal Bin.thy | |
| 414 | "bin_minus(Minus$$0$$0$$1$$0$$1$$0$$1$$1$$1$$1$$0$$0$$1$$1$$1$$1) = \ | |
| 415 | \ Plus$$0$$1$$1$$0$$1$$0$$1$$0$$0$$0$$0$$1$$1$$0$$0$$0$$1"; | |
| 416 | by (simp_tac bin_comp_ss 1); (*0.7 secs*) | |
| 417 | result(); | |
| 418 | ||
| 419 | bin_minus(binary_of_int ~54321); | |
| 420 | ||
| 421 | (* 13*19 = 247 *) | |
| 422 | goal Bin.thy "bin_mult(Plus$$1$$1$$0$$1, Plus$$1$$0$$0$$1$$1) = \ | |
| 423 | \ Plus$$1$$1$$1$$1$$0$$1$$1$$1"; | |
| 424 | by (simp_tac bin_comp_ss 1); (*1.5 secs*) | |
| 425 | result(); | |
| 426 | ||
| 427 | bin_mult(binary_of_int 13, binary_of_int 19); | |
| 428 | ||
| 429 | (* ~84 * 51 = ~4284 *) | |
| 430 | goal Bin.thy | |
| 431 | "bin_mult(Minus$$0$$1$$0$$1$$1$$0$$0, Plus$$1$$1$$0$$0$$1$$1) = \ | |
| 432 | \ Minus$$0$$1$$1$$1$$1$$0$$1$$0$$0$$0$$1$$0$$0"; | |
| 433 | by (simp_tac bin_comp_ss 1); (*2.6 secs*) | |
| 434 | result(); | |
| 435 | ||
| 436 | bin_mult(binary_of_int ~84, binary_of_int 51); | |
| 437 | ||
| 438 | (* 255*255 = 65025; the worst case for 8-bit operands *) | |
| 439 | goal Bin.thy | |
| 440 | "bin_mult(Plus$$1$$1$$1$$1$$1$$1$$1$$1, \ | |
| 441 | \ Plus$$1$$1$$1$$1$$1$$1$$1$$1) = \ | |
| 442 | \ Plus$$1$$1$$1$$1$$1$$1$$1$$0$$0$$0$$0$$0$$0$$0$$0$$1"; | |
| 443 | by (simp_tac bin_comp_ss 1); (*9.8 secs*) | |
| 444 | result(); | |
| 445 | ||
| 446 | bin_mult(binary_of_int 255, binary_of_int 255); | |
| 447 | ||
| 448 | (* 1359 * ~2468 = ~3354012 *) | |
| 449 | goal Bin.thy | |
| 450 | "bin_mult(Plus$$1$$0$$1$$0$$1$$0$$0$$1$$1$$1$$1, \ | |
| 451 | \ Minus$$0$$1$$1$$0$$0$$1$$0$$1$$1$$1$$0$$0) = \ | |
| 452 | \ Minus$$0$$0$$1$$1$$0$$0$$1$$1$$0$$1$$0$$0$$1$$0$$0$$1$$1$$0$$0$$1$$0$$0"; | |
| 453 | by (simp_tac bin_comp_ss 1); (*13.7 secs*) | |
| 454 | result(); | |
| 455 | ||
| 456 | bin_mult(binary_of_int 1359, binary_of_int ~2468); |