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