| author | berghofe | 
| Fri, 23 Jun 2000 12:24:37 +0200 | |
| changeset 9114 | de99e37effda | 
| parent 9108 | 9fff97d29837 | 
| child 9214 | 9454f30eacc7 | 
| 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 | 
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7033diff
changeset | 2 | ID: $Id$ | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 3 | Authors: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 4 | David Spelt, University of Twente | 
| 6060 | 5 | Tobias Nipkow, Technical University Munich | 
| 1632 | 6 | Copyright 1994 University of Cambridge | 
| 6060 | 7 | Copyright 1996 University of Twente | 
| 8 | Copyright 1999 TU Munich | |
| 1632 | 9 | |
| 7708 | 10 | Arithmetic on binary integers. | 
| 1632 | 11 | *) | 
| 12 | ||
| 13 | (** extra rules for bin_succ, bin_pred, bin_add, bin_mult **) | |
| 14 | ||
| 7008 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 15 | Goal "NCons Pls False = Pls"; | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 16 | by (Simp_tac 1); | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 17 | qed "NCons_Pls_0"; | 
| 1632 | 18 | |
| 7008 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 19 | Goal "NCons Pls True = Pls BIT True"; | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 20 | by (Simp_tac 1); | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 21 | qed "NCons_Pls_1"; | 
| 1632 | 22 | |
| 7008 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 23 | Goal "NCons Min False = Min BIT False"; | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 24 | by (Simp_tac 1); | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 25 | qed "NCons_Min_0"; | 
| 1632 | 26 | |
| 7008 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 27 | Goal "NCons Min True = Min"; | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 28 | by (Simp_tac 1); | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 29 | qed "NCons_Min_1"; | 
| 1632 | 30 | |
| 7008 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 31 | Goal "bin_succ(w BIT True) = (bin_succ w) BIT False"; | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 32 | by (Simp_tac 1); | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 33 | qed "bin_succ_1"; | 
| 1632 | 34 | |
| 7008 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 35 | Goal "bin_succ(w BIT False) = NCons w True"; | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 36 | by (Simp_tac 1); | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 37 | qed "bin_succ_0"; | 
| 1632 | 38 | |
| 7008 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 39 | Goal "bin_pred(w BIT True) = NCons w False"; | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 40 | by (Simp_tac 1); | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 41 | qed "bin_pred_1"; | 
| 1632 | 42 | |
| 7008 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 43 | Goal "bin_pred(w BIT False) = (bin_pred w) BIT True"; | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 44 | by (Simp_tac 1); | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 45 | qed "bin_pred_0"; | 
| 1632 | 46 | |
| 7008 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 47 | Goal "bin_minus(w BIT True) = bin_pred (NCons (bin_minus w) False)"; | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 48 | by (Simp_tac 1); | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 49 | qed "bin_minus_1"; | 
| 1632 | 50 | |
| 7008 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 51 | Goal "bin_minus(w BIT False) = (bin_minus w) BIT False"; | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 52 | by (Simp_tac 1); | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 53 | qed "bin_minus_0"; | 
| 1632 | 54 | |
| 5491 | 55 | |
| 1632 | 56 | (*** bin_add: binary addition ***) | 
| 57 | ||
| 7008 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 58 | Goal "bin_add (v BIT True) (w BIT True) = \ | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 59 | \ NCons (bin_add v (bin_succ w)) False"; | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 60 | by (Simp_tac 1); | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 61 | qed "bin_add_BIT_11"; | 
| 1632 | 62 | |
| 7008 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 63 | Goal "bin_add (v BIT True) (w BIT False) = NCons (bin_add v w) True"; | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 64 | by (Simp_tac 1); | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 65 | qed "bin_add_BIT_10"; | 
| 1632 | 66 | |
| 7008 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 67 | Goal "bin_add (v BIT False) (w BIT y) = NCons (bin_add v w) y"; | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 68 | by Auto_tac; | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 69 | qed "bin_add_BIT_0"; | 
| 1632 | 70 | |
| 5551 | 71 | Goal "bin_add w Pls = w"; | 
| 72 | by (induct_tac "w" 1); | |
| 73 | by Auto_tac; | |
| 74 | qed "bin_add_Pls_right"; | |
| 1632 | 75 | |
| 7517 
bad2f36810e1
generalized the theorem bin_add_BIT_Min to bin_add_Min_right
 paulson parents: 
7074diff
changeset | 76 | Goal "bin_add w Min = bin_pred w"; | 
| 
bad2f36810e1
generalized the theorem bin_add_BIT_Min to bin_add_Min_right
 paulson parents: 
7074diff
changeset | 77 | by (induct_tac "w" 1); | 
| 
bad2f36810e1
generalized the theorem bin_add_BIT_Min to bin_add_Min_right
 paulson parents: 
7074diff
changeset | 78 | by Auto_tac; | 
| 
bad2f36810e1
generalized the theorem bin_add_BIT_Min to bin_add_Min_right
 paulson parents: 
7074diff
changeset | 79 | qed "bin_add_Min_right"; | 
| 1632 | 80 | |
| 7008 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 81 | Goal "bin_add (v BIT x) (w BIT y) = \ | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 82 | \ NCons(bin_add v (if x & y then (bin_succ w) else w)) (x~= y)"; | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 83 | by (Simp_tac 1); | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 84 | qed "bin_add_BIT_BIT"; | 
| 1632 | 85 | |
| 86 | ||
| 6036 | 87 | (*** bin_mult: binary multiplication ***) | 
| 1632 | 88 | |
| 7008 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 89 | Goal "bin_mult (v BIT True) w = bin_add (NCons (bin_mult v w) False) w"; | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 90 | by (Simp_tac 1); | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 91 | qed "bin_mult_1"; | 
| 1632 | 92 | |
| 7008 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 93 | Goal "bin_mult (v BIT False) w = NCons (bin_mult v w) False"; | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 94 | by (Simp_tac 1); | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 95 | qed "bin_mult_0"; | 
| 1632 | 96 | |
| 97 | ||
| 98 | (**** The carry/borrow functions, bin_succ and bin_pred ****) | |
| 99 | ||
| 100 | ||
| 6910 | 101 | (**** number_of ****) | 
| 1632 | 102 | |
| 7008 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 103 | Goal "number_of(NCons w b) = (number_of(w BIT b)::int)"; | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 104 | by (induct_tac "w" 1); | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 105 | by (ALLGOALS Asm_simp_tac); | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 106 | qed "number_of_NCons"; | 
| 1632 | 107 | |
| 6910 | 108 | Addsimps [number_of_NCons]; | 
| 1632 | 109 | |
| 7008 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 110 | Goal "number_of(bin_succ w) = int 1 + number_of w"; | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 111 | by (induct_tac "w" 1); | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 112 | by (ALLGOALS (asm_simp_tac (simpset() addsimps zadd_ac))); | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 113 | qed "number_of_succ"; | 
| 5491 | 114 | |
| 7008 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 115 | Goal "number_of(bin_pred w) = - (int 1) + number_of w"; | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 116 | by (induct_tac "w" 1); | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 117 | by (ALLGOALS (asm_simp_tac (simpset() addsimps zadd_ac))); | 
| 
6def5ce146e2
qed_goal -> Goal;  new theorems nat_le_0, nat_le_eq_zle and zdiff_int
 paulson parents: 
6997diff
changeset | 118 | qed "number_of_pred"; | 
| 1632 | 119 | |
| 6910 | 120 | Goal "number_of(bin_minus w) = (- (number_of w)::int)"; | 
| 121 | by (induct_tac "w" 1); | |
| 5491 | 122 | by (Simp_tac 1); | 
| 123 | by (Simp_tac 1); | |
| 124 | by (asm_simp_tac (simpset() | |
| 5551 | 125 | delsimps [bin_pred_Pls, bin_pred_Min, bin_pred_BIT] | 
| 6910 | 126 | addsimps [number_of_succ,number_of_pred, | 
| 5491 | 127 | zadd_assoc]) 1); | 
| 6910 | 128 | qed "number_of_minus"; | 
| 1632 | 129 | |
| 130 | ||
| 9108 | 131 | bind_thms ("bin_add_simps", [bin_add_BIT_BIT, number_of_succ, number_of_pred]);
 | 
| 1632 | 132 | |
| 6036 | 133 | (*This proof is complicated by the mutual recursion*) | 
| 6910 | 134 | Goal "! w. number_of(bin_add v w) = (number_of v + number_of w::int)"; | 
| 5184 | 135 | by (induct_tac "v" 1); | 
| 4686 | 136 | by (simp_tac (simpset() addsimps bin_add_simps) 1); | 
| 137 | by (simp_tac (simpset() addsimps bin_add_simps) 1); | |
| 1632 | 138 | by (rtac allI 1); | 
| 5184 | 139 | by (induct_tac "w" 1); | 
| 5540 | 140 | by (ALLGOALS (asm_simp_tac (simpset() addsimps bin_add_simps @ zadd_ac))); | 
| 6910 | 141 | qed_spec_mp "number_of_add"; | 
| 1632 | 142 | |
| 5779 
5c74f003a68e
Explicit (and improved) simprules for binary arithmetic.
 paulson parents: 
5747diff
changeset | 143 | |
| 
5c74f003a68e
Explicit (and improved) simprules for binary arithmetic.
 paulson parents: 
5747diff
changeset | 144 | (*Subtraction*) | 
| 
5c74f003a68e
Explicit (and improved) simprules for binary arithmetic.
 paulson parents: 
5747diff
changeset | 145 | Goalw [zdiff_def] | 
| 6910 | 146 | "number_of v - number_of w = (number_of(bin_add v (bin_minus w))::int)"; | 
| 147 | by (simp_tac (simpset() addsimps [number_of_add, number_of_minus]) 1); | |
| 148 | qed "diff_number_of_eq"; | |
| 5779 
5c74f003a68e
Explicit (and improved) simprules for binary arithmetic.
 paulson parents: 
5747diff
changeset | 149 | |
| 9108 | 150 | bind_thms ("bin_mult_simps", [zmult_zminus, number_of_minus, number_of_add]);
 | 
| 1632 | 151 | |
| 6910 | 152 | Goal "number_of(bin_mult v w) = (number_of v * number_of w::int)"; | 
| 5184 | 153 | by (induct_tac "v" 1); | 
| 4686 | 154 | by (simp_tac (simpset() addsimps bin_mult_simps) 1); | 
| 155 | by (simp_tac (simpset() addsimps bin_mult_simps) 1); | |
| 5491 | 156 | by (asm_simp_tac | 
| 5540 | 157 | (simpset() addsimps bin_mult_simps @ [zadd_zmult_distrib] @ zadd_ac) 1); | 
| 6910 | 158 | qed "number_of_mult"; | 
| 5491 | 159 | |
| 1632 | 160 | |
| 6941 | 161 | (*The correctness of shifting. But it doesn't seem to give a measurable | 
| 162 | speed-up.*) | |
| 163 | Goal "(#2::int) * number_of w = number_of (w BIT False)"; | |
| 164 | by (induct_tac "w" 1); | |
| 165 | by (ALLGOALS (asm_simp_tac | |
| 166 | (simpset() addsimps bin_mult_simps @ [zadd_zmult_distrib] @ zadd_ac))); | |
| 167 | qed "double_number_of_BIT"; | |
| 168 | ||
| 169 | ||
| 5491 | 170 | (** Simplification rules with integer constants **) | 
| 171 | ||
| 6910 | 172 | Goal "#0 + z = (z::int)"; | 
| 5491 | 173 | by (Simp_tac 1); | 
| 174 | qed "zadd_0"; | |
| 175 | ||
| 6910 | 176 | Goal "z + #0 = (z::int)"; | 
| 5491 | 177 | by (Simp_tac 1); | 
| 178 | qed "zadd_0_right"; | |
| 179 | ||
| 5592 | 180 | Addsimps [zadd_0, zadd_0_right]; | 
| 181 | ||
| 182 | ||
| 183 | (** Converting simple cases of (int n) to numerals **) | |
| 5491 | 184 | |
| 5592 | 185 | (*int 0 = #0 *) | 
| 6910 | 186 | bind_thm ("int_0", number_of_Pls RS sym);
 | 
| 5491 | 187 | |
| 5592 | 188 | Goal "int (Suc n) = #1 + int n"; | 
| 189 | by (simp_tac (simpset() addsimps [zadd_int]) 1); | |
| 190 | qed "int_Suc"; | |
| 5510 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 191 | |
| 6910 | 192 | Goal "- (#0) = (#0::int)"; | 
| 5491 | 193 | by (Simp_tac 1); | 
| 194 | qed "zminus_0"; | |
| 195 | ||
| 196 | Addsimps [zminus_0]; | |
| 197 | ||
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 198 | |
| 6910 | 199 | Goal "(#0::int) - x = -x"; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 200 | by (simp_tac (simpset() addsimps [zdiff_def]) 1); | 
| 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 201 | qed "zdiff0"; | 
| 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 202 | |
| 6910 | 203 | Goal "x - (#0::int) = x"; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 204 | by (simp_tac (simpset() addsimps [zdiff_def]) 1); | 
| 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 205 | qed "zdiff0_right"; | 
| 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 206 | |
| 6910 | 207 | Goal "x - x = (#0::int)"; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 208 | by (simp_tac (simpset() addsimps [zdiff_def]) 1); | 
| 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 209 | qed "zdiff_self"; | 
| 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 210 | |
| 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 211 | Addsimps [zdiff0, zdiff0_right, zdiff_self]; | 
| 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 212 | |
| 6917 | 213 | |
| 214 | (** Special simplification, for constants only **) | |
| 6838 
941c4f70db91
rewrite rules to distribute CONSTANT multiplication over sum and difference;
 paulson parents: 
6716diff
changeset | 215 | |
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7033diff
changeset | 216 | (*Distributive laws for literals*) | 
| 6917 | 217 | Addsimps (map (inst "w" "number_of ?v") | 
| 6838 
941c4f70db91
rewrite rules to distribute CONSTANT multiplication over sum and difference;
 paulson parents: 
6716diff
changeset | 218 | [zadd_zmult_distrib, zadd_zmult_distrib2, | 
| 
941c4f70db91
rewrite rules to distribute CONSTANT multiplication over sum and difference;
 paulson parents: 
6716diff
changeset | 219 | zdiff_zmult_distrib, zdiff_zmult_distrib2]); | 
| 
941c4f70db91
rewrite rules to distribute CONSTANT multiplication over sum and difference;
 paulson parents: 
6716diff
changeset | 220 | |
| 6917 | 221 | Addsimps (map (inst "x" "number_of ?v") | 
| 222 | [zless_zminus, zle_zminus, equation_zminus]); | |
| 223 | Addsimps (map (inst "y" "number_of ?v") | |
| 224 | [zminus_zless, zminus_zle, zminus_equation]); | |
| 225 | ||
| 7074 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7033diff
changeset | 226 | (*Moving negation out of products*) | 
| 
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
 paulson parents: 
7033diff
changeset | 227 | Addsimps [zmult_zminus, zmult_zminus_right]; | 
| 6917 | 228 | |
| 6838 
941c4f70db91
rewrite rules to distribute CONSTANT multiplication over sum and difference;
 paulson parents: 
6716diff
changeset | 229 | (** Special-case simplification for small constants **) | 
| 
941c4f70db91
rewrite rules to distribute CONSTANT multiplication over sum and difference;
 paulson parents: 
6716diff
changeset | 230 | |
| 6910 | 231 | Goal "#0 * z = (#0::int)"; | 
| 5491 | 232 | by (Simp_tac 1); | 
| 233 | qed "zmult_0"; | |
| 234 | ||
| 6910 | 235 | Goal "z * #0 = (#0::int)"; | 
| 6838 
941c4f70db91
rewrite rules to distribute CONSTANT multiplication over sum and difference;
 paulson parents: 
6716diff
changeset | 236 | by (Simp_tac 1); | 
| 
941c4f70db91
rewrite rules to distribute CONSTANT multiplication over sum and difference;
 paulson parents: 
6716diff
changeset | 237 | qed "zmult_0_right"; | 
| 
941c4f70db91
rewrite rules to distribute CONSTANT multiplication over sum and difference;
 paulson parents: 
6716diff
changeset | 238 | |
| 6910 | 239 | Goal "#1 * z = (z::int)"; | 
| 5491 | 240 | by (Simp_tac 1); | 
| 241 | qed "zmult_1"; | |
| 242 | ||
| 6910 | 243 | Goal "z * #1 = (z::int)"; | 
| 6838 
941c4f70db91
rewrite rules to distribute CONSTANT multiplication over sum and difference;
 paulson parents: 
6716diff
changeset | 244 | by (Simp_tac 1); | 
| 
941c4f70db91
rewrite rules to distribute CONSTANT multiplication over sum and difference;
 paulson parents: 
6716diff
changeset | 245 | qed "zmult_1_right"; | 
| 
941c4f70db91
rewrite rules to distribute CONSTANT multiplication over sum and difference;
 paulson parents: 
6716diff
changeset | 246 | |
| 6917 | 247 | Goal "#-1 * z = -(z::int)"; | 
| 248 | by (simp_tac (simpset() addsimps zcompare_rls@[zmult_zminus]) 1); | |
| 249 | qed "zmult_minus1"; | |
| 250 | ||
| 251 | Goal "z * #-1 = -(z::int)"; | |
| 252 | by (simp_tac (simpset() addsimps zcompare_rls@[zmult_zminus_right]) 1); | |
| 253 | qed "zmult_minus1_right"; | |
| 254 | ||
| 255 | Addsimps [zmult_0, zmult_0_right, | |
| 256 | zmult_1, zmult_1_right, | |
| 257 | zmult_minus1, zmult_minus1_right]; | |
| 258 | ||
| 8785 | 259 | (*Negation of a coefficient*) | 
| 260 | Goal "- (number_of w) * z = number_of(bin_minus w) * (z::int)"; | |
| 261 | by (simp_tac (simpset_of Int.thy addsimps [number_of_minus, zmult_zminus]) 1); | |
| 262 | qed "zminus_number_of_zmult"; | |
| 263 | ||
| 264 | Addsimps [zminus_number_of_zmult]; | |
| 265 | ||
| 6917 | 266 | |
| 267 | (** Inequality reasoning **) | |
| 5491 | 268 | |
| 6989 | 269 | Goal "(m*n = (#0::int)) = (m = #0 | n = #0)"; | 
| 270 | by (stac (int_0 RS sym) 1 THEN rtac zmult_eq_int0_iff 1); | |
| 271 | qed "zmult_eq_0_iff"; | |
| 272 | ||
| 9062 | 273 | Goal "((#0::int) = m*n) = (#0 = m | #0 = n)"; | 
| 274 | by (stac eq_commute 1 THEN stac zmult_eq_0_iff 1); | |
| 275 | by Auto_tac; | |
| 276 | qed "zmult_0_eq_iff"; | |
| 277 | ||
| 278 | Addsimps [zmult_eq_0_iff, zmult_0_eq_iff]; | |
| 279 | ||
| 6910 | 280 | Goal "(w < z + (#1::int)) = (w<z | w=z)"; | 
| 5592 | 281 | by (simp_tac (simpset() addsimps [zless_add_int_Suc_eq]) 1); | 
| 5491 | 282 | qed "zless_add1_eq"; | 
| 283 | ||
| 6910 | 284 | Goal "(w + (#1::int) <= z) = (w<z)"; | 
| 5592 | 285 | by (simp_tac (simpset() addsimps [add_int_Suc_zle_eq]) 1); | 
| 5491 | 286 | qed "add1_zle_eq"; | 
| 6997 | 287 | |
| 288 | Goal "((#1::int) + w <= z) = (w<z)"; | |
| 289 | by (stac zadd_commute 1); | |
| 290 | by (rtac add1_zle_eq 1); | |
| 291 | qed "add1_left_zle_eq"; | |
| 5491 | 292 | |
| 5540 | 293 | Goal "neg x = (x < #0)"; | 
| 6917 | 294 | by (simp_tac (simpset() addsimps [neg_eq_less_int0]) 1); | 
| 5540 | 295 | qed "neg_eq_less_0"; | 
| 5510 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 296 | |
| 6989 | 297 | Goal "(~neg x) = (#0 <= x)"; | 
| 6917 | 298 | by (simp_tac (simpset() addsimps [not_neg_eq_ge_int0]) 1); | 
| 5540 | 299 | qed "not_neg_eq_ge_0"; | 
| 5510 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 300 | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 301 | Goal "#0 <= int m"; | 
| 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 302 | by (Simp_tac 1); | 
| 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 303 | qed "zero_zle_int"; | 
| 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 304 | AddIffs [zero_zle_int]; | 
| 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 305 | |
| 5510 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 306 | |
| 5747 | 307 | (** Needed because (int 0) rewrites to #0. | 
| 308 | Can these be generalized without evaluating large numbers?**) | |
| 309 | ||
| 310 | Goal "~ (int k < #0)"; | |
| 311 | by (Simp_tac 1); | |
| 312 | qed "int_less_0_conv"; | |
| 313 | ||
| 314 | Goal "(int k <= #0) = (k=0)"; | |
| 315 | by (Simp_tac 1); | |
| 316 | qed "int_le_0_conv"; | |
| 317 | ||
| 318 | Goal "(int k = #0) = (k=0)"; | |
| 319 | by (Simp_tac 1); | |
| 320 | qed "int_eq_0_conv"; | |
| 321 | ||
| 322 | Goal "(#0 = int k) = (k=0)"; | |
| 323 | by Auto_tac; | |
| 324 | qed "int_eq_0_conv'"; | |
| 325 | ||
| 326 | Addsimps [int_less_0_conv, int_le_0_conv, int_eq_0_conv, int_eq_0_conv']; | |
| 327 | ||
| 328 | ||
| 5491 | 329 | (** Simplification rules for comparison of binary numbers (Norbert Voelker) **) | 
| 330 | ||
| 331 | (** Equals (=) **) | |
| 1632 | 332 | |
| 5491 | 333 | Goalw [iszero_def] | 
| 6997 | 334 | "((number_of x::int) = number_of y) = \ | 
| 335 | \ iszero (number_of (bin_add x (bin_minus y)))"; | |
| 5491 | 336 | by (simp_tac (simpset() addsimps | 
| 6910 | 337 | (zcompare_rls @ [number_of_add, number_of_minus])) 1); | 
| 338 | qed "eq_number_of_eq"; | |
| 5491 | 339 | |
| 6910 | 340 | Goalw [iszero_def] "iszero ((number_of Pls)::int)"; | 
| 5491 | 341 | by (Simp_tac 1); | 
| 6910 | 342 | qed "iszero_number_of_Pls"; | 
| 5491 | 343 | |
| 6910 | 344 | Goalw [iszero_def] "~ iszero ((number_of Min)::int)"; | 
| 5491 | 345 | by (Simp_tac 1); | 
| 6910 | 346 | qed "nonzero_number_of_Min"; | 
| 5491 | 347 | |
| 348 | Goalw [iszero_def] | |
| 6910 | 349 | "iszero (number_of (w BIT x)) = (~x & iszero (number_of w::int))"; | 
| 5491 | 350 | by (Simp_tac 1); | 
| 6910 | 351 | by (int_case_tac "number_of w" 1); | 
| 5491 | 352 | by (ALLGOALS (asm_simp_tac | 
| 5540 | 353 | (simpset() addsimps zcompare_rls @ | 
| 354 | [zminus_zadd_distrib RS sym, | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 355 | zadd_int]))); | 
| 6910 | 356 | qed "iszero_number_of_BIT"; | 
| 5491 | 357 | |
| 6910 | 358 | Goal "iszero (number_of (w BIT False)) = iszero (number_of w::int)"; | 
| 359 | by (simp_tac (HOL_ss addsimps [iszero_number_of_BIT]) 1); | |
| 360 | qed "iszero_number_of_0"; | |
| 5779 
5c74f003a68e
Explicit (and improved) simprules for binary arithmetic.
 paulson parents: 
5747diff
changeset | 361 | |
| 6910 | 362 | Goal "~ iszero (number_of (w BIT True)::int)"; | 
| 363 | by (simp_tac (HOL_ss addsimps [iszero_number_of_BIT]) 1); | |
| 364 | qed "iszero_number_of_1"; | |
| 5779 
5c74f003a68e
Explicit (and improved) simprules for binary arithmetic.
 paulson parents: 
5747diff
changeset | 365 | |
| 
5c74f003a68e
Explicit (and improved) simprules for binary arithmetic.
 paulson parents: 
5747diff
changeset | 366 | |
| 5491 | 367 | |
| 368 | (** Less-than (<) **) | |
| 369 | ||
| 370 | Goalw [zless_def,zdiff_def] | |
| 6910 | 371 | "(number_of x::int) < number_of y \ | 
| 372 | \ = neg (number_of (bin_add x (bin_minus y)))"; | |
| 5491 | 373 | by (simp_tac (simpset() addsimps bin_mult_simps) 1); | 
| 6910 | 374 | qed "less_number_of_eq_neg"; | 
| 5491 | 375 | |
| 6910 | 376 | Goal "~ neg (number_of Pls)"; | 
| 5491 | 377 | by (Simp_tac 1); | 
| 6910 | 378 | qed "not_neg_number_of_Pls"; | 
| 5491 | 379 | |
| 6910 | 380 | Goal "neg (number_of Min)"; | 
| 5491 | 381 | by (Simp_tac 1); | 
| 6910 | 382 | qed "neg_number_of_Min"; | 
| 5491 | 383 | |
| 6910 | 384 | Goal "neg (number_of (w BIT x)) = neg (number_of w)"; | 
| 5491 | 385 | by (Asm_simp_tac 1); | 
| 6910 | 386 | by (int_case_tac "number_of w" 1); | 
| 5491 | 387 | by (ALLGOALS (asm_simp_tac | 
| 6917 | 388 | (simpset() addsimps [zadd_int, neg_eq_less_int0, | 
| 5540 | 389 | symmetric zdiff_def] @ zcompare_rls))); | 
| 6910 | 390 | qed "neg_number_of_BIT"; | 
| 5491 | 391 | |
| 392 | ||
| 393 | (** Less-than-or-equals (<=) **) | |
| 394 | ||
| 7033 | 395 | Goal "(number_of x <= (number_of y::int)) = \ | 
| 396 | \ (~ number_of y < (number_of x::int))"; | |
| 397 | by (rtac (linorder_not_less RS sym) 1); | |
| 6910 | 398 | qed "le_number_of_eq_not_less"; | 
| 5491 | 399 | |
| 5540 | 400 | (*Delete the original rewrites, with their clumsy conditional expressions*) | 
| 5551 | 401 | Delsimps [bin_succ_BIT, bin_pred_BIT, bin_minus_BIT, | 
| 402 | NCons_Pls, NCons_Min, bin_add_BIT, bin_mult_BIT]; | |
| 5491 | 403 | |
| 404 | (*Hide the binary representation of integer constants*) | |
| 6910 | 405 | Delsimps [number_of_Pls, number_of_Min, number_of_BIT]; | 
| 5491 | 406 | |
| 8787 | 407 | (*Simplification of arithmetic operations on integer constants. | 
| 408 | Note that bin_pred_Pls, etc. were added to the simpset by primrec.*) | |
| 409 | ||
| 9108 | 410 | bind_thms ("NCons_simps", [NCons_Pls_0, NCons_Pls_1, NCons_Min_0, NCons_Min_1, NCons_BIT]);
 | 
| 8787 | 411 | |
| 9108 | 412 | bind_thms ("bin_arith_extra_simps",
 | 
| 6910 | 413 | [number_of_add RS sym, | 
| 414 | number_of_minus RS sym, | |
| 415 | number_of_mult RS sym, | |
| 5779 
5c74f003a68e
Explicit (and improved) simprules for binary arithmetic.
 paulson parents: 
5747diff
changeset | 416 | bin_succ_1, bin_succ_0, | 
| 
5c74f003a68e
Explicit (and improved) simprules for binary arithmetic.
 paulson parents: 
5747diff
changeset | 417 | bin_pred_1, bin_pred_0, | 
| 
5c74f003a68e
Explicit (and improved) simprules for binary arithmetic.
 paulson parents: 
5747diff
changeset | 418 | bin_minus_1, bin_minus_0, | 
| 7517 
bad2f36810e1
generalized the theorem bin_add_BIT_Min to bin_add_Min_right
 paulson parents: 
7074diff
changeset | 419 | bin_add_Pls_right, bin_add_Min_right, | 
| 5779 
5c74f003a68e
Explicit (and improved) simprules for binary arithmetic.
 paulson parents: 
5747diff
changeset | 420 | bin_add_BIT_0, bin_add_BIT_10, bin_add_BIT_11, | 
| 6910 | 421 | diff_number_of_eq, | 
| 9108 | 422 | bin_mult_1, bin_mult_0] @ NCons_simps); | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 423 | |
| 5779 
5c74f003a68e
Explicit (and improved) simprules for binary arithmetic.
 paulson parents: 
5747diff
changeset | 424 | (*For making a minimal simpset, one must include these default simprules | 
| 6910 | 425 | of thy. Also include simp_thms, or at least (~False)=True*) | 
| 9108 | 426 | bind_thms ("bin_arith_simps",
 | 
| 5779 
5c74f003a68e
Explicit (and improved) simprules for binary arithmetic.
 paulson parents: 
5747diff
changeset | 427 | [bin_pred_Pls, bin_pred_Min, | 
| 
5c74f003a68e
Explicit (and improved) simprules for binary arithmetic.
 paulson parents: 
5747diff
changeset | 428 | bin_succ_Pls, bin_succ_Min, | 
| 
5c74f003a68e
Explicit (and improved) simprules for binary arithmetic.
 paulson parents: 
5747diff
changeset | 429 | bin_add_Pls, bin_add_Min, | 
| 
5c74f003a68e
Explicit (and improved) simprules for binary arithmetic.
 paulson parents: 
5747diff
changeset | 430 | bin_minus_Pls, bin_minus_Min, | 
| 9108 | 431 | bin_mult_Pls, bin_mult_Min] @ bin_arith_extra_simps); | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 432 | |
| 5779 
5c74f003a68e
Explicit (and improved) simprules for binary arithmetic.
 paulson parents: 
5747diff
changeset | 433 | (*Simplification of relational operations*) | 
| 9108 | 434 | bind_thms ("bin_rel_simps",
 | 
| 6910 | 435 | [eq_number_of_eq, iszero_number_of_Pls, nonzero_number_of_Min, | 
| 436 | iszero_number_of_0, iszero_number_of_1, | |
| 437 | less_number_of_eq_neg, | |
| 438 | not_neg_number_of_Pls, neg_number_of_Min, neg_number_of_BIT, | |
| 9108 | 439 | le_number_of_eq_not_less]); | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 440 | |
| 5779 
5c74f003a68e
Explicit (and improved) simprules for binary arithmetic.
 paulson parents: 
5747diff
changeset | 441 | Addsimps bin_arith_extra_simps; | 
| 
5c74f003a68e
Explicit (and improved) simprules for binary arithmetic.
 paulson parents: 
5747diff
changeset | 442 | Addsimps bin_rel_simps; | 
| 5510 
ad120f7c52ad
improved (but still flawed) treatment of binary arithmetic
 paulson parents: 
5491diff
changeset | 443 | |
| 6997 | 444 | |
| 8764 | 445 | (** Simplification of arithmetic when nested to the right **) | 
| 6997 | 446 | |
| 8764 | 447 | Goal "number_of v + (number_of w + z) = (number_of(bin_add v w) + z::int)"; | 
| 448 | by (simp_tac (simpset() addsimps [zadd_assoc RS sym]) 1); | |
| 449 | qed "add_number_of_left"; | |
| 450 | ||
| 451 | Goal "number_of v * (number_of w * z) = (number_of(bin_mult v w) * z::int)"; | |
| 452 | by (simp_tac (simpset() addsimps [zmult_assoc RS sym]) 1); | |
| 453 | qed "mult_number_of_left"; | |
| 6997 | 454 | |
| 455 | Goalw [zdiff_def] | |
| 456 | "number_of v + (number_of w - c) = number_of(bin_add v w) - (c::int)"; | |
| 8764 | 457 | by (rtac add_number_of_left 1); | 
| 458 | qed "add_number_of_diff1"; | |
| 6997 | 459 | |
| 460 | Goal "number_of v + (c - number_of w) = \ | |
| 461 | \ number_of (bin_add v (bin_minus w)) + (c::int)"; | |
| 462 | by (stac (diff_number_of_eq RS sym) 1); | |
| 463 | by Auto_tac; | |
| 8764 | 464 | qed "add_number_of_diff2"; | 
| 6997 | 465 | |
| 8764 | 466 | Addsimps [add_number_of_left, mult_number_of_left, | 
| 467 | add_number_of_diff1, add_number_of_diff2]; |