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