| author | wenzelm | 
| Tue, 19 Nov 2019 15:45:41 +0100 | |
| changeset 71141 | b1c555d3cd71 | 
| parent 70193 | 49a65e3f04c9 | 
| child 71181 | 8331063570d6 | 
| permissions | -rw-r--r-- | 
| 65363 | 1 | (* Title: HOL/Word/Bits_Int.thy | 
| 2 | Author: Jeremy Dawson and Gerwin Klein, NICTA | |
| 24333 | 3 | |
| 65363 | 4 | Definitions and basic theorems for bit-wise logical operations | 
| 5 | for integers expressed using Pls, Min, BIT, | |
| 6 | and converting them to and from lists of bools. | |
| 7 | *) | |
| 24333 | 8 | |
| 70190 | 9 | section \<open>Bitwise Operations on integers\<close> | 
| 24350 | 10 | |
| 54854 
3324a0078636
prefer "Bits" as theory name for abstract bit operations, similar to "Orderings", "Lattices", "Groups" etc.
 haftmann parents: 
54848diff
changeset | 11 | theory Bits_Int | 
| 70190 | 12 | imports Bits Misc_Auxiliary | 
| 24333 | 13 | begin | 
| 14 | ||
| 70190 | 15 | subsection \<open>Implicit bit representation of \<^typ>\<open>int\<close>\<close> | 
| 16 | ||
| 17 | definition Bit :: "int \<Rightarrow> bool \<Rightarrow> int" (infixl "BIT" 90) | |
| 18 | where "k BIT b = (if b then 1 else 0) + k + k" | |
| 19 | ||
| 20 | lemma Bit_B0: "k BIT False = k + k" | |
| 21 | by (simp add: Bit_def) | |
| 22 | ||
| 23 | lemma Bit_B1: "k BIT True = k + k + 1" | |
| 24 | by (simp add: Bit_def) | |
| 25 | ||
| 26 | lemma Bit_B0_2t: "k BIT False = 2 * k" | |
| 27 | by (rule trans, rule Bit_B0) simp | |
| 28 | ||
| 29 | lemma Bit_B1_2t: "k BIT True = 2 * k + 1" | |
| 30 | by (rule trans, rule Bit_B1) simp | |
| 31 | ||
| 32 | lemma uminus_Bit_eq: | |
| 33 | "- k BIT b = (- k - of_bool b) BIT b" | |
| 34 | by (cases b) (simp_all add: Bit_def) | |
| 35 | ||
| 36 | lemma power_BIT: "2 ^ Suc n - 1 = (2 ^ n - 1) BIT True" | |
| 37 | by (simp add: Bit_B1) | |
| 38 | ||
| 39 | definition bin_last :: "int \<Rightarrow> bool" | |
| 40 | where "bin_last w \<longleftrightarrow> w mod 2 = 1" | |
| 41 | ||
| 42 | lemma bin_last_odd: "bin_last = odd" | |
| 43 | by (rule ext) (simp add: bin_last_def even_iff_mod_2_eq_zero) | |
| 44 | ||
| 45 | definition bin_rest :: "int \<Rightarrow> int" | |
| 46 | where "bin_rest w = w div 2" | |
| 47 | ||
| 48 | lemma bin_rl_simp [simp]: "bin_rest w BIT bin_last w = w" | |
| 49 | unfolding bin_rest_def bin_last_def Bit_def | |
| 50 | by (cases "w mod 2 = 0") (use div_mult_mod_eq [of w 2] in simp_all) | |
| 51 | ||
| 52 | lemma bin_rest_BIT [simp]: "bin_rest (x BIT b) = x" | |
| 53 | unfolding bin_rest_def Bit_def | |
| 54 | by (cases b) simp_all | |
| 55 | ||
| 56 | lemma bin_last_BIT [simp]: "bin_last (x BIT b) = b" | |
| 57 | unfolding bin_last_def Bit_def | |
| 58 | by (cases b) simp_all | |
| 59 | ||
| 60 | lemma BIT_eq_iff [iff]: "u BIT b = v BIT c \<longleftrightarrow> u = v \<and> b = c" | |
| 61 | by (auto simp: Bit_def) arith+ | |
| 62 | ||
| 63 | lemma BIT_bin_simps [simp]: | |
| 64 | "numeral k BIT False = numeral (Num.Bit0 k)" | |
| 65 | "numeral k BIT True = numeral (Num.Bit1 k)" | |
| 66 | "(- numeral k) BIT False = - numeral (Num.Bit0 k)" | |
| 67 | "(- numeral k) BIT True = - numeral (Num.BitM k)" | |
| 68 | unfolding numeral.simps numeral_BitM | |
| 69 | by (simp_all add: Bit_def del: arith_simps add_numeral_special diff_numeral_special) | |
| 70 | ||
| 71 | lemma BIT_special_simps [simp]: | |
| 72 | shows "0 BIT False = 0" | |
| 73 | and "0 BIT True = 1" | |
| 74 | and "1 BIT False = 2" | |
| 75 | and "1 BIT True = 3" | |
| 76 | and "(- 1) BIT False = - 2" | |
| 77 | and "(- 1) BIT True = - 1" | |
| 78 | by (simp_all add: Bit_def) | |
| 79 | ||
| 80 | lemma Bit_eq_0_iff: "w BIT b = 0 \<longleftrightarrow> w = 0 \<and> \<not> b" | |
| 81 | by (auto simp: Bit_def) arith | |
| 82 | ||
| 83 | lemma Bit_eq_m1_iff: "w BIT b = -1 \<longleftrightarrow> w = -1 \<and> b" | |
| 84 | by (auto simp: Bit_def) arith | |
| 85 | ||
| 86 | lemma BitM_inc: "Num.BitM (Num.inc w) = Num.Bit1 w" | |
| 87 | by (induct w) simp_all | |
| 88 | ||
| 89 | lemma expand_BIT: | |
| 90 | "numeral (Num.Bit0 w) = numeral w BIT False" | |
| 91 | "numeral (Num.Bit1 w) = numeral w BIT True" | |
| 92 | "- numeral (Num.Bit0 w) = (- numeral w) BIT False" | |
| 93 | "- numeral (Num.Bit1 w) = (- numeral (w + Num.One)) BIT True" | |
| 94 | by (simp_all add: add_One BitM_inc) | |
| 95 | ||
| 96 | lemma bin_last_numeral_simps [simp]: | |
| 97 | "\<not> bin_last 0" | |
| 98 | "bin_last 1" | |
| 99 | "bin_last (- 1)" | |
| 100 | "bin_last Numeral1" | |
| 101 | "\<not> bin_last (numeral (Num.Bit0 w))" | |
| 102 | "bin_last (numeral (Num.Bit1 w))" | |
| 103 | "\<not> bin_last (- numeral (Num.Bit0 w))" | |
| 104 | "bin_last (- numeral (Num.Bit1 w))" | |
| 105 | by (simp_all add: bin_last_def zmod_zminus1_eq_if) (auto simp add: divmod_def) | |
| 106 | ||
| 107 | lemma bin_rest_numeral_simps [simp]: | |
| 108 | "bin_rest 0 = 0" | |
| 109 | "bin_rest 1 = 0" | |
| 110 | "bin_rest (- 1) = - 1" | |
| 111 | "bin_rest Numeral1 = 0" | |
| 112 | "bin_rest (numeral (Num.Bit0 w)) = numeral w" | |
| 113 | "bin_rest (numeral (Num.Bit1 w)) = numeral w" | |
| 114 | "bin_rest (- numeral (Num.Bit0 w)) = - numeral w" | |
| 115 | "bin_rest (- numeral (Num.Bit1 w)) = - numeral (w + Num.One)" | |
| 116 | by (simp_all add: bin_rest_def zdiv_zminus1_eq_if) (auto simp add: divmod_def) | |
| 117 | ||
| 118 | lemma less_Bits: "v BIT b < w BIT c \<longleftrightarrow> v < w \<or> v \<le> w \<and> \<not> b \<and> c" | |
| 119 | by (auto simp: Bit_def) | |
| 120 | ||
| 121 | lemma le_Bits: "v BIT b \<le> w BIT c \<longleftrightarrow> v < w \<or> v \<le> w \<and> (\<not> b \<or> c)" | |
| 122 | by (auto simp: Bit_def) | |
| 123 | ||
| 124 | lemma pred_BIT_simps [simp]: | |
| 125 | "x BIT False - 1 = (x - 1) BIT True" | |
| 126 | "x BIT True - 1 = x BIT False" | |
| 127 | by (simp_all add: Bit_B0_2t Bit_B1_2t) | |
| 128 | ||
| 129 | lemma succ_BIT_simps [simp]: | |
| 130 | "x BIT False + 1 = x BIT True" | |
| 131 | "x BIT True + 1 = (x + 1) BIT False" | |
| 132 | by (simp_all add: Bit_B0_2t Bit_B1_2t) | |
| 133 | ||
| 134 | lemma add_BIT_simps [simp]: | |
| 135 | "x BIT False + y BIT False = (x + y) BIT False" | |
| 136 | "x BIT False + y BIT True = (x + y) BIT True" | |
| 137 | "x BIT True + y BIT False = (x + y) BIT True" | |
| 138 | "x BIT True + y BIT True = (x + y + 1) BIT False" | |
| 139 | by (simp_all add: Bit_B0_2t Bit_B1_2t) | |
| 140 | ||
| 141 | lemma mult_BIT_simps [simp]: | |
| 142 | "x BIT False * y = (x * y) BIT False" | |
| 143 | "x * y BIT False = (x * y) BIT False" | |
| 144 | "x BIT True * y = (x * y) BIT False + y" | |
| 145 | by (simp_all add: Bit_B0_2t Bit_B1_2t algebra_simps) | |
| 146 | ||
| 147 | lemma B_mod_2': "X = 2 \<Longrightarrow> (w BIT True) mod X = 1 \<and> (w BIT False) mod X = 0" | |
| 148 | by (simp add: Bit_B0 Bit_B1) | |
| 149 | ||
| 150 | lemma bin_ex_rl: "\<exists>w b. w BIT b = bin" | |
| 151 | by (metis bin_rl_simp) | |
| 152 | ||
| 153 | lemma bin_exhaust: "(\<And>x b. bin = x BIT b \<Longrightarrow> Q) \<Longrightarrow> Q" | |
| 154 | by (metis bin_ex_rl) | |
| 155 | ||
| 156 | lemma bin_abs_lem: "bin = (w BIT b) \<Longrightarrow> bin \<noteq> -1 \<longrightarrow> bin \<noteq> 0 \<longrightarrow> nat \<bar>w\<bar> < nat \<bar>bin\<bar>" | |
| 157 | apply clarsimp | |
| 158 | apply (unfold Bit_def) | |
| 159 | apply (cases b) | |
| 160 | apply (clarsimp, arith) | |
| 161 | apply (clarsimp, arith) | |
| 162 | done | |
| 163 | ||
| 164 | lemma bin_induct: | |
| 165 | assumes PPls: "P 0" | |
| 166 | and PMin: "P (- 1)" | |
| 167 | and PBit: "\<And>bin bit. P bin \<Longrightarrow> P (bin BIT bit)" | |
| 168 | shows "P bin" | |
| 169 | apply (rule_tac P=P and a=bin and f1="nat \<circ> abs" in wf_measure [THEN wf_induct]) | |
| 170 | apply (simp add: measure_def inv_image_def) | |
| 171 | apply (case_tac x rule: bin_exhaust) | |
| 172 | apply (frule bin_abs_lem) | |
| 173 | apply (auto simp add : PPls PMin PBit) | |
| 174 | done | |
| 175 | ||
| 176 | lemma Bit_div2 [simp]: "(w BIT b) div 2 = w" | |
| 177 | unfolding bin_rest_def [symmetric] by (rule bin_rest_BIT) | |
| 178 | ||
| 179 | lemma bin_rl_eqI: "\<lbrakk>bin_rest x = bin_rest y; bin_last x = bin_last y\<rbrakk> \<Longrightarrow> x = y" | |
| 180 | by (metis (mono_tags) BIT_eq_iff bin_ex_rl bin_last_BIT bin_rest_BIT) | |
| 181 | ||
| 182 | lemma twice_conv_BIT: "2 * x = x BIT False" | |
| 183 | by (rule bin_rl_eqI) (simp_all, simp_all add: bin_rest_def bin_last_def) | |
| 184 | ||
| 185 | lemma BIT_lt0 [simp]: "x BIT b < 0 \<longleftrightarrow> x < 0" | |
| 186 | by(cases b)(auto simp add: Bit_def) | |
| 187 | ||
| 188 | lemma BIT_ge0 [simp]: "x BIT b \<ge> 0 \<longleftrightarrow> x \<ge> 0" | |
| 189 | by(cases b)(auto simp add: Bit_def) | |
| 190 | ||
| 191 | lemma [simp]: | |
| 192 | shows bin_rest_lt0: "bin_rest i < 0 \<longleftrightarrow> i < 0" | |
| 193 | and bin_rest_ge_0: "bin_rest i \<ge> 0 \<longleftrightarrow> i \<ge> 0" | |
| 194 | by(auto simp add: bin_rest_def) | |
| 195 | ||
| 196 | lemma bin_rest_gt_0 [simp]: "bin_rest x > 0 \<longleftrightarrow> x > 1" | |
| 197 | by(simp add: bin_rest_def add1_zle_eq pos_imp_zdiv_pos_iff) (metis add1_zle_eq one_add_one) | |
| 198 | ||
| 199 | ||
| 200 | subsection \<open>Explicit bit representation of \<^typ>\<open>int\<close>\<close> | |
| 201 | ||
| 202 | primrec bl_to_bin_aux :: "bool list \<Rightarrow> int \<Rightarrow> int" | |
| 203 | where | |
| 204 | Nil: "bl_to_bin_aux [] w = w" | |
| 205 | | Cons: "bl_to_bin_aux (b # bs) w = bl_to_bin_aux bs (w BIT b)" | |
| 206 | ||
| 207 | definition bl_to_bin :: "bool list \<Rightarrow> int" | |
| 208 | where "bl_to_bin bs = bl_to_bin_aux bs 0" | |
| 209 | ||
| 210 | primrec bin_to_bl_aux :: "nat \<Rightarrow> int \<Rightarrow> bool list \<Rightarrow> bool list" | |
| 211 | where | |
| 212 | Z: "bin_to_bl_aux 0 w bl = bl" | |
| 213 | | Suc: "bin_to_bl_aux (Suc n) w bl = bin_to_bl_aux n (bin_rest w) ((bin_last w) # bl)" | |
| 214 | ||
| 215 | definition bin_to_bl :: "nat \<Rightarrow> int \<Rightarrow> bool list" | |
| 216 | where "bin_to_bl n w = bin_to_bl_aux n w []" | |
| 217 | ||
| 218 | lemma bin_to_bl_aux_zero_minus_simp [simp]: | |
| 219 | "0 < n \<Longrightarrow> bin_to_bl_aux n 0 bl = bin_to_bl_aux (n - 1) 0 (False # bl)" | |
| 220 | by (cases n) auto | |
| 221 | ||
| 222 | lemma bin_to_bl_aux_minus1_minus_simp [simp]: | |
| 223 | "0 < n \<Longrightarrow> bin_to_bl_aux n (- 1) bl = bin_to_bl_aux (n - 1) (- 1) (True # bl)" | |
| 224 | by (cases n) auto | |
| 225 | ||
| 226 | lemma bin_to_bl_aux_one_minus_simp [simp]: | |
| 227 | "0 < n \<Longrightarrow> bin_to_bl_aux n 1 bl = bin_to_bl_aux (n - 1) 0 (True # bl)" | |
| 228 | by (cases n) auto | |
| 229 | ||
| 230 | lemma bin_to_bl_aux_Bit_minus_simp [simp]: | |
| 231 | "0 < n \<Longrightarrow> bin_to_bl_aux n (w BIT b) bl = bin_to_bl_aux (n - 1) w (b # bl)" | |
| 232 | by (cases n) auto | |
| 233 | ||
| 234 | lemma bin_to_bl_aux_Bit0_minus_simp [simp]: | |
| 235 | "0 < n \<Longrightarrow> | |
| 236 | bin_to_bl_aux n (numeral (Num.Bit0 w)) bl = bin_to_bl_aux (n - 1) (numeral w) (False # bl)" | |
| 237 | by (cases n) auto | |
| 238 | ||
| 239 | lemma bin_to_bl_aux_Bit1_minus_simp [simp]: | |
| 240 | "0 < n \<Longrightarrow> | |
| 241 | bin_to_bl_aux n (numeral (Num.Bit1 w)) bl = bin_to_bl_aux (n - 1) (numeral w) (True # bl)" | |
| 242 | by (cases n) auto | |
| 243 | ||
| 244 | lemma bl_to_bin_aux_append: "bl_to_bin_aux (bs @ cs) w = bl_to_bin_aux cs (bl_to_bin_aux bs w)" | |
| 245 | by (induct bs arbitrary: w) auto | |
| 246 | ||
| 247 | lemma bin_to_bl_aux_append: "bin_to_bl_aux n w bs @ cs = bin_to_bl_aux n w (bs @ cs)" | |
| 248 | by (induct n arbitrary: w bs) auto | |
| 249 | ||
| 250 | lemma bl_to_bin_append: "bl_to_bin (bs @ cs) = bl_to_bin_aux cs (bl_to_bin bs)" | |
| 251 | unfolding bl_to_bin_def by (rule bl_to_bin_aux_append) | |
| 252 | ||
| 253 | lemma bin_to_bl_aux_alt: "bin_to_bl_aux n w bs = bin_to_bl n w @ bs" | |
| 254 | by (simp add: bin_to_bl_def bin_to_bl_aux_append) | |
| 255 | ||
| 256 | lemma bin_to_bl_0 [simp]: "bin_to_bl 0 bs = []" | |
| 257 | by (auto simp: bin_to_bl_def) | |
| 258 | ||
| 259 | lemma size_bin_to_bl_aux: "length (bin_to_bl_aux n w bs) = n + length bs" | |
| 260 | by (induct n arbitrary: w bs) auto | |
| 261 | ||
| 262 | lemma size_bin_to_bl [simp]: "length (bin_to_bl n w) = n" | |
| 263 | by (simp add: bin_to_bl_def size_bin_to_bl_aux) | |
| 264 | ||
| 265 | lemma bl_bin_bl': "bin_to_bl (n + length bs) (bl_to_bin_aux bs w) = bin_to_bl_aux n w bs" | |
| 266 | apply (induct bs arbitrary: w n) | |
| 267 | apply auto | |
| 268 | apply (simp_all only: add_Suc [symmetric]) | |
| 269 | apply (auto simp add: bin_to_bl_def) | |
| 270 | done | |
| 271 | ||
| 272 | lemma bl_bin_bl [simp]: "bin_to_bl (length bs) (bl_to_bin bs) = bs" | |
| 273 | unfolding bl_to_bin_def | |
| 274 | apply (rule box_equals) | |
| 275 | apply (rule bl_bin_bl') | |
| 276 | prefer 2 | |
| 277 | apply (rule bin_to_bl_aux.Z) | |
| 278 | apply simp | |
| 279 | done | |
| 280 | ||
| 281 | lemma bl_to_bin_inj: "bl_to_bin bs = bl_to_bin cs \<Longrightarrow> length bs = length cs \<Longrightarrow> bs = cs" | |
| 282 | apply (rule_tac box_equals) | |
| 283 | defer | |
| 284 | apply (rule bl_bin_bl) | |
| 285 | apply (rule bl_bin_bl) | |
| 286 | apply simp | |
| 287 | done | |
| 288 | ||
| 289 | lemma bl_to_bin_False [simp]: "bl_to_bin (False # bl) = bl_to_bin bl" | |
| 290 | by (auto simp: bl_to_bin_def) | |
| 291 | ||
| 292 | lemma bl_to_bin_Nil [simp]: "bl_to_bin [] = 0" | |
| 293 | by (auto simp: bl_to_bin_def) | |
| 294 | ||
| 295 | lemma bin_to_bl_zero_aux: "bin_to_bl_aux n 0 bl = replicate n False @ bl" | |
| 296 | by (induct n arbitrary: bl) (auto simp: replicate_app_Cons_same) | |
| 297 | ||
| 298 | lemma bin_to_bl_zero: "bin_to_bl n 0 = replicate n False" | |
| 299 | by (simp add: bin_to_bl_def bin_to_bl_zero_aux) | |
| 300 | ||
| 301 | lemma bin_to_bl_minus1_aux: "bin_to_bl_aux n (- 1) bl = replicate n True @ bl" | |
| 302 | by (induct n arbitrary: bl) (auto simp: replicate_app_Cons_same) | |
| 303 | ||
| 304 | lemma bin_to_bl_minus1: "bin_to_bl n (- 1) = replicate n True" | |
| 305 | by (simp add: bin_to_bl_def bin_to_bl_minus1_aux) | |
| 306 | ||
| 307 | lemma bl_to_bin_BIT: | |
| 308 | "bl_to_bin bs BIT b = bl_to_bin (bs @ [b])" | |
| 309 | by (simp add: bl_to_bin_append) | |
| 310 | ||
| 311 | ||
| 312 | subsection \<open>Bit projection\<close> | |
| 313 | ||
| 314 | primrec bin_nth :: "int \<Rightarrow> nat \<Rightarrow> bool" | |
| 315 | where | |
| 316 | Z: "bin_nth w 0 \<longleftrightarrow> bin_last w" | |
| 317 | | Suc: "bin_nth w (Suc n) \<longleftrightarrow> bin_nth (bin_rest w) n" | |
| 318 | ||
| 319 | lemma bin_nth_eq_mod: | |
| 320 | "bin_nth w n \<longleftrightarrow> odd (w div 2 ^ n)" | |
| 321 | by (induction n arbitrary: w) (simp_all add: bin_last_def bin_rest_def odd_iff_mod_2_eq_one zdiv_zmult2_eq) | |
| 322 | ||
| 323 | lemma bin_nth_eq_iff: "bin_nth x = bin_nth y \<longleftrightarrow> x = y" | |
| 324 | proof - | |
| 325 | have bin_nth_lem [rule_format]: "\<forall>y. bin_nth x = bin_nth y \<longrightarrow> x = y" | |
| 326 | apply (induct x rule: bin_induct) | |
| 327 | apply safe | |
| 328 | apply (erule rev_mp) | |
| 329 | apply (induct_tac y rule: bin_induct) | |
| 330 | apply safe | |
| 331 | apply (drule_tac x=0 in fun_cong, force) | |
| 332 | apply (erule notE, rule ext, drule_tac x="Suc x" in fun_cong, force) | |
| 333 | apply (drule_tac x=0 in fun_cong, force) | |
| 334 | apply (erule rev_mp) | |
| 335 | apply (induct_tac y rule: bin_induct) | |
| 336 | apply safe | |
| 337 | apply (drule_tac x=0 in fun_cong, force) | |
| 338 | apply (erule notE, rule ext, drule_tac x="Suc x" in fun_cong, force) | |
| 339 | apply (metis Bit_eq_m1_iff Z bin_last_BIT) | |
| 340 | apply (case_tac y rule: bin_exhaust) | |
| 341 | apply clarify | |
| 342 | apply (erule allE) | |
| 343 | apply (erule impE) | |
| 344 | prefer 2 | |
| 345 | apply (erule conjI) | |
| 346 | apply (drule_tac x=0 in fun_cong, force) | |
| 347 | apply (rule ext) | |
| 348 | apply (drule_tac x="Suc x" for x in fun_cong, force) | |
| 349 | done | |
| 350 | show ?thesis | |
| 351 | by (auto elim: bin_nth_lem) | |
| 352 | qed | |
| 353 | ||
| 354 | lemma bin_eqI: | |
| 355 | "x = y" if "\<And>n. bin_nth x n \<longleftrightarrow> bin_nth y n" | |
| 356 | using that bin_nth_eq_iff [of x y] by (simp add: fun_eq_iff) | |
| 357 | ||
| 358 | lemma bin_eq_iff: "x = y \<longleftrightarrow> (\<forall>n. bin_nth x n = bin_nth y n)" | |
| 359 | using bin_nth_eq_iff by auto | |
| 360 | ||
| 361 | lemma bin_nth_zero [simp]: "\<not> bin_nth 0 n" | |
| 362 | by (induct n) auto | |
| 363 | ||
| 364 | lemma bin_nth_1 [simp]: "bin_nth 1 n \<longleftrightarrow> n = 0" | |
| 365 | by (cases n) simp_all | |
| 366 | ||
| 367 | lemma bin_nth_minus1 [simp]: "bin_nth (- 1) n" | |
| 368 | by (induct n) auto | |
| 369 | ||
| 370 | lemma bin_nth_0_BIT: "bin_nth (w BIT b) 0 \<longleftrightarrow> b" | |
| 371 | by auto | |
| 372 | ||
| 373 | lemma bin_nth_Suc_BIT: "bin_nth (w BIT b) (Suc n) = bin_nth w n" | |
| 374 | by auto | |
| 375 | ||
| 376 | lemma bin_nth_minus [simp]: "0 < n \<Longrightarrow> bin_nth (w BIT b) n = bin_nth w (n - 1)" | |
| 377 | by (cases n) auto | |
| 378 | ||
| 379 | lemma bin_nth_numeral: "bin_rest x = y \<Longrightarrow> bin_nth x (numeral n) = bin_nth y (pred_numeral n)" | |
| 380 | by (simp add: numeral_eq_Suc) | |
| 381 | ||
| 382 | lemmas bin_nth_numeral_simps [simp] = | |
| 383 | bin_nth_numeral [OF bin_rest_numeral_simps(2)] | |
| 384 | bin_nth_numeral [OF bin_rest_numeral_simps(5)] | |
| 385 | bin_nth_numeral [OF bin_rest_numeral_simps(6)] | |
| 386 | bin_nth_numeral [OF bin_rest_numeral_simps(7)] | |
| 387 | bin_nth_numeral [OF bin_rest_numeral_simps(8)] | |
| 388 | ||
| 389 | lemmas bin_nth_simps = | |
| 390 | bin_nth.Z bin_nth.Suc bin_nth_zero bin_nth_minus1 | |
| 391 | bin_nth_numeral_simps | |
| 392 | ||
| 393 | lemma nth_2p_bin: "bin_nth (2 ^ n) m = (m = n)" \<comment> \<open>for use when simplifying with \<open>bin_nth_Bit\<close>\<close> | |
| 394 | apply (induct n arbitrary: m) | |
| 395 | apply clarsimp | |
| 396 | apply safe | |
| 397 | apply (case_tac m) | |
| 398 | apply (auto simp: Bit_B0_2t [symmetric]) | |
| 399 | done | |
| 400 | ||
| 401 | lemma nth_rest_power_bin: "bin_nth ((bin_rest ^^ k) w) n = bin_nth w (n + k)" | |
| 402 | apply (induct k arbitrary: n) | |
| 403 | apply clarsimp | |
| 404 | apply clarsimp | |
| 405 | apply (simp only: bin_nth.Suc [symmetric] add_Suc) | |
| 406 | done | |
| 407 | ||
| 408 | lemma bin_nth_numeral_unfold: | |
| 409 | "bin_nth (numeral (num.Bit0 x)) n \<longleftrightarrow> n > 0 \<and> bin_nth (numeral x) (n - 1)" | |
| 410 | "bin_nth (numeral (num.Bit1 x)) n \<longleftrightarrow> (n > 0 \<longrightarrow> bin_nth (numeral x) (n - 1))" | |
| 411 | by(case_tac [!] n) simp_all | |
| 412 | ||
| 413 | ||
| 414 | subsection \<open>Truncating\<close> | |
| 415 | ||
| 416 | definition bin_sign :: "int \<Rightarrow> int" | |
| 417 | where "bin_sign k = (if k \<ge> 0 then 0 else - 1)" | |
| 418 | ||
| 419 | lemma bin_sign_simps [simp]: | |
| 420 | "bin_sign 0 = 0" | |
| 421 | "bin_sign 1 = 0" | |
| 422 | "bin_sign (- 1) = - 1" | |
| 423 | "bin_sign (numeral k) = 0" | |
| 424 | "bin_sign (- numeral k) = -1" | |
| 425 | "bin_sign (w BIT b) = bin_sign w" | |
| 426 | by (simp_all add: bin_sign_def Bit_def) | |
| 427 | ||
| 428 | lemma bin_sign_rest [simp]: "bin_sign (bin_rest w) = bin_sign w" | |
| 429 | by (cases w rule: bin_exhaust) auto | |
| 430 | ||
| 431 | primrec bintrunc :: "nat \<Rightarrow> int \<Rightarrow> int" | |
| 432 | where | |
| 433 | Z : "bintrunc 0 bin = 0" | |
| 434 | | Suc : "bintrunc (Suc n) bin = bintrunc n (bin_rest bin) BIT (bin_last bin)" | |
| 435 | ||
| 436 | primrec sbintrunc :: "nat \<Rightarrow> int \<Rightarrow> int" | |
| 437 | where | |
| 438 | Z : "sbintrunc 0 bin = (if bin_last bin then -1 else 0)" | |
| 439 | | Suc : "sbintrunc (Suc n) bin = sbintrunc n (bin_rest bin) BIT (bin_last bin)" | |
| 440 | ||
| 441 | lemma bintrunc_mod2p: "bintrunc n w = w mod 2 ^ n" | |
| 442 | by (induct n arbitrary: w) (auto simp add: bin_last_def bin_rest_def Bit_def zmod_zmult2_eq) | |
| 443 | ||
| 444 | lemma sbintrunc_mod2p: "sbintrunc n w = (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n" | |
| 445 | proof (induction n arbitrary: w) | |
| 446 | case 0 | |
| 447 | then show ?case | |
| 448 | by (auto simp add: bin_last_odd odd_iff_mod_2_eq_one) | |
| 449 | next | |
| 450 | case (Suc n) | |
| 451 | moreover have "((bin_rest w + 2 ^ n) mod (2 * 2 ^ n) - 2 ^ n) BIT bin_last w = | |
| 452 | (w + 2 * 2 ^ n) mod (4 * 2 ^ n) - 2 * 2 ^ n" | |
| 453 | proof (cases w rule: parity_cases) | |
| 454 | case even | |
| 455 | then show ?thesis | |
| 456 | by (simp add: bin_last_odd bin_rest_def Bit_B0_2t mult_mod_right) | |
| 457 | next | |
| 458 | case odd | |
| 459 | then have "2 * (w div 2) = w - 1" | |
| 460 | using minus_mod_eq_mult_div [of w 2] by simp | |
| 461 | moreover have "(2 * 2 ^ n + w - 1) mod (2 * 2 * 2 ^ n) + 1 = (2 * 2 ^ n + w) mod (2 * 2 * 2 ^ n)" | |
| 462 | using odd emep1 [of "2 * 2 ^ n + w - 1" "2 * 2 * 2 ^ n"] by simp | |
| 463 | ultimately show ?thesis | |
| 464 | using odd by (simp add: bin_last_odd bin_rest_def Bit_B1_2t mult_mod_right) (simp add: algebra_simps) | |
| 465 | qed | |
| 466 | ultimately show ?case | |
| 467 | by simp | |
| 468 | qed | |
| 469 | ||
| 470 | lemma sign_bintr: "bin_sign (bintrunc n w) = 0" | |
| 471 | by (simp add: bintrunc_mod2p bin_sign_def) | |
| 472 | ||
| 473 | lemma bintrunc_n_0 [simp]: "bintrunc n 0 = 0" | |
| 474 | by (simp add: bintrunc_mod2p) | |
| 475 | ||
| 476 | lemma sbintrunc_n_0 [simp]: "sbintrunc n 0 = 0" | |
| 477 | by (simp add: sbintrunc_mod2p) | |
| 478 | ||
| 479 | lemma sbintrunc_n_minus1 [simp]: "sbintrunc n (- 1) = -1" | |
| 480 | by (induct n) auto | |
| 481 | ||
| 482 | lemma bintrunc_Suc_numeral: | |
| 483 | "bintrunc (Suc n) 1 = 1" | |
| 484 | "bintrunc (Suc n) (- 1) = bintrunc n (- 1) BIT True" | |
| 485 | "bintrunc (Suc n) (numeral (Num.Bit0 w)) = bintrunc n (numeral w) BIT False" | |
| 486 | "bintrunc (Suc n) (numeral (Num.Bit1 w)) = bintrunc n (numeral w) BIT True" | |
| 487 | "bintrunc (Suc n) (- numeral (Num.Bit0 w)) = bintrunc n (- numeral w) BIT False" | |
| 488 | "bintrunc (Suc n) (- numeral (Num.Bit1 w)) = bintrunc n (- numeral (w + Num.One)) BIT True" | |
| 489 | by simp_all | |
| 490 | ||
| 491 | lemma sbintrunc_0_numeral [simp]: | |
| 492 | "sbintrunc 0 1 = -1" | |
| 493 | "sbintrunc 0 (numeral (Num.Bit0 w)) = 0" | |
| 494 | "sbintrunc 0 (numeral (Num.Bit1 w)) = -1" | |
| 495 | "sbintrunc 0 (- numeral (Num.Bit0 w)) = 0" | |
| 496 | "sbintrunc 0 (- numeral (Num.Bit1 w)) = -1" | |
| 497 | by simp_all | |
| 498 | ||
| 499 | lemma sbintrunc_Suc_numeral: | |
| 500 | "sbintrunc (Suc n) 1 = 1" | |
| 501 | "sbintrunc (Suc n) (numeral (Num.Bit0 w)) = sbintrunc n (numeral w) BIT False" | |
| 502 | "sbintrunc (Suc n) (numeral (Num.Bit1 w)) = sbintrunc n (numeral w) BIT True" | |
| 503 | "sbintrunc (Suc n) (- numeral (Num.Bit0 w)) = sbintrunc n (- numeral w) BIT False" | |
| 504 | "sbintrunc (Suc n) (- numeral (Num.Bit1 w)) = sbintrunc n (- numeral (w + Num.One)) BIT True" | |
| 505 | by simp_all | |
| 506 | ||
| 507 | lemma bin_sign_lem: "(bin_sign (sbintrunc n bin) = -1) = bin_nth bin n" | |
| 508 | apply (induct n arbitrary: bin) | |
| 509 | apply (case_tac bin rule: bin_exhaust, case_tac b, auto) | |
| 510 | done | |
| 511 | ||
| 512 | lemma nth_bintr: "bin_nth (bintrunc m w) n \<longleftrightarrow> n < m \<and> bin_nth w n" | |
| 513 | apply (induct n arbitrary: w m) | |
| 514 | apply (case_tac m, auto)[1] | |
| 515 | apply (case_tac m, auto)[1] | |
| 516 | done | |
| 517 | ||
| 518 | lemma nth_sbintr: "bin_nth (sbintrunc m w) n = (if n < m then bin_nth w n else bin_nth w m)" | |
| 519 | apply (induct n arbitrary: w m) | |
| 520 | apply (case_tac m) | |
| 521 | apply simp_all | |
| 522 | apply (case_tac m) | |
| 523 | apply simp_all | |
| 524 | done | |
| 525 | ||
| 526 | lemma bin_nth_Bit: "bin_nth (w BIT b) n \<longleftrightarrow> n = 0 \<and> b \<or> (\<exists>m. n = Suc m \<and> bin_nth w m)" | |
| 527 | by (cases n) auto | |
| 528 | ||
| 529 | lemma bin_nth_Bit0: | |
| 530 | "bin_nth (numeral (Num.Bit0 w)) n \<longleftrightarrow> | |
| 531 | (\<exists>m. n = Suc m \<and> bin_nth (numeral w) m)" | |
| 532 | using bin_nth_Bit [where w="numeral w" and b="False"] by simp | |
| 533 | ||
| 534 | lemma bin_nth_Bit1: | |
| 535 | "bin_nth (numeral (Num.Bit1 w)) n \<longleftrightarrow> | |
| 536 | n = 0 \<or> (\<exists>m. n = Suc m \<and> bin_nth (numeral w) m)" | |
| 537 | using bin_nth_Bit [where w="numeral w" and b="True"] by simp | |
| 538 | ||
| 539 | lemma bintrunc_bintrunc_l: "n \<le> m \<Longrightarrow> bintrunc m (bintrunc n w) = bintrunc n w" | |
| 540 | by (rule bin_eqI) (auto simp: nth_bintr) | |
| 541 | ||
| 542 | lemma sbintrunc_sbintrunc_l: "n \<le> m \<Longrightarrow> sbintrunc m (sbintrunc n w) = sbintrunc n w" | |
| 543 | by (rule bin_eqI) (auto simp: nth_sbintr) | |
| 544 | ||
| 545 | lemma bintrunc_bintrunc_ge: "n \<le> m \<Longrightarrow> bintrunc n (bintrunc m w) = bintrunc n w" | |
| 546 | by (rule bin_eqI) (auto simp: nth_bintr) | |
| 547 | ||
| 548 | lemma bintrunc_bintrunc_min [simp]: "bintrunc m (bintrunc n w) = bintrunc (min m n) w" | |
| 549 | by (rule bin_eqI) (auto simp: nth_bintr) | |
| 550 | ||
| 551 | lemma sbintrunc_sbintrunc_min [simp]: "sbintrunc m (sbintrunc n w) = sbintrunc (min m n) w" | |
| 552 | by (rule bin_eqI) (auto simp: nth_sbintr min.absorb1 min.absorb2) | |
| 553 | ||
| 554 | lemmas bintrunc_Pls = | |
| 555 | bintrunc.Suc [where bin="0", simplified bin_last_numeral_simps bin_rest_numeral_simps] | |
| 556 | ||
| 557 | lemmas bintrunc_Min [simp] = | |
| 558 | bintrunc.Suc [where bin="-1", simplified bin_last_numeral_simps bin_rest_numeral_simps] | |
| 559 | ||
| 560 | lemmas bintrunc_BIT [simp] = | |
| 561 | bintrunc.Suc [where bin="w BIT b", simplified bin_last_BIT bin_rest_BIT] for w b | |
| 562 | ||
| 563 | lemmas bintrunc_Sucs = bintrunc_Pls bintrunc_Min bintrunc_BIT | |
| 564 | bintrunc_Suc_numeral | |
| 565 | ||
| 566 | lemmas sbintrunc_Suc_Pls = | |
| 567 | sbintrunc.Suc [where bin="0", simplified bin_last_numeral_simps bin_rest_numeral_simps] | |
| 568 | ||
| 569 | lemmas sbintrunc_Suc_Min = | |
| 570 | sbintrunc.Suc [where bin="-1", simplified bin_last_numeral_simps bin_rest_numeral_simps] | |
| 571 | ||
| 572 | lemmas sbintrunc_Suc_BIT [simp] = | |
| 573 | sbintrunc.Suc [where bin="w BIT b", simplified bin_last_BIT bin_rest_BIT] for w b | |
| 574 | ||
| 575 | lemmas sbintrunc_Sucs = sbintrunc_Suc_Pls sbintrunc_Suc_Min sbintrunc_Suc_BIT | |
| 576 | sbintrunc_Suc_numeral | |
| 577 | ||
| 578 | lemmas sbintrunc_Pls = | |
| 579 | sbintrunc.Z [where bin="0", simplified bin_last_numeral_simps bin_rest_numeral_simps] | |
| 580 | ||
| 581 | lemmas sbintrunc_Min = | |
| 582 | sbintrunc.Z [where bin="-1", simplified bin_last_numeral_simps bin_rest_numeral_simps] | |
| 583 | ||
| 584 | lemmas sbintrunc_0_BIT_B0 [simp] = | |
| 585 | sbintrunc.Z [where bin="w BIT False", simplified bin_last_numeral_simps bin_rest_numeral_simps] | |
| 586 | for w | |
| 587 | ||
| 588 | lemmas sbintrunc_0_BIT_B1 [simp] = | |
| 589 | sbintrunc.Z [where bin="w BIT True", simplified bin_last_BIT bin_rest_numeral_simps] | |
| 590 | for w | |
| 591 | ||
| 592 | lemmas sbintrunc_0_simps = | |
| 593 | sbintrunc_Pls sbintrunc_Min sbintrunc_0_BIT_B0 sbintrunc_0_BIT_B1 | |
| 594 | ||
| 595 | lemmas bintrunc_simps = bintrunc.Z bintrunc_Sucs | |
| 596 | lemmas sbintrunc_simps = sbintrunc_0_simps sbintrunc_Sucs | |
| 597 | ||
| 598 | lemma bintrunc_minus: "0 < n \<Longrightarrow> bintrunc (Suc (n - 1)) w = bintrunc n w" | |
| 599 | by auto | |
| 600 | ||
| 601 | lemma sbintrunc_minus: "0 < n \<Longrightarrow> sbintrunc (Suc (n - 1)) w = sbintrunc n w" | |
| 602 | by auto | |
| 603 | ||
| 604 | lemmas bintrunc_minus_simps = | |
| 605 | bintrunc_Sucs [THEN [2] bintrunc_minus [symmetric, THEN trans]] | |
| 606 | lemmas sbintrunc_minus_simps = | |
| 607 | sbintrunc_Sucs [THEN [2] sbintrunc_minus [symmetric, THEN trans]] | |
| 608 | ||
| 609 | lemmas thobini1 = arg_cong [where f = "\<lambda>w. w BIT b"] for b | |
| 610 | ||
| 611 | lemmas bintrunc_BIT_I = trans [OF bintrunc_BIT thobini1] | |
| 612 | lemmas bintrunc_Min_I = trans [OF bintrunc_Min thobini1] | |
| 613 | ||
| 614 | lemmas bmsts = bintrunc_minus_simps(1-3) [THEN thobini1 [THEN [2] trans]] | |
| 615 | lemmas bintrunc_Pls_minus_I = bmsts(1) | |
| 616 | lemmas bintrunc_Min_minus_I = bmsts(2) | |
| 617 | lemmas bintrunc_BIT_minus_I = bmsts(3) | |
| 618 | ||
| 619 | lemma bintrunc_Suc_lem: "bintrunc (Suc n) x = y \<Longrightarrow> m = Suc n \<Longrightarrow> bintrunc m x = y" | |
| 620 | by auto | |
| 621 | ||
| 622 | lemmas bintrunc_Suc_Ialts = | |
| 623 | bintrunc_Min_I [THEN bintrunc_Suc_lem] | |
| 624 | bintrunc_BIT_I [THEN bintrunc_Suc_lem] | |
| 625 | ||
| 626 | lemmas sbintrunc_BIT_I = trans [OF sbintrunc_Suc_BIT thobini1] | |
| 627 | ||
| 628 | lemmas sbintrunc_Suc_Is = | |
| 629 | sbintrunc_Sucs(1-3) [THEN thobini1 [THEN [2] trans]] | |
| 630 | ||
| 631 | lemmas sbintrunc_Suc_minus_Is = | |
| 632 | sbintrunc_minus_simps(1-3) [THEN thobini1 [THEN [2] trans]] | |
| 633 | ||
| 634 | lemma sbintrunc_Suc_lem: "sbintrunc (Suc n) x = y \<Longrightarrow> m = Suc n \<Longrightarrow> sbintrunc m x = y" | |
| 635 | by auto | |
| 636 | ||
| 637 | lemmas sbintrunc_Suc_Ialts = | |
| 638 | sbintrunc_Suc_Is [THEN sbintrunc_Suc_lem] | |
| 639 | ||
| 640 | lemma sbintrunc_bintrunc_lt: "m > n \<Longrightarrow> sbintrunc n (bintrunc m w) = sbintrunc n w" | |
| 641 | by (rule bin_eqI) (auto simp: nth_sbintr nth_bintr) | |
| 642 | ||
| 643 | lemma bintrunc_sbintrunc_le: "m \<le> Suc n \<Longrightarrow> bintrunc m (sbintrunc n w) = bintrunc m w" | |
| 644 | apply (rule bin_eqI) | |
| 645 | using le_Suc_eq less_Suc_eq_le apply (auto simp: nth_sbintr nth_bintr) | |
| 646 | done | |
| 647 | ||
| 648 | lemmas bintrunc_sbintrunc [simp] = order_refl [THEN bintrunc_sbintrunc_le] | |
| 649 | lemmas sbintrunc_bintrunc [simp] = lessI [THEN sbintrunc_bintrunc_lt] | |
| 650 | lemmas bintrunc_bintrunc [simp] = order_refl [THEN bintrunc_bintrunc_l] | |
| 651 | lemmas sbintrunc_sbintrunc [simp] = order_refl [THEN sbintrunc_sbintrunc_l] | |
| 652 | ||
| 653 | lemma bintrunc_sbintrunc' [simp]: "0 < n \<Longrightarrow> bintrunc n (sbintrunc (n - 1) w) = bintrunc n w" | |
| 654 | by (cases n) (auto simp del: bintrunc.Suc) | |
| 655 | ||
| 656 | lemma sbintrunc_bintrunc' [simp]: "0 < n \<Longrightarrow> sbintrunc (n - 1) (bintrunc n w) = sbintrunc (n - 1) w" | |
| 657 | by (cases n) (auto simp del: bintrunc.Suc) | |
| 658 | ||
| 659 | lemma bin_sbin_eq_iff: "bintrunc (Suc n) x = bintrunc (Suc n) y \<longleftrightarrow> sbintrunc n x = sbintrunc n y" | |
| 660 | apply (rule iffI) | |
| 661 | apply (rule box_equals [OF _ sbintrunc_bintrunc sbintrunc_bintrunc]) | |
| 662 | apply simp | |
| 663 | apply (rule box_equals [OF _ bintrunc_sbintrunc bintrunc_sbintrunc]) | |
| 664 | apply simp | |
| 665 | done | |
| 666 | ||
| 667 | lemma bin_sbin_eq_iff': | |
| 668 | "0 < n \<Longrightarrow> bintrunc n x = bintrunc n y \<longleftrightarrow> sbintrunc (n - 1) x = sbintrunc (n - 1) y" | |
| 669 | by (cases n) (simp_all add: bin_sbin_eq_iff del: bintrunc.Suc) | |
| 670 | ||
| 671 | lemmas bintrunc_sbintruncS0 [simp] = bintrunc_sbintrunc' [unfolded One_nat_def] | |
| 672 | lemmas sbintrunc_bintruncS0 [simp] = sbintrunc_bintrunc' [unfolded One_nat_def] | |
| 673 | ||
| 674 | lemmas bintrunc_bintrunc_l' = le_add1 [THEN bintrunc_bintrunc_l] | |
| 675 | lemmas sbintrunc_sbintrunc_l' = le_add1 [THEN sbintrunc_sbintrunc_l] | |
| 676 | ||
| 677 | (* although bintrunc_minus_simps, if added to default simpset, | |
| 678 | tends to get applied where it's not wanted in developing the theories, | |
| 679 | we get a version for when the word length is given literally *) | |
| 680 | ||
| 681 | lemmas nat_non0_gr = | |
| 682 | trans [OF iszero_def [THEN Not_eq_iff [THEN iffD2]] refl] | |
| 683 | ||
| 684 | lemma bintrunc_numeral: | |
| 685 | "bintrunc (numeral k) x = bintrunc (pred_numeral k) (bin_rest x) BIT bin_last x" | |
| 686 | by (simp add: numeral_eq_Suc) | |
| 687 | ||
| 688 | lemma sbintrunc_numeral: | |
| 689 | "sbintrunc (numeral k) x = sbintrunc (pred_numeral k) (bin_rest x) BIT bin_last x" | |
| 690 | by (simp add: numeral_eq_Suc) | |
| 691 | ||
| 692 | lemma bintrunc_numeral_simps [simp]: | |
| 693 | "bintrunc (numeral k) (numeral (Num.Bit0 w)) = bintrunc (pred_numeral k) (numeral w) BIT False" | |
| 694 | "bintrunc (numeral k) (numeral (Num.Bit1 w)) = bintrunc (pred_numeral k) (numeral w) BIT True" | |
| 695 | "bintrunc (numeral k) (- numeral (Num.Bit0 w)) = bintrunc (pred_numeral k) (- numeral w) BIT False" | |
| 696 | "bintrunc (numeral k) (- numeral (Num.Bit1 w)) = | |
| 697 | bintrunc (pred_numeral k) (- numeral (w + Num.One)) BIT True" | |
| 698 | "bintrunc (numeral k) 1 = 1" | |
| 699 | by (simp_all add: bintrunc_numeral) | |
| 700 | ||
| 701 | lemma sbintrunc_numeral_simps [simp]: | |
| 702 | "sbintrunc (numeral k) (numeral (Num.Bit0 w)) = sbintrunc (pred_numeral k) (numeral w) BIT False" | |
| 703 | "sbintrunc (numeral k) (numeral (Num.Bit1 w)) = sbintrunc (pred_numeral k) (numeral w) BIT True" | |
| 704 | "sbintrunc (numeral k) (- numeral (Num.Bit0 w)) = | |
| 705 | sbintrunc (pred_numeral k) (- numeral w) BIT False" | |
| 706 | "sbintrunc (numeral k) (- numeral (Num.Bit1 w)) = | |
| 707 | sbintrunc (pred_numeral k) (- numeral (w + Num.One)) BIT True" | |
| 708 | "sbintrunc (numeral k) 1 = 1" | |
| 709 | by (simp_all add: sbintrunc_numeral) | |
| 710 | ||
| 711 | lemma no_bintr_alt1: "bintrunc n = (\<lambda>w. w mod 2 ^ n :: int)" | |
| 712 | by (rule ext) (rule bintrunc_mod2p) | |
| 713 | ||
| 714 | lemma range_bintrunc: "range (bintrunc n) = {i. 0 \<le> i \<and> i < 2 ^ n}"
 | |
| 715 | apply (unfold no_bintr_alt1) | |
| 716 | apply (auto simp add: image_iff) | |
| 717 | apply (rule exI) | |
| 718 | apply (rule sym) | |
| 719 | using int_mod_lem [symmetric, of "2 ^ n"] | |
| 720 | apply auto | |
| 721 | done | |
| 722 | ||
| 723 | lemma no_sbintr_alt2: "sbintrunc n = (\<lambda>w. (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n :: int)" | |
| 724 | by (rule ext) (simp add : sbintrunc_mod2p) | |
| 725 | ||
| 726 | lemma range_sbintrunc: "range (sbintrunc n) = {i. - (2 ^ n) \<le> i \<and> i < 2 ^ n}"
 | |
| 727 | apply (unfold no_sbintr_alt2) | |
| 728 | apply (auto simp add: image_iff eq_diff_eq) | |
| 729 | ||
| 730 | apply (rule exI) | |
| 731 | apply (auto intro: int_mod_lem [THEN iffD1, symmetric]) | |
| 732 | done | |
| 733 | ||
| 734 | lemma sb_inc_lem: "a + 2^k < 0 \<Longrightarrow> a + 2^k + 2^(Suc k) \<le> (a + 2^k) mod 2^(Suc k)" | |
| 735 | for a :: int | |
| 736 | using int_mod_ge' [where n = "2 ^ (Suc k)" and b = "a + 2 ^ k"] | |
| 737 | by simp | |
| 738 | ||
| 739 | lemma sb_inc_lem': "a < - (2^k) \<Longrightarrow> a + 2^k + 2^(Suc k) \<le> (a + 2^k) mod 2^(Suc k)" | |
| 740 | for a :: int | |
| 741 | by (rule sb_inc_lem) simp | |
| 742 | ||
| 743 | lemma sbintrunc_inc: "x < - (2^n) \<Longrightarrow> x + 2^(Suc n) \<le> sbintrunc n x" | |
| 744 | unfolding no_sbintr_alt2 by (drule sb_inc_lem') simp | |
| 745 | ||
| 746 | lemma sb_dec_lem: "0 \<le> - (2 ^ k) + a \<Longrightarrow> (a + 2 ^ k) mod (2 * 2 ^ k) \<le> - (2 ^ k) + a" | |
| 747 | for a :: int | |
| 748 | using int_mod_le'[where n = "2 ^ (Suc k)" and b = "a + 2 ^ k"] by simp | |
| 749 | ||
| 750 | lemma sb_dec_lem': "2 ^ k \<le> a \<Longrightarrow> (a + 2 ^ k) mod (2 * 2 ^ k) \<le> - (2 ^ k) + a" | |
| 751 | for a :: int | |
| 752 | by (rule sb_dec_lem) simp | |
| 753 | ||
| 754 | lemma sbintrunc_dec: "x \<ge> (2 ^ n) \<Longrightarrow> x - 2 ^ (Suc n) >= sbintrunc n x" | |
| 755 | unfolding no_sbintr_alt2 by (drule sb_dec_lem') simp | |
| 756 | ||
| 757 | lemma bintr_ge0: "0 \<le> bintrunc n w" | |
| 758 | by (simp add: bintrunc_mod2p) | |
| 759 | ||
| 760 | lemma bintr_lt2p: "bintrunc n w < 2 ^ n" | |
| 761 | by (simp add: bintrunc_mod2p) | |
| 762 | ||
| 763 | lemma bintr_Min: "bintrunc n (- 1) = 2 ^ n - 1" | |
| 764 | by (simp add: bintrunc_mod2p m1mod2k) | |
| 765 | ||
| 766 | lemma sbintr_ge: "- (2 ^ n) \<le> sbintrunc n w" | |
| 767 | by (simp add: sbintrunc_mod2p) | |
| 768 | ||
| 769 | lemma sbintr_lt: "sbintrunc n w < 2 ^ n" | |
| 770 | by (simp add: sbintrunc_mod2p) | |
| 771 | ||
| 772 | lemma sign_Pls_ge_0: "bin_sign bin = 0 \<longleftrightarrow> bin \<ge> 0" | |
| 773 | for bin :: int | |
| 774 | by (simp add: bin_sign_def) | |
| 775 | ||
| 776 | lemma sign_Min_lt_0: "bin_sign bin = -1 \<longleftrightarrow> bin < 0" | |
| 777 | for bin :: int | |
| 778 | by (simp add: bin_sign_def) | |
| 779 | ||
| 780 | lemma bin_rest_trunc: "bin_rest (bintrunc n bin) = bintrunc (n - 1) (bin_rest bin)" | |
| 781 | by (induct n arbitrary: bin) auto | |
| 782 | ||
| 783 | lemma bin_rest_power_trunc: | |
| 784 | "(bin_rest ^^ k) (bintrunc n bin) = bintrunc (n - k) ((bin_rest ^^ k) bin)" | |
| 785 | by (induct k) (auto simp: bin_rest_trunc) | |
| 786 | ||
| 787 | lemma bin_rest_trunc_i: "bintrunc n (bin_rest bin) = bin_rest (bintrunc (Suc n) bin)" | |
| 788 | by auto | |
| 789 | ||
| 790 | lemma bin_rest_strunc: "bin_rest (sbintrunc (Suc n) bin) = sbintrunc n (bin_rest bin)" | |
| 791 | by (induct n arbitrary: bin) auto | |
| 792 | ||
| 793 | lemma bintrunc_rest [simp]: "bintrunc n (bin_rest (bintrunc n bin)) = bin_rest (bintrunc n bin)" | |
| 794 | apply (induct n arbitrary: bin) | |
| 795 | apply simp | |
| 796 | apply (case_tac bin rule: bin_exhaust) | |
| 797 | apply (auto simp: bintrunc_bintrunc_l) | |
| 798 | done | |
| 799 | ||
| 800 | lemma sbintrunc_rest [simp]: "sbintrunc n (bin_rest (sbintrunc n bin)) = bin_rest (sbintrunc n bin)" | |
| 801 | apply (induct n arbitrary: bin) | |
| 802 | apply simp | |
| 803 | apply (case_tac bin rule: bin_exhaust) | |
| 804 | apply (auto simp: bintrunc_bintrunc_l split: bool.splits) | |
| 805 | done | |
| 806 | ||
| 807 | lemma bintrunc_rest': "bintrunc n \<circ> bin_rest \<circ> bintrunc n = bin_rest \<circ> bintrunc n" | |
| 808 | by (rule ext) auto | |
| 809 | ||
| 810 | lemma sbintrunc_rest': "sbintrunc n \<circ> bin_rest \<circ> sbintrunc n = bin_rest \<circ> sbintrunc n" | |
| 811 | by (rule ext) auto | |
| 812 | ||
| 813 | lemma rco_lem: "f \<circ> g \<circ> f = g \<circ> f \<Longrightarrow> f \<circ> (g \<circ> f) ^^ n = g ^^ n \<circ> f" | |
| 814 | apply (rule ext) | |
| 815 | apply (induct_tac n) | |
| 816 | apply (simp_all (no_asm)) | |
| 817 | apply (drule fun_cong) | |
| 818 | apply (unfold o_def) | |
| 819 | apply (erule trans) | |
| 820 | apply simp | |
| 821 | done | |
| 822 | ||
| 823 | lemmas rco_bintr = bintrunc_rest' | |
| 824 | [THEN rco_lem [THEN fun_cong], unfolded o_def] | |
| 825 | lemmas rco_sbintr = sbintrunc_rest' | |
| 826 | [THEN rco_lem [THEN fun_cong], unfolded o_def] | |
| 827 | ||
| 828 | ||
| 829 | subsection \<open>Splitting and concatenation\<close> | |
| 830 | ||
| 831 | primrec bin_split :: "nat \<Rightarrow> int \<Rightarrow> int \<times> int" | |
| 832 | where | |
| 833 | Z: "bin_split 0 w = (w, 0)" | |
| 834 | | Suc: "bin_split (Suc n) w = | |
| 835 | (let (w1, w2) = bin_split n (bin_rest w) | |
| 836 | in (w1, w2 BIT bin_last w))" | |
| 837 | ||
| 838 | lemma [code]: | |
| 839 | "bin_split (Suc n) w = (let (w1, w2) = bin_split n (bin_rest w) in (w1, w2 BIT bin_last w))" | |
| 840 | "bin_split 0 w = (w, 0)" | |
| 841 | by simp_all | |
| 842 | ||
| 843 | primrec bin_cat :: "int \<Rightarrow> nat \<Rightarrow> int \<Rightarrow> int" | |
| 844 | where | |
| 845 | Z: "bin_cat w 0 v = w" | |
| 846 | | Suc: "bin_cat w (Suc n) v = bin_cat w n (bin_rest v) BIT bin_last v" | |
| 847 | ||
| 848 | lemma bin_sign_cat: "bin_sign (bin_cat x n y) = bin_sign x" | |
| 849 | by (induct n arbitrary: y) auto | |
| 850 | ||
| 851 | lemma bin_cat_Suc_Bit: "bin_cat w (Suc n) (v BIT b) = bin_cat w n v BIT b" | |
| 852 | by auto | |
| 853 | ||
| 854 | lemma bin_cat_assoc: "bin_cat (bin_cat x m y) n z = bin_cat x (m + n) (bin_cat y n z)" | |
| 855 | by (induct n arbitrary: z) auto | |
| 856 | ||
| 857 | lemma bin_cat_assoc_sym: "bin_cat x m (bin_cat y n z) = bin_cat (bin_cat x (m - n) y) (min m n) z" | |
| 858 | apply (induct n arbitrary: z m) | |
| 859 | apply clarsimp | |
| 860 | apply (case_tac m, auto) | |
| 861 | done | |
| 862 | ||
| 863 | definition bin_rcat :: "nat \<Rightarrow> int list \<Rightarrow> int" | |
| 864 | where "bin_rcat n = foldl (\<lambda>u v. bin_cat u n v) 0" | |
| 865 | ||
| 866 | fun bin_rsplit_aux :: "nat \<Rightarrow> nat \<Rightarrow> int \<Rightarrow> int list \<Rightarrow> int list" | |
| 867 | where "bin_rsplit_aux n m c bs = | |
| 868 | (if m = 0 \<or> n = 0 then bs | |
| 869 | else | |
| 870 | let (a, b) = bin_split n c | |
| 871 | in bin_rsplit_aux n (m - n) a (b # bs))" | |
| 872 | ||
| 873 | definition bin_rsplit :: "nat \<Rightarrow> nat \<times> int \<Rightarrow> int list" | |
| 874 | where "bin_rsplit n w = bin_rsplit_aux n (fst w) (snd w) []" | |
| 875 | ||
| 876 | fun bin_rsplitl_aux :: "nat \<Rightarrow> nat \<Rightarrow> int \<Rightarrow> int list \<Rightarrow> int list" | |
| 877 | where "bin_rsplitl_aux n m c bs = | |
| 878 | (if m = 0 \<or> n = 0 then bs | |
| 879 | else | |
| 880 | let (a, b) = bin_split (min m n) c | |
| 881 | in bin_rsplitl_aux n (m - n) a (b # bs))" | |
| 882 | ||
| 883 | definition bin_rsplitl :: "nat \<Rightarrow> nat \<times> int \<Rightarrow> int list" | |
| 884 | where "bin_rsplitl n w = bin_rsplitl_aux n (fst w) (snd w) []" | |
| 885 | ||
| 886 | declare bin_rsplit_aux.simps [simp del] | |
| 887 | declare bin_rsplitl_aux.simps [simp del] | |
| 888 | ||
| 889 | lemma bin_nth_cat: | |
| 890 | "bin_nth (bin_cat x k y) n = | |
| 891 | (if n < k then bin_nth y n else bin_nth x (n - k))" | |
| 892 | apply (induct k arbitrary: n y) | |
| 893 | apply clarsimp | |
| 894 | apply (case_tac n, auto) | |
| 895 | done | |
| 896 | ||
| 897 | lemma bin_nth_split: | |
| 898 | "bin_split n c = (a, b) \<Longrightarrow> | |
| 899 | (\<forall>k. bin_nth a k = bin_nth c (n + k)) \<and> | |
| 900 | (\<forall>k. bin_nth b k = (k < n \<and> bin_nth c k))" | |
| 901 | apply (induct n arbitrary: b c) | |
| 902 | apply clarsimp | |
| 903 | apply (clarsimp simp: Let_def split: prod.split_asm) | |
| 904 | apply (case_tac k) | |
| 905 | apply auto | |
| 906 | done | |
| 907 | ||
| 908 | lemma bin_cat_zero [simp]: "bin_cat 0 n w = bintrunc n w" | |
| 909 | by (induct n arbitrary: w) auto | |
| 910 | ||
| 911 | lemma bintr_cat1: "bintrunc (k + n) (bin_cat a n b) = bin_cat (bintrunc k a) n b" | |
| 912 | by (induct n arbitrary: b) auto | |
| 913 | ||
| 914 | lemma bintr_cat: "bintrunc m (bin_cat a n b) = | |
| 915 | bin_cat (bintrunc (m - n) a) n (bintrunc (min m n) b)" | |
| 916 | by (rule bin_eqI) (auto simp: bin_nth_cat nth_bintr) | |
| 917 | ||
| 918 | lemma bintr_cat_same [simp]: "bintrunc n (bin_cat a n b) = bintrunc n b" | |
| 919 | by (auto simp add : bintr_cat) | |
| 920 | ||
| 921 | lemma cat_bintr [simp]: "bin_cat a n (bintrunc n b) = bin_cat a n b" | |
| 922 | by (induct n arbitrary: b) auto | |
| 923 | ||
| 924 | lemma split_bintrunc: "bin_split n c = (a, b) \<Longrightarrow> b = bintrunc n c" | |
| 925 | by (induct n arbitrary: b c) (auto simp: Let_def split: prod.split_asm) | |
| 926 | ||
| 927 | lemma bin_cat_split: "bin_split n w = (u, v) \<Longrightarrow> w = bin_cat u n v" | |
| 928 | by (induct n arbitrary: v w) (auto simp: Let_def split: prod.split_asm) | |
| 929 | ||
| 930 | lemma bin_split_cat: "bin_split n (bin_cat v n w) = (v, bintrunc n w)" | |
| 931 | by (induct n arbitrary: w) auto | |
| 932 | ||
| 933 | lemma bin_split_zero [simp]: "bin_split n 0 = (0, 0)" | |
| 934 | by (induct n) auto | |
| 935 | ||
| 936 | lemma bin_split_minus1 [simp]: | |
| 937 | "bin_split n (- 1) = (- 1, bintrunc n (- 1))" | |
| 938 | by (induct n) auto | |
| 939 | ||
| 940 | lemma bin_split_trunc: | |
| 941 | "bin_split (min m n) c = (a, b) \<Longrightarrow> | |
| 942 | bin_split n (bintrunc m c) = (bintrunc (m - n) a, b)" | |
| 943 | apply (induct n arbitrary: m b c, clarsimp) | |
| 944 | apply (simp add: bin_rest_trunc Let_def split: prod.split_asm) | |
| 945 | apply (case_tac m) | |
| 946 | apply (auto simp: Let_def split: prod.split_asm) | |
| 947 | done | |
| 948 | ||
| 949 | lemma bin_split_trunc1: | |
| 950 | "bin_split n c = (a, b) \<Longrightarrow> | |
| 951 | bin_split n (bintrunc m c) = (bintrunc (m - n) a, bintrunc m b)" | |
| 952 | apply (induct n arbitrary: m b c, clarsimp) | |
| 953 | apply (simp add: bin_rest_trunc Let_def split: prod.split_asm) | |
| 954 | apply (case_tac m) | |
| 955 | apply (auto simp: Let_def split: prod.split_asm) | |
| 956 | done | |
| 957 | ||
| 958 | lemma bin_cat_num: "bin_cat a n b = a * 2 ^ n + bintrunc n b" | |
| 959 | apply (induct n arbitrary: b) | |
| 960 | apply clarsimp | |
| 961 | apply (simp add: Bit_def) | |
| 962 | done | |
| 963 | ||
| 964 | lemma bin_split_num: "bin_split n b = (b div 2 ^ n, b mod 2 ^ n)" | |
| 965 | apply (induct n arbitrary: b) | |
| 966 | apply simp | |
| 967 | apply (simp add: bin_rest_def zdiv_zmult2_eq) | |
| 968 | apply (case_tac b rule: bin_exhaust) | |
| 969 | apply simp | |
| 970 | apply (simp add: Bit_def mod_mult_mult1 pos_zmod_mult_2 add.commute) | |
| 971 | done | |
| 972 | ||
| 973 | lemmas bin_rsplit_aux_simps = bin_rsplit_aux.simps bin_rsplitl_aux.simps | |
| 974 | lemmas rsplit_aux_simps = bin_rsplit_aux_simps | |
| 975 | ||
| 976 | lemmas th_if_simp1 = if_split [where P = "(=) l", THEN iffD1, THEN conjunct1, THEN mp] for l | |
| 977 | lemmas th_if_simp2 = if_split [where P = "(=) l", THEN iffD1, THEN conjunct2, THEN mp] for l | |
| 978 | ||
| 979 | lemmas rsplit_aux_simp1s = rsplit_aux_simps [THEN th_if_simp1] | |
| 980 | ||
| 981 | lemmas rsplit_aux_simp2ls = rsplit_aux_simps [THEN th_if_simp2] | |
| 982 | \<comment> \<open>these safe to \<open>[simp add]\<close> as require calculating \<open>m - n\<close>\<close> | |
| 983 | lemmas bin_rsplit_aux_simp2s [simp] = rsplit_aux_simp2ls [unfolded Let_def] | |
| 984 | lemmas rbscl = bin_rsplit_aux_simp2s (2) | |
| 985 | ||
| 986 | lemmas rsplit_aux_0_simps [simp] = | |
| 987 | rsplit_aux_simp1s [OF disjI1] rsplit_aux_simp1s [OF disjI2] | |
| 988 | ||
| 989 | lemma bin_rsplit_aux_append: "bin_rsplit_aux n m c (bs @ cs) = bin_rsplit_aux n m c bs @ cs" | |
| 990 | apply (induct n m c bs rule: bin_rsplit_aux.induct) | |
| 991 | apply (subst bin_rsplit_aux.simps) | |
| 992 | apply (subst bin_rsplit_aux.simps) | |
| 993 | apply (clarsimp split: prod.split) | |
| 994 | done | |
| 995 | ||
| 996 | lemma bin_rsplitl_aux_append: "bin_rsplitl_aux n m c (bs @ cs) = bin_rsplitl_aux n m c bs @ cs" | |
| 997 | apply (induct n m c bs rule: bin_rsplitl_aux.induct) | |
| 998 | apply (subst bin_rsplitl_aux.simps) | |
| 999 | apply (subst bin_rsplitl_aux.simps) | |
| 1000 | apply (clarsimp split: prod.split) | |
| 1001 | done | |
| 1002 | ||
| 1003 | lemmas rsplit_aux_apps [where bs = "[]"] = | |
| 1004 | bin_rsplit_aux_append bin_rsplitl_aux_append | |
| 1005 | ||
| 1006 | lemmas rsplit_def_auxs = bin_rsplit_def bin_rsplitl_def | |
| 1007 | ||
| 1008 | lemmas rsplit_aux_alts = rsplit_aux_apps | |
| 1009 | [unfolded append_Nil rsplit_def_auxs [symmetric]] | |
| 1010 | ||
| 1011 | lemma bin_split_minus: "0 < n \<Longrightarrow> bin_split (Suc (n - 1)) w = bin_split n w" | |
| 1012 | by auto | |
| 1013 | ||
| 1014 | lemmas bin_split_minus_simp = | |
| 1015 | bin_split.Suc [THEN [2] bin_split_minus [symmetric, THEN trans]] | |
| 1016 | ||
| 1017 | lemma bin_split_pred_simp [simp]: | |
| 1018 | "(0::nat) < numeral bin \<Longrightarrow> | |
| 1019 | bin_split (numeral bin) w = | |
| 1020 | (let (w1, w2) = bin_split (numeral bin - 1) (bin_rest w) | |
| 1021 | in (w1, w2 BIT bin_last w))" | |
| 1022 | by (simp only: bin_split_minus_simp) | |
| 1023 | ||
| 1024 | lemma bin_rsplit_aux_simp_alt: | |
| 1025 | "bin_rsplit_aux n m c bs = | |
| 1026 | (if m = 0 \<or> n = 0 then bs | |
| 1027 | else let (a, b) = bin_split n c in bin_rsplit n (m - n, a) @ b # bs)" | |
| 1028 | apply (simp add: bin_rsplit_aux.simps [of n m c bs]) | |
| 1029 | apply (subst rsplit_aux_alts) | |
| 1030 | apply (simp add: bin_rsplit_def) | |
| 1031 | done | |
| 1032 | ||
| 1033 | lemmas bin_rsplit_simp_alt = | |
| 1034 | trans [OF bin_rsplit_def bin_rsplit_aux_simp_alt] | |
| 1035 | ||
| 1036 | lemmas bthrs = bin_rsplit_simp_alt [THEN [2] trans] | |
| 1037 | ||
| 1038 | lemma bin_rsplit_size_sign' [rule_format]: | |
| 1039 | "n > 0 \<Longrightarrow> rev sw = bin_rsplit n (nw, w) \<Longrightarrow> \<forall>v\<in>set sw. bintrunc n v = v" | |
| 1040 | apply (induct sw arbitrary: nw w) | |
| 1041 | apply clarsimp | |
| 1042 | apply clarsimp | |
| 1043 | apply (drule bthrs) | |
| 1044 | apply (simp (no_asm_use) add: Let_def split: prod.split_asm if_split_asm) | |
| 1045 | apply clarify | |
| 1046 | apply (drule split_bintrunc) | |
| 1047 | apply simp | |
| 1048 | done | |
| 1049 | ||
| 1050 | lemmas bin_rsplit_size_sign = bin_rsplit_size_sign' [OF asm_rl | |
| 1051 | rev_rev_ident [THEN trans] set_rev [THEN equalityD2 [THEN subsetD]]] | |
| 1052 | ||
| 1053 | lemma bin_nth_rsplit [rule_format] : | |
| 1054 | "n > 0 \<Longrightarrow> m < n \<Longrightarrow> | |
| 1055 | \<forall>w k nw. | |
| 1056 | rev sw = bin_rsplit n (nw, w) \<longrightarrow> | |
| 1057 | k < size sw \<longrightarrow> bin_nth (sw ! k) m = bin_nth w (k * n + m)" | |
| 1058 | apply (induct sw) | |
| 1059 | apply clarsimp | |
| 1060 | apply clarsimp | |
| 1061 | apply (drule bthrs) | |
| 1062 | apply (simp (no_asm_use) add: Let_def split: prod.split_asm if_split_asm) | |
| 1063 | apply clarify | |
| 1064 | apply (erule allE, erule impE, erule exI) | |
| 1065 | apply (case_tac k) | |
| 1066 | apply clarsimp | |
| 1067 | prefer 2 | |
| 1068 | apply clarsimp | |
| 1069 | apply (erule allE) | |
| 1070 | apply (erule (1) impE) | |
| 1071 | apply (drule bin_nth_split, erule conjE, erule allE, erule trans, simp add: ac_simps)+ | |
| 1072 | done | |
| 1073 | ||
| 1074 | lemma bin_rsplit_all: "0 < nw \<Longrightarrow> nw \<le> n \<Longrightarrow> bin_rsplit n (nw, w) = [bintrunc n w]" | |
| 1075 | by (auto simp: bin_rsplit_def rsplit_aux_simp2ls split: prod.split dest!: split_bintrunc) | |
| 1076 | ||
| 1077 | lemma bin_rsplit_l [rule_format]: | |
| 1078 | "\<forall>bin. bin_rsplitl n (m, bin) = bin_rsplit n (m, bintrunc m bin)" | |
| 1079 | apply (rule_tac a = "m" in wf_less_than [THEN wf_induct]) | |
| 1080 | apply (simp (no_asm) add: bin_rsplitl_def bin_rsplit_def) | |
| 1081 | apply (rule allI) | |
| 1082 | apply (subst bin_rsplitl_aux.simps) | |
| 1083 | apply (subst bin_rsplit_aux.simps) | |
| 1084 | apply (clarsimp simp: Let_def split: prod.split) | |
| 1085 | apply (drule bin_split_trunc) | |
| 1086 | apply (drule sym [THEN trans], assumption) | |
| 1087 | apply (subst rsplit_aux_alts(1)) | |
| 1088 | apply (subst rsplit_aux_alts(2)) | |
| 1089 | apply clarsimp | |
| 1090 | unfolding bin_rsplit_def bin_rsplitl_def | |
| 1091 | apply simp | |
| 1092 | done | |
| 1093 | ||
| 1094 | lemma bin_rsplit_rcat [rule_format]: | |
| 1095 | "n > 0 \<longrightarrow> bin_rsplit n (n * size ws, bin_rcat n ws) = map (bintrunc n) ws" | |
| 1096 | apply (unfold bin_rsplit_def bin_rcat_def) | |
| 1097 | apply (rule_tac xs = ws in rev_induct) | |
| 1098 | apply clarsimp | |
| 1099 | apply clarsimp | |
| 1100 | apply (subst rsplit_aux_alts) | |
| 1101 | unfolding bin_split_cat | |
| 1102 | apply simp | |
| 1103 | done | |
| 1104 | ||
| 1105 | lemma bin_rsplit_aux_len_le [rule_format] : | |
| 1106 | "\<forall>ws m. n \<noteq> 0 \<longrightarrow> ws = bin_rsplit_aux n nw w bs \<longrightarrow> | |
| 1107 | length ws \<le> m \<longleftrightarrow> nw + length bs * n \<le> m * n" | |
| 1108 | proof - | |
| 1109 | have *: R | |
| 1110 | if d: "i \<le> j \<or> m < j'" | |
| 1111 | and R1: "i * k \<le> j * k \<Longrightarrow> R" | |
| 1112 | and R2: "Suc m * k' \<le> j' * k' \<Longrightarrow> R" | |
| 1113 | for i j j' k k' m :: nat and R | |
| 1114 | using d | |
| 1115 | apply safe | |
| 1116 | apply (rule R1, erule mult_le_mono1) | |
| 1117 | apply (rule R2, erule Suc_le_eq [THEN iffD2 [THEN mult_le_mono1]]) | |
| 1118 | done | |
| 1119 | have **: "0 < sc \<Longrightarrow> sc - n + (n + lb * n) \<le> m * n \<longleftrightarrow> sc + lb * n \<le> m * n" | |
| 1120 | for sc m n lb :: nat | |
| 1121 | apply safe | |
| 1122 | apply arith | |
| 1123 | apply (case_tac "sc \<ge> n") | |
| 1124 | apply arith | |
| 1125 | apply (insert linorder_le_less_linear [of m lb]) | |
| 1126 | apply (erule_tac k=n and k'=n in *) | |
| 1127 | apply arith | |
| 1128 | apply simp | |
| 1129 | done | |
| 1130 | show ?thesis | |
| 1131 | apply (induct n nw w bs rule: bin_rsplit_aux.induct) | |
| 1132 | apply (subst bin_rsplit_aux.simps) | |
| 1133 | apply (simp add: ** Let_def split: prod.split) | |
| 1134 | done | |
| 1135 | qed | |
| 1136 | ||
| 1137 | lemma bin_rsplit_len_le: "n \<noteq> 0 \<longrightarrow> ws = bin_rsplit n (nw, w) \<longrightarrow> length ws \<le> m \<longleftrightarrow> nw \<le> m * n" | |
| 1138 | by (auto simp: bin_rsplit_def bin_rsplit_aux_len_le) | |
| 1139 | ||
| 1140 | lemma bin_rsplit_aux_len: | |
| 1141 | "n \<noteq> 0 \<Longrightarrow> length (bin_rsplit_aux n nw w cs) = (nw + n - 1) div n + length cs" | |
| 1142 | apply (induct n nw w cs rule: bin_rsplit_aux.induct) | |
| 1143 | apply (subst bin_rsplit_aux.simps) | |
| 1144 | apply (clarsimp simp: Let_def split: prod.split) | |
| 1145 | apply (erule thin_rl) | |
| 1146 | apply (case_tac m) | |
| 1147 | apply simp | |
| 1148 | apply (case_tac "m \<le> n") | |
| 1149 | apply (auto simp add: div_add_self2) | |
| 1150 | done | |
| 1151 | ||
| 1152 | lemma bin_rsplit_len: "n \<noteq> 0 \<Longrightarrow> length (bin_rsplit n (nw, w)) = (nw + n - 1) div n" | |
| 1153 | by (auto simp: bin_rsplit_def bin_rsplit_aux_len) | |
| 1154 | ||
| 1155 | lemma bin_rsplit_aux_len_indep: | |
| 1156 | "n \<noteq> 0 \<Longrightarrow> length bs = length cs \<Longrightarrow> | |
| 1157 | length (bin_rsplit_aux n nw v bs) = | |
| 1158 | length (bin_rsplit_aux n nw w cs)" | |
| 1159 | proof (induct n nw w cs arbitrary: v bs rule: bin_rsplit_aux.induct) | |
| 1160 | case (1 n m w cs v bs) | |
| 1161 | show ?case | |
| 1162 | proof (cases "m = 0") | |
| 1163 | case True | |
| 1164 | with \<open>length bs = length cs\<close> show ?thesis by simp | |
| 1165 | next | |
| 1166 | case False | |
| 1167 | from "1.hyps" \<open>m \<noteq> 0\<close> \<open>n \<noteq> 0\<close> | |
| 1168 | have hyp: "\<And>v bs. length bs = Suc (length cs) \<Longrightarrow> | |
| 1169 | length (bin_rsplit_aux n (m - n) v bs) = | |
| 1170 | length (bin_rsplit_aux n (m - n) (fst (bin_split n w)) (snd (bin_split n w) # cs))" | |
| 1171 | by auto | |
| 1172 | from \<open>length bs = length cs\<close> \<open>n \<noteq> 0\<close> show ?thesis | |
| 1173 | by (auto simp add: bin_rsplit_aux_simp_alt Let_def bin_rsplit_len split: prod.split) | |
| 1174 | qed | |
| 1175 | qed | |
| 1176 | ||
| 1177 | lemma bin_rsplit_len_indep: | |
| 1178 | "n \<noteq> 0 \<Longrightarrow> length (bin_rsplit n (nw, v)) = length (bin_rsplit n (nw, w))" | |
| 1179 | apply (unfold bin_rsplit_def) | |
| 1180 | apply (simp (no_asm)) | |
| 1181 | apply (erule bin_rsplit_aux_len_indep) | |
| 1182 | apply (rule refl) | |
| 1183 | done | |
| 1184 | ||
| 1185 | ||
| 61799 | 1186 | subsection \<open>Logical operations\<close> | 
| 24353 | 1187 | |
| 70191 | 1188 | primrec bin_sc :: "nat \<Rightarrow> bool \<Rightarrow> int \<Rightarrow> int" | 
| 1189 | where | |
| 1190 | Z: "bin_sc 0 b w = bin_rest w BIT b" | |
| 1191 | | Suc: "bin_sc (Suc n) b w = bin_sc n b (bin_rest w) BIT bin_last w" | |
| 1192 | ||
| 1193 | lemma bin_nth_sc [simp]: "bin_nth (bin_sc n b w) n \<longleftrightarrow> b" | |
| 1194 | by (induct n arbitrary: w) auto | |
| 1195 | ||
| 1196 | lemma bin_sc_sc_same [simp]: "bin_sc n c (bin_sc n b w) = bin_sc n c w" | |
| 1197 | by (induct n arbitrary: w) auto | |
| 1198 | ||
| 1199 | lemma bin_sc_sc_diff: "m \<noteq> n \<Longrightarrow> bin_sc m c (bin_sc n b w) = bin_sc n b (bin_sc m c w)" | |
| 1200 | apply (induct n arbitrary: w m) | |
| 1201 | apply (case_tac [!] m) | |
| 1202 | apply auto | |
| 1203 | done | |
| 1204 | ||
| 1205 | lemma bin_nth_sc_gen: "bin_nth (bin_sc n b w) m = (if m = n then b else bin_nth w m)" | |
| 1206 | by (induct n arbitrary: w m) (case_tac [!] m, auto) | |
| 1207 | ||
| 1208 | lemma bin_sc_nth [simp]: "bin_sc n (bin_nth w n) w = w" | |
| 1209 | by (induct n arbitrary: w) auto | |
| 1210 | ||
| 1211 | lemma bin_sign_sc [simp]: "bin_sign (bin_sc n b w) = bin_sign w" | |
| 1212 | by (induct n arbitrary: w) auto | |
| 1213 | ||
| 1214 | lemma bin_sc_bintr [simp]: "bintrunc m (bin_sc n x (bintrunc m (w))) = bintrunc m (bin_sc n x w)" | |
| 1215 | apply (induct n arbitrary: w m) | |
| 1216 | apply (case_tac [!] w rule: bin_exhaust) | |
| 1217 | apply (case_tac [!] m, auto) | |
| 1218 | done | |
| 1219 | ||
| 1220 | lemma bin_clr_le: "bin_sc n False w \<le> w" | |
| 1221 | apply (induct n arbitrary: w) | |
| 1222 | apply (case_tac [!] w rule: bin_exhaust) | |
| 1223 | apply (auto simp: le_Bits) | |
| 1224 | done | |
| 1225 | ||
| 1226 | lemma bin_set_ge: "bin_sc n True w \<ge> w" | |
| 1227 | apply (induct n arbitrary: w) | |
| 1228 | apply (case_tac [!] w rule: bin_exhaust) | |
| 1229 | apply (auto simp: le_Bits) | |
| 1230 | done | |
| 1231 | ||
| 1232 | lemma bintr_bin_clr_le: "bintrunc n (bin_sc m False w) \<le> bintrunc n w" | |
| 1233 | apply (induct n arbitrary: w m) | |
| 1234 | apply simp | |
| 1235 | apply (case_tac w rule: bin_exhaust) | |
| 1236 | apply (case_tac m) | |
| 1237 | apply (auto simp: le_Bits) | |
| 1238 | done | |
| 1239 | ||
| 1240 | lemma bintr_bin_set_ge: "bintrunc n (bin_sc m True w) \<ge> bintrunc n w" | |
| 1241 | apply (induct n arbitrary: w m) | |
| 1242 | apply simp | |
| 1243 | apply (case_tac w rule: bin_exhaust) | |
| 1244 | apply (case_tac m) | |
| 1245 | apply (auto simp: le_Bits) | |
| 1246 | done | |
| 1247 | ||
| 1248 | lemma bin_sc_FP [simp]: "bin_sc n False 0 = 0" | |
| 1249 | by (induct n) auto | |
| 1250 | ||
| 1251 | lemma bin_sc_TM [simp]: "bin_sc n True (- 1) = - 1" | |
| 1252 | by (induct n) auto | |
| 1253 | ||
| 1254 | lemmas bin_sc_simps = bin_sc.Z bin_sc.Suc bin_sc_TM bin_sc_FP | |
| 1255 | ||
| 1256 | lemma bin_sc_minus: "0 < n \<Longrightarrow> bin_sc (Suc (n - 1)) b w = bin_sc n b w" | |
| 1257 | by auto | |
| 1258 | ||
| 1259 | lemmas bin_sc_Suc_minus = | |
| 1260 | trans [OF bin_sc_minus [symmetric] bin_sc.Suc] | |
| 1261 | ||
| 1262 | lemma bin_sc_numeral [simp]: | |
| 1263 | "bin_sc (numeral k) b w = | |
| 1264 | bin_sc (pred_numeral k) b (bin_rest w) BIT bin_last w" | |
| 1265 | by (simp add: numeral_eq_Suc) | |
| 1266 | ||
| 1267 | instantiation int :: bit_operations | |
| 25762 | 1268 | begin | 
| 1269 | ||
| 65363 | 1270 | definition int_not_def: "bitNOT = (\<lambda>x::int. - x - 1)" | 
| 46019 | 1271 | |
| 65363 | 1272 | function bitAND_int | 
| 1273 | where "bitAND_int x y = | |
| 1274 | (if x = 0 then 0 else if x = -1 then y | |
| 1275 | else (bin_rest x AND bin_rest y) BIT (bin_last x \<and> bin_last y))" | |
| 46019 | 1276 | by pat_completeness simp | 
| 25762 | 1277 | |
| 46019 | 1278 | termination | 
| 65363 | 1279 | by (relation "measure (nat \<circ> abs \<circ> fst)", simp_all add: bin_rest_def) | 
| 46019 | 1280 | |
| 1281 | declare bitAND_int.simps [simp del] | |
| 25762 | 1282 | |
| 67120 | 1283 | definition int_or_def: "bitOR = (\<lambda>x y::int. NOT (NOT x AND NOT y))" | 
| 25762 | 1284 | |
| 67120 | 1285 | definition int_xor_def: "bitXOR = (\<lambda>x y::int. (x AND NOT y) OR (NOT x AND y))" | 
| 25762 | 1286 | |
| 70191 | 1287 | definition [iff]: "i !! n \<longleftrightarrow> bin_nth i n" | 
| 1288 | ||
| 1289 | definition "lsb i = i !! 0" for i :: int | |
| 1290 | ||
| 1291 | definition "set_bit i n b = bin_sc n b i" | |
| 1292 | ||
| 1293 | definition "shiftl x n = x * 2 ^ n" for x :: int | |
| 1294 | ||
| 1295 | definition "shiftr x n = x div 2 ^ n" for x :: int | |
| 1296 | ||
| 1297 | definition "msb x \<longleftrightarrow> x < 0" for x :: int | |
| 1298 | ||
| 25762 | 1299 | instance .. | 
| 1300 | ||
| 1301 | end | |
| 24353 | 1302 | |
| 70191 | 1303 | |
| 61799 | 1304 | subsubsection \<open>Basic simplification rules\<close> | 
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1305 | |
| 67120 | 1306 | lemma int_not_BIT [simp]: "NOT (w BIT b) = (NOT w) BIT (\<not> b)" | 
| 1307 | by (cases b) (simp_all add: int_not_def Bit_def) | |
| 46016 | 1308 | |
| 24333 | 1309 | lemma int_not_simps [simp]: | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1310 | "NOT (0::int) = -1" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1311 | "NOT (1::int) = -2" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54427diff
changeset | 1312 | "NOT (- 1::int) = 0" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54427diff
changeset | 1313 | "NOT (numeral w::int) = - numeral (w + Num.One)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54427diff
changeset | 1314 | "NOT (- numeral (Num.Bit0 w)::int) = numeral (Num.BitM w)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54427diff
changeset | 1315 | "NOT (- numeral (Num.Bit1 w)::int) = numeral (Num.Bit0 w)" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1316 | unfolding int_not_def by simp_all | 
| 24333 | 1317 | |
| 67120 | 1318 | lemma int_not_not [simp]: "NOT (NOT x) = x" | 
| 1319 | for x :: int | |
| 46017 | 1320 | unfolding int_not_def by simp | 
| 1321 | ||
| 67120 | 1322 | lemma int_and_0 [simp]: "0 AND x = 0" | 
| 1323 | for x :: int | |
| 46019 | 1324 | by (simp add: bitAND_int.simps) | 
| 1325 | ||
| 67120 | 1326 | lemma int_and_m1 [simp]: "-1 AND x = x" | 
| 1327 | for x :: int | |
| 46019 | 1328 | by (simp add: bitAND_int.simps) | 
| 1329 | ||
| 67120 | 1330 | lemma int_and_Bits [simp]: "(x BIT b) AND (y BIT c) = (x AND y) BIT (b \<and> c)" | 
| 1331 | by (subst bitAND_int.simps) (simp add: Bit_eq_0_iff Bit_eq_m1_iff) | |
| 46017 | 1332 | |
| 67120 | 1333 | lemma int_or_zero [simp]: "0 OR x = x" | 
| 1334 | for x :: int | |
| 1335 | by (simp add: int_or_def) | |
| 46018 | 1336 | |
| 67120 | 1337 | lemma int_or_minus1 [simp]: "-1 OR x = -1" | 
| 1338 | for x :: int | |
| 1339 | by (simp add: int_or_def) | |
| 46017 | 1340 | |
| 67120 | 1341 | lemma int_or_Bits [simp]: "(x BIT b) OR (y BIT c) = (x OR y) BIT (b \<or> c)" | 
| 1342 | by (simp add: int_or_def) | |
| 24333 | 1343 | |
| 67120 | 1344 | lemma int_xor_zero [simp]: "0 XOR x = x" | 
| 1345 | for x :: int | |
| 1346 | by (simp add: int_xor_def) | |
| 46018 | 1347 | |
| 67120 | 1348 | lemma int_xor_Bits [simp]: "(x BIT b) XOR (y BIT c) = (x XOR y) BIT ((b \<or> c) \<and> \<not> (b \<and> c))" | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1349 | unfolding int_xor_def by auto | 
| 46018 | 1350 | |
| 67120 | 1351 | |
| 61799 | 1352 | subsubsection \<open>Binary destructors\<close> | 
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1353 | |
| 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1354 | lemma bin_rest_NOT [simp]: "bin_rest (NOT x) = NOT (bin_rest x)" | 
| 67120 | 1355 | by (cases x rule: bin_exhaust) simp | 
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1356 | |
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1357 | lemma bin_last_NOT [simp]: "bin_last (NOT x) \<longleftrightarrow> \<not> bin_last x" | 
| 67120 | 1358 | by (cases x rule: bin_exhaust) simp | 
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1359 | |
| 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1360 | lemma bin_rest_AND [simp]: "bin_rest (x AND y) = bin_rest x AND bin_rest y" | 
| 67120 | 1361 | by (cases x rule: bin_exhaust, cases y rule: bin_exhaust) simp | 
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1362 | |
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1363 | lemma bin_last_AND [simp]: "bin_last (x AND y) \<longleftrightarrow> bin_last x \<and> bin_last y" | 
| 67120 | 1364 | by (cases x rule: bin_exhaust, cases y rule: bin_exhaust) simp | 
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1365 | |
| 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1366 | lemma bin_rest_OR [simp]: "bin_rest (x OR y) = bin_rest x OR bin_rest y" | 
| 67120 | 1367 | by (cases x rule: bin_exhaust, cases y rule: bin_exhaust) simp | 
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1368 | |
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1369 | lemma bin_last_OR [simp]: "bin_last (x OR y) \<longleftrightarrow> bin_last x \<or> bin_last y" | 
| 67120 | 1370 | by (cases x rule: bin_exhaust, cases y rule: bin_exhaust) simp | 
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1371 | |
| 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1372 | lemma bin_rest_XOR [simp]: "bin_rest (x XOR y) = bin_rest x XOR bin_rest y" | 
| 67120 | 1373 | by (cases x rule: bin_exhaust, cases y rule: bin_exhaust) simp | 
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1374 | |
| 67120 | 1375 | lemma bin_last_XOR [simp]: | 
| 1376 | "bin_last (x XOR y) \<longleftrightarrow> (bin_last x \<or> bin_last y) \<and> \<not> (bin_last x \<and> bin_last y)" | |
| 1377 | by (cases x rule: bin_exhaust, cases y rule: bin_exhaust) simp | |
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1378 | |
| 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1379 | lemma bin_nth_ops: | 
| 67120 | 1380 | "\<And>x y. bin_nth (x AND y) n \<longleftrightarrow> bin_nth x n \<and> bin_nth y n" | 
| 1381 | "\<And>x y. bin_nth (x OR y) n \<longleftrightarrow> bin_nth x n \<or> bin_nth y n" | |
| 1382 | "\<And>x y. bin_nth (x XOR y) n \<longleftrightarrow> bin_nth x n \<noteq> bin_nth y n" | |
| 1383 | "\<And>x. bin_nth (NOT x) n \<longleftrightarrow> \<not> bin_nth x n" | |
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1384 | by (induct n) auto | 
| 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1385 | |
| 67120 | 1386 | |
| 61799 | 1387 | subsubsection \<open>Derived properties\<close> | 
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1388 | |
| 67120 | 1389 | lemma int_xor_minus1 [simp]: "-1 XOR x = NOT x" | 
| 1390 | for x :: int | |
| 46018 | 1391 | by (auto simp add: bin_eq_iff bin_nth_ops) | 
| 1392 | ||
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1393 | lemma int_xor_extra_simps [simp]: | 
| 67120 | 1394 | "w XOR 0 = w" | 
| 1395 | "w XOR -1 = NOT w" | |
| 1396 | for w :: int | |
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1397 | by (auto simp add: bin_eq_iff bin_nth_ops) | 
| 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1398 | |
| 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1399 | lemma int_or_extra_simps [simp]: | 
| 67120 | 1400 | "w OR 0 = w" | 
| 1401 | "w OR -1 = -1" | |
| 1402 | for w :: int | |
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1403 | by (auto simp add: bin_eq_iff bin_nth_ops) | 
| 24333 | 1404 | |
| 37667 | 1405 | lemma int_and_extra_simps [simp]: | 
| 67120 | 1406 | "w AND 0 = 0" | 
| 1407 | "w AND -1 = w" | |
| 1408 | for w :: int | |
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1409 | by (auto simp add: bin_eq_iff bin_nth_ops) | 
| 24333 | 1410 | |
| 67120 | 1411 | text \<open>Commutativity of the above.\<close> | 
| 24333 | 1412 | lemma bin_ops_comm: | 
| 67120 | 1413 | fixes x y :: int | 
| 1414 | shows int_and_comm: "x AND y = y AND x" | |
| 1415 | and int_or_comm: "x OR y = y OR x" | |
| 1416 | and int_xor_comm: "x XOR y = y XOR x" | |
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1417 | by (auto simp add: bin_eq_iff bin_nth_ops) | 
| 24333 | 1418 | |
| 1419 | lemma bin_ops_same [simp]: | |
| 67120 | 1420 | "x AND x = x" | 
| 1421 | "x OR x = x" | |
| 1422 | "x XOR x = 0" | |
| 1423 | for x :: int | |
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1424 | by (auto simp add: bin_eq_iff bin_nth_ops) | 
| 24333 | 1425 | |
| 65363 | 1426 | lemmas bin_log_esimps = | 
| 24333 | 1427 | int_and_extra_simps int_or_extra_simps int_xor_extra_simps | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1428 | int_and_0 int_and_m1 int_or_zero int_or_minus1 int_xor_zero int_xor_minus1 | 
| 24333 | 1429 | |
| 67120 | 1430 | |
| 1431 | subsubsection \<open>Basic properties of logical (bit-wise) operations\<close> | |
| 24333 | 1432 | |
| 67120 | 1433 | lemma bbw_ao_absorb: "x AND (y OR x) = x \<and> x OR (y AND x) = x" | 
| 1434 | for x y :: int | |
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1435 | by (auto simp add: bin_eq_iff bin_nth_ops) | 
| 24333 | 1436 | |
| 1437 | lemma bbw_ao_absorbs_other: | |
| 67120 | 1438 | "x AND (x OR y) = x \<and> (y AND x) OR x = x" | 
| 1439 | "(y OR x) AND x = x \<and> x OR (x AND y) = x" | |
| 1440 | "(x OR y) AND x = x \<and> (x AND y) OR x = x" | |
| 1441 | for x y :: int | |
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1442 | by (auto simp add: bin_eq_iff bin_nth_ops) | 
| 24353 | 1443 | |
| 24333 | 1444 | lemmas bbw_ao_absorbs [simp] = bbw_ao_absorb bbw_ao_absorbs_other | 
| 1445 | ||
| 67120 | 1446 | lemma int_xor_not: "(NOT x) XOR y = NOT (x XOR y) \<and> x XOR (NOT y) = NOT (x XOR y)" | 
| 1447 | for x y :: int | |
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1448 | by (auto simp add: bin_eq_iff bin_nth_ops) | 
| 24333 | 1449 | |
| 67120 | 1450 | lemma int_and_assoc: "(x AND y) AND z = x AND (y AND z)" | 
| 1451 | for x y z :: int | |
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1452 | by (auto simp add: bin_eq_iff bin_nth_ops) | 
| 24333 | 1453 | |
| 67120 | 1454 | lemma int_or_assoc: "(x OR y) OR z = x OR (y OR z)" | 
| 1455 | for x y z :: int | |
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1456 | by (auto simp add: bin_eq_iff bin_nth_ops) | 
| 24333 | 1457 | |
| 67120 | 1458 | lemma int_xor_assoc: "(x XOR y) XOR z = x XOR (y XOR z)" | 
| 1459 | for x y z :: int | |
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1460 | by (auto simp add: bin_eq_iff bin_nth_ops) | 
| 24333 | 1461 | |
| 1462 | lemmas bbw_assocs = int_and_assoc int_or_assoc int_xor_assoc | |
| 1463 | ||
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1464 | (* BH: Why are these declared as simp rules??? *) | 
| 65363 | 1465 | lemma bbw_lcs [simp]: | 
| 67120 | 1466 | "y AND (x AND z) = x AND (y AND z)" | 
| 1467 | "y OR (x OR z) = x OR (y OR z)" | |
| 1468 | "y XOR (x XOR z) = x XOR (y XOR z)" | |
| 1469 | for x y :: int | |
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1470 | by (auto simp add: bin_eq_iff bin_nth_ops) | 
| 24333 | 1471 | |
| 65363 | 1472 | lemma bbw_not_dist: | 
| 67120 | 1473 | "NOT (x OR y) = (NOT x) AND (NOT y)" | 
| 1474 | "NOT (x AND y) = (NOT x) OR (NOT y)" | |
| 1475 | for x y :: int | |
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1476 | by (auto simp add: bin_eq_iff bin_nth_ops) | 
| 24333 | 1477 | |
| 67120 | 1478 | lemma bbw_oa_dist: "(x AND y) OR z = (x OR z) AND (y OR z)" | 
| 1479 | for x y z :: int | |
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1480 | by (auto simp add: bin_eq_iff bin_nth_ops) | 
| 24333 | 1481 | |
| 67120 | 1482 | lemma bbw_ao_dist: "(x OR y) AND z = (x AND z) OR (y AND z)" | 
| 1483 | for x y z :: int | |
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1484 | by (auto simp add: bin_eq_iff bin_nth_ops) | 
| 24333 | 1485 | |
| 24367 
3e29eafabe16
AC rules for bitwise logical operators no longer declared simp
 huffman parents: 
24366diff
changeset | 1486 | (* | 
| 
3e29eafabe16
AC rules for bitwise logical operators no longer declared simp
 huffman parents: 
24366diff
changeset | 1487 | Why were these declared simp??? | 
| 65363 | 1488 | declare bin_ops_comm [simp] bbw_assocs [simp] | 
| 24367 
3e29eafabe16
AC rules for bitwise logical operators no longer declared simp
 huffman parents: 
24366diff
changeset | 1489 | *) | 
| 24333 | 1490 | |
| 67120 | 1491 | |
| 61799 | 1492 | subsubsection \<open>Simplification with numerals\<close> | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1493 | |
| 67120 | 1494 | text \<open>Cases for \<open>0\<close> and \<open>-1\<close> are already covered by other simp rules.\<close> | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1495 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1496 | lemma bin_rest_neg_numeral_BitM [simp]: | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54427diff
changeset | 1497 | "bin_rest (- numeral (Num.BitM w)) = - numeral w" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1498 | by (simp only: BIT_bin_simps [symmetric] bin_rest_BIT) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1499 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1500 | lemma bin_last_neg_numeral_BitM [simp]: | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1501 | "bin_last (- numeral (Num.BitM w))" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1502 | by (simp only: BIT_bin_simps [symmetric] bin_last_BIT) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1503 | |
| 67120 | 1504 | (* FIXME: The rule sets below are very large (24 rules for each | 
| 1505 | operator). Is there a simpler way to do this? *) | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1506 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1507 | lemma int_and_numerals [simp]: | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1508 | "numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (numeral x AND numeral y) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1509 | "numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (numeral x AND numeral y) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1510 | "numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (numeral x AND numeral y) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1511 | "numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = (numeral x AND numeral y) BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1512 | "numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (numeral x AND - numeral y) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1513 | "numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (numeral x AND - numeral (y + Num.One)) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1514 | "numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (numeral x AND - numeral y) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1515 | "numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = (numeral x AND - numeral (y + Num.One)) BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1516 | "- numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (- numeral x AND numeral y) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1517 | "- numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (- numeral x AND numeral y) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1518 | "- numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (- numeral (x + Num.One) AND numeral y) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1519 | "- numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = (- numeral (x + Num.One) AND numeral y) BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1520 | "- numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (- numeral x AND - numeral y) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1521 | "- numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (- numeral x AND - numeral (y + Num.One)) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1522 | "- numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (- numeral (x + Num.One) AND - numeral y) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1523 | "- numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = (- numeral (x + Num.One) AND - numeral (y + Num.One)) BIT True" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1524 | "(1::int) AND numeral (Num.Bit0 y) = 0" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1525 | "(1::int) AND numeral (Num.Bit1 y) = 1" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54427diff
changeset | 1526 | "(1::int) AND - numeral (Num.Bit0 y) = 0" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54427diff
changeset | 1527 | "(1::int) AND - numeral (Num.Bit1 y) = 1" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1528 | "numeral (Num.Bit0 x) AND (1::int) = 0" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1529 | "numeral (Num.Bit1 x) AND (1::int) = 1" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54427diff
changeset | 1530 | "- numeral (Num.Bit0 x) AND (1::int) = 0" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54427diff
changeset | 1531 | "- numeral (Num.Bit1 x) AND (1::int) = 1" | 
| 67120 | 1532 | by (rule bin_rl_eqI; simp)+ | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1533 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1534 | lemma int_or_numerals [simp]: | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1535 | "numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (numeral x OR numeral y) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1536 | "numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = (numeral x OR numeral y) BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1537 | "numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = (numeral x OR numeral y) BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1538 | "numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = (numeral x OR numeral y) BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1539 | "numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (numeral x OR - numeral y) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1540 | "numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = (numeral x OR - numeral (y + Num.One)) BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1541 | "numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = (numeral x OR - numeral y) BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1542 | "numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = (numeral x OR - numeral (y + Num.One)) BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1543 | "- numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (- numeral x OR numeral y) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1544 | "- numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = (- numeral x OR numeral y) BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1545 | "- numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = (- numeral (x + Num.One) OR numeral y) BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1546 | "- numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = (- numeral (x + Num.One) OR numeral y) BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1547 | "- numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (- numeral x OR - numeral y) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1548 | "- numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = (- numeral x OR - numeral (y + Num.One)) BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1549 | "- numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = (- numeral (x + Num.One) OR - numeral y) BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1550 | "- numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = (- numeral (x + Num.One) OR - numeral (y + Num.One)) BIT True" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1551 | "(1::int) OR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1552 | "(1::int) OR numeral (Num.Bit1 y) = numeral (Num.Bit1 y)" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54427diff
changeset | 1553 | "(1::int) OR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54427diff
changeset | 1554 | "(1::int) OR - numeral (Num.Bit1 y) = - numeral (Num.Bit1 y)" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1555 | "numeral (Num.Bit0 x) OR (1::int) = numeral (Num.Bit1 x)" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1556 | "numeral (Num.Bit1 x) OR (1::int) = numeral (Num.Bit1 x)" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54427diff
changeset | 1557 | "- numeral (Num.Bit0 x) OR (1::int) = - numeral (Num.BitM x)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54427diff
changeset | 1558 | "- numeral (Num.Bit1 x) OR (1::int) = - numeral (Num.Bit1 x)" | 
| 67120 | 1559 | by (rule bin_rl_eqI; simp)+ | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1560 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1561 | lemma int_xor_numerals [simp]: | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1562 | "numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (numeral x XOR numeral y) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1563 | "numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = (numeral x XOR numeral y) BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1564 | "numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = (numeral x XOR numeral y) BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1565 | "numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (numeral x XOR numeral y) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1566 | "numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (numeral x XOR - numeral y) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1567 | "numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = (numeral x XOR - numeral (y + Num.One)) BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1568 | "numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = (numeral x XOR - numeral y) BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1569 | "numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (numeral x XOR - numeral (y + Num.One)) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1570 | "- numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (- numeral x XOR numeral y) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1571 | "- numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = (- numeral x XOR numeral y) BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1572 | "- numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = (- numeral (x + Num.One) XOR numeral y) BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1573 | "- numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (- numeral (x + Num.One) XOR numeral y) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1574 | "- numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (- numeral x XOR - numeral y) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1575 | "- numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = (- numeral x XOR - numeral (y + Num.One)) BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1576 | "- numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = (- numeral (x + Num.One) XOR - numeral y) BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1577 | "- numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (- numeral (x + Num.One) XOR - numeral (y + Num.One)) BIT False" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1578 | "(1::int) XOR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1579 | "(1::int) XOR numeral (Num.Bit1 y) = numeral (Num.Bit0 y)" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54427diff
changeset | 1580 | "(1::int) XOR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54427diff
changeset | 1581 | "(1::int) XOR - numeral (Num.Bit1 y) = - numeral (Num.Bit0 (y + Num.One))" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1582 | "numeral (Num.Bit0 x) XOR (1::int) = numeral (Num.Bit1 x)" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1583 | "numeral (Num.Bit1 x) XOR (1::int) = numeral (Num.Bit0 x)" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54427diff
changeset | 1584 | "- numeral (Num.Bit0 x) XOR (1::int) = - numeral (Num.BitM x)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54427diff
changeset | 1585 | "- numeral (Num.Bit1 x) XOR (1::int) = - numeral (Num.Bit0 (x + Num.One))" | 
| 67120 | 1586 | by (rule bin_rl_eqI; simp)+ | 
| 1587 | ||
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1588 | |
| 61799 | 1589 | subsubsection \<open>Interactions with arithmetic\<close> | 
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1590 | |
| 67120 | 1591 | lemma plus_and_or [rule_format]: "\<forall>y::int. (x AND y) + (x OR y) = x + y" | 
| 24333 | 1592 | apply (induct x rule: bin_induct) | 
| 1593 | apply clarsimp | |
| 1594 | apply clarsimp | |
| 1595 | apply clarsimp | |
| 1596 | apply (case_tac y rule: bin_exhaust) | |
| 1597 | apply clarsimp | |
| 1598 | apply (unfold Bit_def) | |
| 1599 | apply clarsimp | |
| 1600 | apply (erule_tac x = "x" in allE) | |
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 1601 | apply simp | 
| 24333 | 1602 | done | 
| 1603 | ||
| 67120 | 1604 | lemma le_int_or: "bin_sign y = 0 \<Longrightarrow> x \<le> x OR y" | 
| 1605 | for x y :: int | |
| 37667 | 1606 | apply (induct y arbitrary: x rule: bin_induct) | 
| 24333 | 1607 | apply clarsimp | 
| 1608 | apply clarsimp | |
| 1609 | apply (case_tac x rule: bin_exhaust) | |
| 1610 | apply (case_tac b) | |
| 1611 | apply (case_tac [!] bit) | |
| 46604 
9f9e85264e4d
make uses of bin_sign respect int/bin distinction
 huffman parents: 
46023diff
changeset | 1612 | apply (auto simp: le_Bits) | 
| 24333 | 1613 | done | 
| 1614 | ||
| 1615 | lemmas int_and_le = | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 1616 | xtrans(3) [OF bbw_ao_absorbs (2) [THEN conjunct2, symmetric] le_int_or] | 
| 24333 | 1617 | |
| 67120 | 1618 | text \<open>Interaction between bit-wise and arithmetic: good example of \<open>bin_induction\<close>.\<close> | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1619 | lemma bin_add_not: "x + NOT x = (-1::int)" | 
| 24364 | 1620 | apply (induct x rule: bin_induct) | 
| 1621 | apply clarsimp | |
| 1622 | apply clarsimp | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46610diff
changeset | 1623 | apply (case_tac bit, auto) | 
| 24364 | 1624 | done | 
| 1625 | ||
| 70169 | 1626 | lemma mod_BIT: "bin BIT bit mod 2 ^ Suc n = (bin mod 2 ^ n) BIT bit" | 
| 1627 | proof - | |
| 1628 | have "2 * (bin mod 2 ^ n) + 1 = (2 * bin mod 2 ^ Suc n) + 1" | |
| 1629 | by (simp add: mod_mult_mult1) | |
| 1630 | also have "\<dots> = ((2 * bin mod 2 ^ Suc n) + 1) mod 2 ^ Suc n" | |
| 70170 | 1631 | by (simp add: ac_simps pos_zmod_mult_2) | 
| 70169 | 1632 | also have "\<dots> = (2 * bin + 1) mod 2 ^ Suc n" | 
| 1633 | by (simp only: mod_simps) | |
| 1634 | finally show ?thesis | |
| 1635 | by (auto simp add: Bit_def) | |
| 1636 | qed | |
| 1637 | ||
| 1638 | lemma AND_mod: "x AND 2 ^ n - 1 = x mod 2 ^ n" | |
| 1639 | for x :: int | |
| 1640 | proof (induct x arbitrary: n rule: bin_induct) | |
| 1641 | case 1 | |
| 1642 | then show ?case | |
| 1643 | by simp | |
| 1644 | next | |
| 1645 | case 2 | |
| 1646 | then show ?case | |
| 1647 | by (simp, simp add: m1mod2k) | |
| 1648 | next | |
| 1649 | case (3 bin bit) | |
| 1650 | show ?case | |
| 1651 | proof (cases n) | |
| 1652 | case 0 | |
| 1653 | then show ?thesis by simp | |
| 1654 | next | |
| 1655 | case (Suc m) | |
| 1656 | with 3 show ?thesis | |
| 1657 | by (simp only: power_BIT mod_BIT int_and_Bits) simp | |
| 1658 | qed | |
| 1659 | qed | |
| 1660 | ||
| 67120 | 1661 | |
| 70172 | 1662 | subsubsection \<open>Comparison\<close> | 
| 1663 | ||
| 1664 | lemma AND_lower [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close> | |
| 1665 | fixes x y :: int | |
| 1666 | assumes "0 \<le> x" | |
| 1667 | shows "0 \<le> x AND y" | |
| 1668 | using assms | |
| 1669 | proof (induct x arbitrary: y rule: bin_induct) | |
| 1670 | case 1 | |
| 1671 | then show ?case by simp | |
| 1672 | next | |
| 1673 | case 2 | |
| 1674 | then show ?case by (simp only: Min_def) | |
| 1675 | next | |
| 1676 | case (3 bin bit) | |
| 1677 | show ?case | |
| 1678 | proof (cases y rule: bin_exhaust) | |
| 1679 | case (1 bin' bit') | |
| 1680 | from 3 have "0 \<le> bin" | |
| 1681 | by (cases bit) (simp_all add: Bit_def) | |
| 1682 | then have "0 \<le> bin AND bin'" by (rule 3) | |
| 1683 | with 1 show ?thesis | |
| 70190 | 1684 | by simp | 
| 70172 | 1685 | qed | 
| 1686 | qed | |
| 1687 | ||
| 1688 | lemma OR_lower [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close> | |
| 1689 | fixes x y :: int | |
| 1690 | assumes "0 \<le> x" "0 \<le> y" | |
| 1691 | shows "0 \<le> x OR y" | |
| 1692 | using assms | |
| 1693 | proof (induct x arbitrary: y rule: bin_induct) | |
| 1694 | case (3 bin bit) | |
| 1695 | show ?case | |
| 1696 | proof (cases y rule: bin_exhaust) | |
| 1697 | case (1 bin' bit') | |
| 1698 | from 3 have "0 \<le> bin" | |
| 1699 | by (cases bit) (simp_all add: Bit_def) | |
| 1700 | moreover from 1 3 have "0 \<le> bin'" | |
| 1701 | by (cases bit') (simp_all add: Bit_def) | |
| 1702 | ultimately have "0 \<le> bin OR bin'" by (rule 3) | |
| 1703 | with 1 show ?thesis | |
| 70190 | 1704 | by simp | 
| 70172 | 1705 | qed | 
| 1706 | qed simp_all | |
| 1707 | ||
| 1708 | lemma XOR_lower [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close> | |
| 1709 | fixes x y :: int | |
| 1710 | assumes "0 \<le> x" "0 \<le> y" | |
| 1711 | shows "0 \<le> x XOR y" | |
| 1712 | using assms | |
| 1713 | proof (induct x arbitrary: y rule: bin_induct) | |
| 1714 | case (3 bin bit) | |
| 1715 | show ?case | |
| 1716 | proof (cases y rule: bin_exhaust) | |
| 1717 | case (1 bin' bit') | |
| 1718 | from 3 have "0 \<le> bin" | |
| 1719 | by (cases bit) (simp_all add: Bit_def) | |
| 1720 | moreover from 1 3 have "0 \<le> bin'" | |
| 1721 | by (cases bit') (simp_all add: Bit_def) | |
| 1722 | ultimately have "0 \<le> bin XOR bin'" by (rule 3) | |
| 1723 | with 1 show ?thesis | |
| 70190 | 1724 | by simp | 
| 70172 | 1725 | qed | 
| 1726 | next | |
| 1727 | case 2 | |
| 1728 | then show ?case by (simp only: Min_def) | |
| 1729 | qed simp | |
| 1730 | ||
| 1731 | lemma AND_upper1 [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close> | |
| 1732 | fixes x y :: int | |
| 1733 | assumes "0 \<le> x" | |
| 1734 | shows "x AND y \<le> x" | |
| 1735 | using assms | |
| 1736 | proof (induct x arbitrary: y rule: bin_induct) | |
| 1737 | case (3 bin bit) | |
| 1738 | show ?case | |
| 1739 | proof (cases y rule: bin_exhaust) | |
| 1740 | case (1 bin' bit') | |
| 1741 | from 3 have "0 \<le> bin" | |
| 1742 | by (cases bit) (simp_all add: Bit_def) | |
| 1743 | then have "bin AND bin' \<le> bin" by (rule 3) | |
| 1744 | with 1 show ?thesis | |
| 1745 | by simp (simp add: Bit_def) | |
| 1746 | qed | |
| 1747 | next | |
| 1748 | case 2 | |
| 1749 | then show ?case by (simp only: Min_def) | |
| 1750 | qed simp | |
| 1751 | ||
| 1752 | lemmas AND_upper1' [simp] = order_trans [OF AND_upper1] \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close> | |
| 1753 | lemmas AND_upper1'' [simp] = order_le_less_trans [OF AND_upper1] \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close> | |
| 1754 | ||
| 1755 | lemma AND_upper2 [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close> | |
| 1756 | fixes x y :: int | |
| 1757 | assumes "0 \<le> y" | |
| 1758 | shows "x AND y \<le> y" | |
| 1759 | using assms | |
| 1760 | proof (induct y arbitrary: x rule: bin_induct) | |
| 1761 | case 1 | |
| 1762 | then show ?case by simp | |
| 1763 | next | |
| 1764 | case 2 | |
| 1765 | then show ?case by (simp only: Min_def) | |
| 1766 | next | |
| 1767 | case (3 bin bit) | |
| 1768 | show ?case | |
| 1769 | proof (cases x rule: bin_exhaust) | |
| 1770 | case (1 bin' bit') | |
| 1771 | from 3 have "0 \<le> bin" | |
| 1772 | by (cases bit) (simp_all add: Bit_def) | |
| 1773 | then have "bin' AND bin \<le> bin" by (rule 3) | |
| 1774 | with 1 show ?thesis | |
| 1775 | by simp (simp add: Bit_def) | |
| 1776 | qed | |
| 1777 | qed | |
| 1778 | ||
| 1779 | lemmas AND_upper2' [simp] = order_trans [OF AND_upper2] \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close> | |
| 1780 | lemmas AND_upper2'' [simp] = order_le_less_trans [OF AND_upper2] \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close> | |
| 1781 | ||
| 1782 | lemma OR_upper: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close> | |
| 1783 | fixes x y :: int | |
| 1784 | assumes "0 \<le> x" "x < 2 ^ n" "y < 2 ^ n" | |
| 1785 | shows "x OR y < 2 ^ n" | |
| 1786 | using assms | |
| 1787 | proof (induct x arbitrary: y n rule: bin_induct) | |
| 1788 | case (3 bin bit) | |
| 1789 | show ?case | |
| 1790 | proof (cases y rule: bin_exhaust) | |
| 1791 | case (1 bin' bit') | |
| 1792 | show ?thesis | |
| 1793 | proof (cases n) | |
| 1794 | case 0 | |
| 70190 | 1795 | with 3 have "bin BIT bit = 0" | 
| 1796 | by (simp add: Bit_def) | |
| 70172 | 1797 | then have "bin = 0" and "\<not> bit" | 
| 1798 | by (auto simp add: Bit_def split: if_splits) arith | |
| 1799 | then show ?thesis using 0 1 \<open>y < 2 ^ n\<close> | |
| 1800 | by simp | |
| 1801 | next | |
| 1802 | case (Suc m) | |
| 1803 | from 3 have "0 \<le> bin" | |
| 1804 | by (cases bit) (simp_all add: Bit_def) | |
| 1805 | moreover from 3 Suc have "bin < 2 ^ m" | |
| 1806 | by (cases bit) (simp_all add: Bit_def) | |
| 1807 | moreover from 1 3 Suc have "bin' < 2 ^ m" | |
| 1808 | by (cases bit') (simp_all add: Bit_def) | |
| 1809 | ultimately have "bin OR bin' < 2 ^ m" by (rule 3) | |
| 1810 | with 1 Suc show ?thesis | |
| 1811 | by simp (simp add: Bit_def) | |
| 1812 | qed | |
| 1813 | qed | |
| 1814 | qed simp_all | |
| 1815 | ||
| 1816 | lemma XOR_upper: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close> | |
| 1817 | fixes x y :: int | |
| 1818 | assumes "0 \<le> x" "x < 2 ^ n" "y < 2 ^ n" | |
| 1819 | shows "x XOR y < 2 ^ n" | |
| 1820 | using assms | |
| 1821 | proof (induct x arbitrary: y n rule: bin_induct) | |
| 1822 | case 1 | |
| 1823 | then show ?case by simp | |
| 1824 | next | |
| 1825 | case 2 | |
| 1826 | then show ?case by (simp only: Min_def) | |
| 1827 | next | |
| 1828 | case (3 bin bit) | |
| 1829 | show ?case | |
| 1830 | proof (cases y rule: bin_exhaust) | |
| 1831 | case (1 bin' bit') | |
| 1832 | show ?thesis | |
| 1833 | proof (cases n) | |
| 1834 | case 0 | |
| 70190 | 1835 | with 3 have "bin BIT bit = 0" | 
| 1836 | by (simp add: Bit_def) | |
| 70172 | 1837 | then have "bin = 0" and "\<not> bit" | 
| 1838 | by (auto simp add: Bit_def split: if_splits) arith | |
| 1839 | then show ?thesis using 0 1 \<open>y < 2 ^ n\<close> | |
| 1840 | by simp | |
| 1841 | next | |
| 1842 | case (Suc m) | |
| 1843 | from 3 have "0 \<le> bin" | |
| 1844 | by (cases bit) (simp_all add: Bit_def) | |
| 1845 | moreover from 3 Suc have "bin < 2 ^ m" | |
| 1846 | by (cases bit) (simp_all add: Bit_def) | |
| 1847 | moreover from 1 3 Suc have "bin' < 2 ^ m" | |
| 1848 | by (cases bit') (simp_all add: Bit_def) | |
| 1849 | ultimately have "bin XOR bin' < 2 ^ m" by (rule 3) | |
| 1850 | with 1 Suc show ?thesis | |
| 1851 | by simp (simp add: Bit_def) | |
| 1852 | qed | |
| 1853 | qed | |
| 1854 | qed | |
| 1855 | ||
| 1856 | ||
| 1857 | ||
| 61799 | 1858 | subsubsection \<open>Truncating results of bit-wise operations\<close> | 
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1859 | |
| 65363 | 1860 | lemma bin_trunc_ao: | 
| 67120 | 1861 | "bintrunc n x AND bintrunc n y = bintrunc n (x AND y)" | 
| 1862 | "bintrunc n x OR bintrunc n y = bintrunc n (x OR y)" | |
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1863 | by (auto simp add: bin_eq_iff bin_nth_ops nth_bintr) | 
| 24364 | 1864 | |
| 67120 | 1865 | lemma bin_trunc_xor: "bintrunc n (bintrunc n x XOR bintrunc n y) = bintrunc n (x XOR y)" | 
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1866 | by (auto simp add: bin_eq_iff bin_nth_ops nth_bintr) | 
| 24364 | 1867 | |
| 67120 | 1868 | lemma bin_trunc_not: "bintrunc n (NOT (bintrunc n x)) = bintrunc n (NOT x)" | 
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 1869 | by (auto simp add: bin_eq_iff bin_nth_ops nth_bintr) | 
| 24364 | 1870 | |
| 67120 | 1871 | text \<open>Want theorems of the form of \<open>bin_trunc_xor\<close>.\<close> | 
| 1872 | lemma bintr_bintr_i: "x = bintrunc n y \<Longrightarrow> bintrunc n x = bintrunc n y" | |
| 24364 | 1873 | by auto | 
| 1874 | ||
| 1875 | lemmas bin_trunc_and = bin_trunc_ao(1) [THEN bintr_bintr_i] | |
| 1876 | lemmas bin_trunc_or = bin_trunc_ao(2) [THEN bintr_bintr_i] | |
| 1877 | ||
| 70190 | 1878 | |
| 1879 | subsubsection \<open>More lemmas\<close> | |
| 1880 | ||
| 1881 | lemma not_int_cmp_0 [simp]: | |
| 1882 | fixes i :: int shows | |
| 1883 | "0 < NOT i \<longleftrightarrow> i < -1" | |
| 1884 | "0 \<le> NOT i \<longleftrightarrow> i < 0" | |
| 1885 | "NOT i < 0 \<longleftrightarrow> i \<ge> 0" | |
| 1886 | "NOT i \<le> 0 \<longleftrightarrow> i \<ge> -1" | |
| 1887 | by(simp_all add: int_not_def) arith+ | |
| 1888 | ||
| 1889 | lemma bbw_ao_dist2: "(x :: int) AND (y OR z) = x AND y OR x AND z" | |
| 1890 | by(metis int_and_comm bbw_ao_dist) | |
| 1891 | ||
| 1892 | lemmas int_and_ac = bbw_lcs(1) int_and_comm int_and_assoc | |
| 1893 | ||
| 1894 | lemma int_nand_same [simp]: fixes x :: int shows "x AND NOT x = 0" | |
| 1895 | by(induct x y\<equiv>"NOT x" rule: bitAND_int.induct)(subst bitAND_int.simps, clarsimp) | |
| 1896 | ||
| 1897 | lemma int_nand_same_middle: fixes x :: int shows "x AND y AND NOT x = 0" | |
| 1898 | by (metis bbw_lcs(1) int_and_0 int_nand_same) | |
| 1899 | ||
| 1900 | lemma and_xor_dist: fixes x :: int shows | |
| 1901 | "x AND (y XOR z) = (x AND y) XOR (x AND z)" | |
| 1902 | by(simp add: int_xor_def bbw_ao_dist2 bbw_not_dist int_and_ac int_nand_same_middle) | |
| 1903 | ||
| 1904 | lemma int_and_lt0 [simp]: fixes x y :: int shows | |
| 1905 | "x AND y < 0 \<longleftrightarrow> x < 0 \<and> y < 0" | |
| 1906 | by(induct x y rule: bitAND_int.induct)(subst bitAND_int.simps, simp) | |
| 1907 | ||
| 1908 | lemma int_and_ge0 [simp]: fixes x y :: int shows | |
| 1909 | "x AND y \<ge> 0 \<longleftrightarrow> x \<ge> 0 \<or> y \<ge> 0" | |
| 1910 | by (metis int_and_lt0 linorder_not_less) | |
| 1911 | ||
| 1912 | lemma int_and_1: fixes x :: int shows "x AND 1 = x mod 2" | |
| 1913 | by(subst bitAND_int.simps)(simp add: Bit_def bin_last_def zmod_minus1) | |
| 1914 | ||
| 1915 | lemma int_1_and: fixes x :: int shows "1 AND x = x mod 2" | |
| 1916 | by(subst int_and_comm)(simp add: int_and_1) | |
| 1917 | ||
| 1918 | lemma int_or_lt0 [simp]: fixes x y :: int shows | |
| 1919 | "x OR y < 0 \<longleftrightarrow> x < 0 \<or> y < 0" | |
| 1920 | by(simp add: int_or_def) | |
| 1921 | ||
| 1922 | lemma int_xor_lt0 [simp]: fixes x y :: int shows | |
| 1923 | "x XOR y < 0 \<longleftrightarrow> ((x < 0) \<noteq> (y < 0))" | |
| 1924 | by(auto simp add: int_xor_def) | |
| 1925 | ||
| 1926 | lemma int_xor_ge0 [simp]: fixes x y :: int shows | |
| 1927 | "x XOR y \<ge> 0 \<longleftrightarrow> ((x \<ge> 0) \<longleftrightarrow> (y \<ge> 0))" | |
| 1928 | by (metis int_xor_lt0 linorder_not_le) | |
| 1929 | ||
| 1930 | lemma bin_last_conv_AND: | |
| 1931 | "bin_last i \<longleftrightarrow> i AND 1 \<noteq> 0" | |
| 1932 | proof - | |
| 1933 | obtain x b where "i = x BIT b" by(cases i rule: bin_exhaust) | |
| 1934 | hence "i AND 1 = 0 BIT b" | |
| 1935 | by(simp add: BIT_special_simps(2)[symmetric] del: BIT_special_simps(2)) | |
| 1936 | thus ?thesis using \<open>i = x BIT b\<close> by(cases b) simp_all | |
| 1937 | qed | |
| 1938 | ||
| 1939 | lemma bitval_bin_last: | |
| 1940 | "of_bool (bin_last i) = i AND 1" | |
| 1941 | proof - | |
| 1942 | obtain x b where "i = x BIT b" by(cases i rule: bin_exhaust) | |
| 1943 | hence "i AND 1 = 0 BIT b" | |
| 1944 | by(simp add: BIT_special_simps(2)[symmetric] del: BIT_special_simps(2)) | |
| 1945 | thus ?thesis by(cases b)(simp_all add: bin_last_conv_AND) | |
| 1946 | qed | |
| 1947 | ||
| 1948 | lemma bin_sign_and: | |
| 1949 | "bin_sign (i AND j) = - (bin_sign i * bin_sign j)" | |
| 1950 | by(simp add: bin_sign_def) | |
| 1951 | ||
| 1952 | lemma minus_BIT_0: fixes x y :: int shows "x BIT b - y BIT False = (x - y) BIT b" | |
| 1953 | by(simp add: Bit_def) | |
| 1954 | ||
| 1955 | lemma int_not_neg_numeral: "NOT (- numeral n) = (Num.sub n num.One :: int)" | |
| 1956 | by(simp add: int_not_def) | |
| 1957 | ||
| 1958 | lemma int_neg_numeral_pOne_conv_not: "- numeral (n + num.One) = (NOT (numeral n) :: int)" | |
| 1959 | by(simp add: int_not_def) | |
| 70169 | 1960 | |
| 67120 | 1961 | |
| 61799 | 1962 | subsection \<open>Setting and clearing bits\<close> | 
| 24364 | 1963 | |
| 70191 | 1964 | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 1965 | |
| 70183 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1966 | lemma int_lsb_BIT [simp]: fixes x :: int shows | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1967 | "lsb (x BIT b) \<longleftrightarrow> b" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1968 | by(simp add: lsb_int_def) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1969 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1970 | lemma bin_last_conv_lsb: "bin_last = lsb" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1971 | by(clarsimp simp add: lsb_int_def fun_eq_iff) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1972 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1973 | lemma int_lsb_numeral [simp]: | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1974 | "lsb (0 :: int) = False" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1975 | "lsb (1 :: int) = True" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1976 | "lsb (Numeral1 :: int) = True" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1977 | "lsb (- 1 :: int) = True" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1978 | "lsb (- Numeral1 :: int) = True" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1979 | "lsb (numeral (num.Bit0 w) :: int) = False" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1980 | "lsb (numeral (num.Bit1 w) :: int) = True" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1981 | "lsb (- numeral (num.Bit0 w) :: int) = False" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1982 | "lsb (- numeral (num.Bit1 w) :: int) = True" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1983 | by(simp_all add: lsb_int_def) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1984 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1985 | lemma int_set_bit_0 [simp]: fixes x :: int shows | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1986 | "set_bit x 0 b = bin_rest x BIT b" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1987 | by(auto simp add: set_bit_int_def intro: bin_rl_eqI) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1988 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1989 | lemma int_set_bit_Suc: fixes x :: int shows | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1990 | "set_bit x (Suc n) b = set_bit (bin_rest x) n b BIT bin_last x" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1991 | by(auto simp add: set_bit_int_def twice_conv_BIT intro: bin_rl_eqI) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1992 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1993 | lemma bin_last_set_bit: | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1994 | "bin_last (set_bit x n b) = (if n > 0 then bin_last x else b)" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1995 | by(cases n)(simp_all add: int_set_bit_Suc) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1996 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1997 | lemma bin_rest_set_bit: | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1998 | "bin_rest (set_bit x n b) = (if n > 0 then set_bit (bin_rest x) (n - 1) b else bin_rest x)" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 1999 | by(cases n)(simp_all add: int_set_bit_Suc) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2000 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2001 | lemma int_set_bit_numeral: fixes x :: int shows | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2002 | "set_bit x (numeral w) b = set_bit (bin_rest x) (pred_numeral w) b BIT bin_last x" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2003 | by(simp add: set_bit_int_def) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2004 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2005 | lemmas int_set_bit_numerals [simp] = | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2006 | int_set_bit_numeral[where x="numeral w'"] | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2007 | int_set_bit_numeral[where x="- numeral w'"] | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2008 | int_set_bit_numeral[where x="Numeral1"] | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2009 | int_set_bit_numeral[where x="1"] | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2010 | int_set_bit_numeral[where x="0"] | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2011 | int_set_bit_Suc[where x="numeral w'"] | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2012 | int_set_bit_Suc[where x="- numeral w'"] | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2013 | int_set_bit_Suc[where x="Numeral1"] | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2014 | int_set_bit_Suc[where x="1"] | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2015 | int_set_bit_Suc[where x="0"] | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2016 | for w' | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2017 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2018 | lemma int_shiftl_BIT: fixes x :: int | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2019 | shows int_shiftl0 [simp]: "x << 0 = x" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2020 | and int_shiftl_Suc [simp]: "x << Suc n = (x << n) BIT False" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2021 | by(auto simp add: shiftl_int_def Bit_def) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2022 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2023 | lemma int_0_shiftl [simp]: "0 << n = (0 :: int)" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2024 | by(induct n) simp_all | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2025 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2026 | lemma bin_last_shiftl: "bin_last (x << n) \<longleftrightarrow> n = 0 \<and> bin_last x" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2027 | by(cases n)(simp_all) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2028 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2029 | lemma bin_rest_shiftl: "bin_rest (x << n) = (if n > 0 then x << (n - 1) else bin_rest x)" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2030 | by(cases n)(simp_all) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2031 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2032 | lemma bin_nth_shiftl [simp]: "bin_nth (x << n) m \<longleftrightarrow> n \<le> m \<and> bin_nth x (m - n)" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2033 | proof(induct n arbitrary: x m) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2034 | case (Suc n) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2035 | thus ?case by(cases m) simp_all | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2036 | qed simp | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2037 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2038 | lemma int_shiftr_BIT [simp]: fixes x :: int | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2039 | shows int_shiftr0: "x >> 0 = x" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2040 | and int_shiftr_Suc: "x BIT b >> Suc n = x >> n" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2041 | proof - | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2042 | show "x >> 0 = x" by (simp add: shiftr_int_def) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2043 | show "x BIT b >> Suc n = x >> n" by (cases b) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2044 | (simp_all add: shiftr_int_def Bit_def add.commute pos_zdiv_mult_2) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2045 | qed | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2046 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2047 | lemma bin_last_shiftr: "bin_last (x >> n) \<longleftrightarrow> x !! n" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2048 | proof(induct n arbitrary: x) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2049 | case 0 thus ?case by simp | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2050 | next | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2051 | case (Suc n) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2052 | thus ?case by(cases x rule: bin_exhaust) simp | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2053 | qed | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2054 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2055 | lemma bin_rest_shiftr [simp]: "bin_rest (x >> n) = x >> Suc n" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2056 | proof(induct n arbitrary: x) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2057 | case 0 | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2058 | thus ?case by(cases x rule: bin_exhaust) auto | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2059 | next | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2060 | case (Suc n) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2061 | thus ?case by(cases x rule: bin_exhaust) auto | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2062 | qed | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2063 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2064 | lemma bin_nth_shiftr [simp]: "bin_nth (x >> n) m = bin_nth x (n + m)" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2065 | proof(induct n arbitrary: x m) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2066 | case (Suc n) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2067 | thus ?case by(cases x rule: bin_exhaust) simp_all | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2068 | qed simp | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2069 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2070 | lemma bin_nth_conv_AND: | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2071 | fixes x :: int shows | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2072 | "bin_nth x n \<longleftrightarrow> x AND (1 << n) \<noteq> 0" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2073 | proof(induct n arbitrary: x) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2074 | case 0 | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2075 | thus ?case by(simp add: int_and_1 bin_last_def) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2076 | next | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2077 | case (Suc n) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2078 | thus ?case by(cases x rule: bin_exhaust)(simp_all add: bin_nth_ops Bit_eq_0_iff) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2079 | qed | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2080 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2081 | lemma int_shiftl_numeral [simp]: | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2082 | "(numeral w :: int) << numeral w' = numeral (num.Bit0 w) << pred_numeral w'" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2083 | "(- numeral w :: int) << numeral w' = - numeral (num.Bit0 w) << pred_numeral w'" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2084 | by(simp_all add: numeral_eq_Suc Bit_def shiftl_int_def) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2085 | (metis add_One mult_inc semiring_norm(11) semiring_norm(13) semiring_norm(2) semiring_norm(6) semiring_norm(87))+ | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2086 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2087 | lemma int_shiftl_One_numeral [simp]: "(1 :: int) << numeral w = 2 << pred_numeral w" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2088 | by(metis int_shiftl_numeral numeral_One) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2089 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2090 | lemma shiftl_ge_0 [simp]: fixes i :: int shows "i << n \<ge> 0 \<longleftrightarrow> i \<ge> 0" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2091 | by(induct n) simp_all | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2092 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2093 | lemma shiftl_lt_0 [simp]: fixes i :: int shows "i << n < 0 \<longleftrightarrow> i < 0" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2094 | by (metis not_le shiftl_ge_0) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2095 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2096 | lemma int_shiftl_test_bit: "(n << i :: int) !! m \<longleftrightarrow> m \<ge> i \<and> n !! (m - i)" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2097 | proof(induction i) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2098 | case (Suc n) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2099 | thus ?case by(cases m) simp_all | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2100 | qed simp | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2101 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2102 | lemma int_0shiftr [simp]: "(0 :: int) >> x = 0" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2103 | by(simp add: shiftr_int_def) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2104 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2105 | lemma int_minus1_shiftr [simp]: "(-1 :: int) >> x = -1" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2106 | by(simp add: shiftr_int_def div_eq_minus1) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2107 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2108 | lemma int_shiftr_ge_0 [simp]: fixes i :: int shows "i >> n \<ge> 0 \<longleftrightarrow> i \<ge> 0" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2109 | proof(induct n arbitrary: i) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2110 | case (Suc n) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2111 | thus ?case by(cases i rule: bin_exhaust) simp_all | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2112 | qed simp | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2113 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2114 | lemma int_shiftr_lt_0 [simp]: fixes i :: int shows "i >> n < 0 \<longleftrightarrow> i < 0" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2115 | by (metis int_shiftr_ge_0 not_less) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2116 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2117 | lemma int_shiftr_numeral [simp]: | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2118 | "(1 :: int) >> numeral w' = 0" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2119 | "(numeral num.One :: int) >> numeral w' = 0" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2120 | "(numeral (num.Bit0 w) :: int) >> numeral w' = numeral w >> pred_numeral w'" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2121 | "(numeral (num.Bit1 w) :: int) >> numeral w' = numeral w >> pred_numeral w'" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2122 | "(- numeral (num.Bit0 w) :: int) >> numeral w' = - numeral w >> pred_numeral w'" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2123 | "(- numeral (num.Bit1 w) :: int) >> numeral w' = - numeral (Num.inc w) >> pred_numeral w'" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2124 | by (simp_all only: numeral_One expand_BIT numeral_eq_Suc int_shiftr_Suc BIT_special_simps(2)[symmetric] int_0shiftr add_One uminus_Bit_eq) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2125 | (simp_all add: add_One) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2126 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2127 | lemma int_shiftr_numeral_Suc0 [simp]: | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2128 | "(1 :: int) >> Suc 0 = 0" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2129 | "(numeral num.One :: int) >> Suc 0 = 0" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2130 | "(numeral (num.Bit0 w) :: int) >> Suc 0 = numeral w" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2131 | "(numeral (num.Bit1 w) :: int) >> Suc 0 = numeral w" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2132 | "(- numeral (num.Bit0 w) :: int) >> Suc 0 = - numeral w" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2133 | "(- numeral (num.Bit1 w) :: int) >> Suc 0 = - numeral (Num.inc w)" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2134 | by(simp_all only: One_nat_def[symmetric] numeral_One[symmetric] int_shiftr_numeral pred_numeral_simps int_shiftr0) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2135 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2136 | lemma bin_nth_minus_p2: | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2137 | assumes sign: "bin_sign x = 0" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2138 | and y: "y = 1 << n" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2139 | and m: "m < n" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2140 | and x: "x < y" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2141 | shows "bin_nth (x - y) m = bin_nth x m" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2142 | using sign m x unfolding y | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2143 | proof(induction m arbitrary: x y n) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2144 | case 0 | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2145 | thus ?case | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2146 | by(simp add: bin_last_def shiftl_int_def) (metis (hide_lams, no_types) mod_diff_right_eq mod_self neq0_conv numeral_One power_eq_0_iff power_mod diff_zero zero_neq_numeral) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2147 | next | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2148 | case (Suc m) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2149 | from \<open>Suc m < n\<close> obtain n' where [simp]: "n = Suc n'" by(cases n) auto | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2150 | obtain x' b where [simp]: "x = x' BIT b" by(cases x rule: bin_exhaust) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2151 | from \<open>bin_sign x = 0\<close> have "bin_sign x' = 0" by simp | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2152 | moreover from \<open>x < 1 << n\<close> have "x' < 1 << n'" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2153 | by(cases b)(simp_all add: Bit_def shiftl_int_def) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2154 | moreover have "(2 * x' + of_bool b - 2 * 2 ^ n') div 2 = x' + (- (2 ^ n') + of_bool b div 2)" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2155 | by(simp only: add_diff_eq[symmetric] add.commute div_mult_self2[OF zero_neq_numeral[symmetric]]) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2156 | ultimately show ?case using Suc.IH[of x' n'] Suc.prems | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2157 | by(cases b)(simp_all add: Bit_def bin_rest_def shiftl_int_def) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2158 | qed | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2159 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2160 | lemma bin_clr_conv_NAND: | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2161 | "bin_sc n False i = i AND NOT (1 << n)" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2162 | by(induct n arbitrary: i)(auto intro: bin_rl_eqI) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2163 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2164 | lemma bin_set_conv_OR: | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2165 | "bin_sc n True i = i OR (1 << n)" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2166 | by(induct n arbitrary: i)(auto intro: bin_rl_eqI) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2167 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2168 | lemma msb_conv_bin_sign: "msb x \<longleftrightarrow> bin_sign x = -1" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2169 | by(simp add: bin_sign_def not_le msb_int_def) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2170 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2171 | lemma msb_BIT [simp]: "msb (x BIT b) = msb x" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2172 | by(simp add: msb_int_def) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2173 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2174 | lemma msb_bin_rest [simp]: "msb (bin_rest x) = msb x" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2175 | by(simp add: msb_int_def) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2176 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2177 | lemma int_msb_and [simp]: "msb ((x :: int) AND y) \<longleftrightarrow> msb x \<and> msb y" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2178 | by(simp add: msb_int_def) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2179 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2180 | lemma int_msb_or [simp]: "msb ((x :: int) OR y) \<longleftrightarrow> msb x \<or> msb y" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2181 | by(simp add: msb_int_def) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2182 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2183 | lemma int_msb_xor [simp]: "msb ((x :: int) XOR y) \<longleftrightarrow> msb x \<noteq> msb y" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2184 | by(simp add: msb_int_def) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2185 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2186 | lemma int_msb_not [simp]: "msb (NOT (x :: int)) \<longleftrightarrow> \<not> msb x" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2187 | by(simp add: msb_int_def not_less) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2188 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2189 | lemma msb_shiftl [simp]: "msb ((x :: int) << n) \<longleftrightarrow> msb x" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2190 | by(simp add: msb_int_def) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2191 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2192 | lemma msb_shiftr [simp]: "msb ((x :: int) >> r) \<longleftrightarrow> msb x" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2193 | by(simp add: msb_int_def) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2194 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2195 | lemma msb_bin_sc [simp]: "msb (bin_sc n b x) \<longleftrightarrow> msb x" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2196 | by(simp add: msb_conv_bin_sign) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2197 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2198 | lemma msb_set_bit [simp]: "msb (set_bit (x :: int) n b) \<longleftrightarrow> msb x" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2199 | by(simp add: msb_conv_bin_sign set_bit_int_def) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2200 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2201 | lemma msb_0 [simp]: "msb (0 :: int) = False" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2202 | by(simp add: msb_int_def) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2203 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2204 | lemma msb_1 [simp]: "msb (1 :: int) = False" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2205 | by(simp add: msb_int_def) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2206 | |
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2207 | lemma msb_numeral [simp]: | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2208 | "msb (numeral n :: int) = False" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2209 | "msb (- numeral n :: int) = True" | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2210 | by(simp_all add: msb_int_def) | 
| 
3ea80c950023
incorporated various material from the AFP into the distribution
 haftmann parents: 
70175diff
changeset | 2211 | |
| 70190 | 2212 | |
| 2213 | subsection \<open>Semantic interpretation of \<^typ>\<open>bool list\<close> as \<^typ>\<open>int\<close>\<close> | |
| 2214 | ||
| 2215 | lemma bin_bl_bin': "bl_to_bin (bin_to_bl_aux n w bs) = bl_to_bin_aux bs (bintrunc n w)" | |
| 2216 | by (induct n arbitrary: w bs) (auto simp: bl_to_bin_def) | |
| 2217 | ||
| 2218 | lemma bin_bl_bin [simp]: "bl_to_bin (bin_to_bl n w) = bintrunc n w" | |
| 2219 | by (auto simp: bin_to_bl_def bin_bl_bin') | |
| 2220 | ||
| 2221 | lemma bl_to_bin_rep_F: "bl_to_bin (replicate n False @ bl) = bl_to_bin bl" | |
| 2222 | by (simp add: bin_to_bl_zero_aux [symmetric] bin_bl_bin') (simp add: bl_to_bin_def) | |
| 2223 | ||
| 2224 | lemma bin_to_bl_trunc [simp]: "n \<le> m \<Longrightarrow> bin_to_bl n (bintrunc m w) = bin_to_bl n w" | |
| 2225 | by (auto intro: bl_to_bin_inj) | |
| 2226 | ||
| 2227 | lemma bin_to_bl_aux_bintr: | |
| 2228 | "bin_to_bl_aux n (bintrunc m bin) bl = | |
| 2229 | replicate (n - m) False @ bin_to_bl_aux (min n m) bin bl" | |
| 2230 | apply (induct n arbitrary: m bin bl) | |
| 2231 | apply clarsimp | |
| 2232 | apply clarsimp | |
| 2233 | apply (case_tac "m") | |
| 2234 | apply (clarsimp simp: bin_to_bl_zero_aux) | |
| 2235 | apply (erule thin_rl) | |
| 2236 | apply (induct_tac n) | |
| 2237 | apply auto | |
| 2238 | done | |
| 2239 | ||
| 2240 | lemma bin_to_bl_bintr: | |
| 2241 | "bin_to_bl n (bintrunc m bin) = replicate (n - m) False @ bin_to_bl (min n m) bin" | |
| 2242 | unfolding bin_to_bl_def by (rule bin_to_bl_aux_bintr) | |
| 2243 | ||
| 2244 | lemma bl_to_bin_rep_False: "bl_to_bin (replicate n False) = 0" | |
| 2245 | by (induct n) auto | |
| 2246 | ||
| 2247 | lemma len_bin_to_bl_aux: "length (bin_to_bl_aux n w bs) = n + length bs" | |
| 2248 | by (fact size_bin_to_bl_aux) | |
| 2249 | ||
| 2250 | lemma len_bin_to_bl: "length (bin_to_bl n w) = n" | |
| 2251 | by (fact size_bin_to_bl) (* FIXME: duplicate *) | |
| 2252 | ||
| 2253 | lemma sign_bl_bin': "bin_sign (bl_to_bin_aux bs w) = bin_sign w" | |
| 2254 | by (induct bs arbitrary: w) auto | |
| 2255 | ||
| 2256 | lemma sign_bl_bin: "bin_sign (bl_to_bin bs) = 0" | |
| 2257 | by (simp add: bl_to_bin_def sign_bl_bin') | |
| 2258 | ||
| 2259 | lemma bl_sbin_sign_aux: "hd (bin_to_bl_aux (Suc n) w bs) = (bin_sign (sbintrunc n w) = -1)" | |
| 2260 | apply (induct n arbitrary: w bs) | |
| 2261 | apply clarsimp | |
| 2262 | apply (cases w rule: bin_exhaust) | |
| 2263 | apply simp | |
| 2264 | done | |
| 2265 | ||
| 2266 | lemma bl_sbin_sign: "hd (bin_to_bl (Suc n) w) = (bin_sign (sbintrunc n w) = -1)" | |
| 2267 | unfolding bin_to_bl_def by (rule bl_sbin_sign_aux) | |
| 2268 | ||
| 2269 | lemma bin_nth_of_bl_aux: | |
| 2270 | "bin_nth (bl_to_bin_aux bl w) n = | |
| 2271 | (n < size bl \<and> rev bl ! n \<or> n \<ge> length bl \<and> bin_nth w (n - size bl))" | |
| 2272 | apply (induct bl arbitrary: w) | |
| 2273 | apply clarsimp | |
| 2274 | apply clarsimp | |
| 2275 | apply (cut_tac x=n and y="size bl" in linorder_less_linear) | |
| 2276 | apply (erule disjE, simp add: nth_append)+ | |
| 2277 | apply auto | |
| 2278 | done | |
| 2279 | ||
| 2280 | lemma bin_nth_of_bl: "bin_nth (bl_to_bin bl) n = (n < length bl \<and> rev bl ! n)" | |
| 2281 | by (simp add: bl_to_bin_def bin_nth_of_bl_aux) | |
| 2282 | ||
| 2283 | lemma bin_nth_bl: "n < m \<Longrightarrow> bin_nth w n = nth (rev (bin_to_bl m w)) n" | |
| 2284 | apply (induct n arbitrary: m w) | |
| 2285 | apply clarsimp | |
| 2286 | apply (case_tac m, clarsimp) | |
| 2287 | apply (clarsimp simp: bin_to_bl_def) | |
| 2288 | apply (simp add: bin_to_bl_aux_alt) | |
| 2289 | apply clarsimp | |
| 2290 | apply (case_tac m, clarsimp) | |
| 2291 | apply (clarsimp simp: bin_to_bl_def) | |
| 2292 | apply (simp add: bin_to_bl_aux_alt) | |
| 2293 | done | |
| 2294 | ||
| 2295 | lemma nth_bin_to_bl_aux: | |
| 2296 | "n < m + length bl \<Longrightarrow> (bin_to_bl_aux m w bl) ! n = | |
| 2297 | (if n < m then bin_nth w (m - 1 - n) else bl ! (n - m))" | |
| 2298 | apply (induct m arbitrary: w n bl) | |
| 2299 | apply clarsimp | |
| 2300 | apply clarsimp | |
| 2301 | apply (case_tac w rule: bin_exhaust) | |
| 2302 | apply simp | |
| 2303 | done | |
| 2304 | ||
| 2305 | lemma nth_bin_to_bl: "n < m \<Longrightarrow> (bin_to_bl m w) ! n = bin_nth w (m - Suc n)" | |
| 2306 | by (simp add: bin_to_bl_def nth_bin_to_bl_aux) | |
| 2307 | ||
| 2308 | lemma bl_to_bin_lt2p_aux: "bl_to_bin_aux bs w < (w + 1) * (2 ^ length bs)" | |
| 2309 | apply (induct bs arbitrary: w) | |
| 2310 | apply clarsimp | |
| 2311 | apply clarsimp | |
| 2312 | apply (drule meta_spec, erule xtrans(8) [rotated], simp add: Bit_def)+ | |
| 2313 | done | |
| 2314 | ||
| 2315 | lemma bl_to_bin_lt2p_drop: "bl_to_bin bs < 2 ^ length (dropWhile Not bs)" | |
| 2316 | proof (induct bs) | |
| 2317 | case Nil | |
| 2318 | then show ?case by simp | |
| 2319 | next | |
| 2320 | case (Cons b bs) | |
| 2321 | with bl_to_bin_lt2p_aux[where w=1] show ?case | |
| 2322 | by (simp add: bl_to_bin_def) | |
| 2323 | qed | |
| 2324 | ||
| 2325 | lemma bl_to_bin_lt2p: "bl_to_bin bs < 2 ^ length bs" | |
| 2326 | by (metis bin_bl_bin bintr_lt2p bl_bin_bl) | |
| 2327 | ||
| 2328 | lemma bl_to_bin_ge2p_aux: "bl_to_bin_aux bs w \<ge> w * (2 ^ length bs)" | |
| 2329 | apply (induct bs arbitrary: w) | |
| 2330 | apply clarsimp | |
| 2331 | apply clarsimp | |
| 2332 | apply (drule meta_spec, erule order_trans [rotated], | |
| 2333 | simp add: Bit_B0_2t Bit_B1_2t algebra_simps)+ | |
| 2334 | apply (simp add: Bit_def) | |
| 2335 | done | |
| 2336 | ||
| 2337 | lemma bl_to_bin_ge0: "bl_to_bin bs \<ge> 0" | |
| 2338 | apply (unfold bl_to_bin_def) | |
| 2339 | apply (rule xtrans(4)) | |
| 2340 | apply (rule bl_to_bin_ge2p_aux) | |
| 2341 | apply simp | |
| 2342 | done | |
| 2343 | ||
| 2344 | lemma butlast_rest_bin: "butlast (bin_to_bl n w) = bin_to_bl (n - 1) (bin_rest w)" | |
| 2345 | apply (unfold bin_to_bl_def) | |
| 2346 | apply (cases w rule: bin_exhaust) | |
| 2347 | apply (cases n, clarsimp) | |
| 2348 | apply clarsimp | |
| 2349 | apply (auto simp add: bin_to_bl_aux_alt) | |
| 2350 | done | |
| 2351 | ||
| 2352 | lemma butlast_bin_rest: "butlast bl = bin_to_bl (length bl - Suc 0) (bin_rest (bl_to_bin bl))" | |
| 2353 | using butlast_rest_bin [where w="bl_to_bin bl" and n="length bl"] by simp | |
| 2354 | ||
| 2355 | lemma butlast_rest_bl2bin_aux: | |
| 2356 | "bl \<noteq> [] \<Longrightarrow> bl_to_bin_aux (butlast bl) w = bin_rest (bl_to_bin_aux bl w)" | |
| 2357 | by (induct bl arbitrary: w) auto | |
| 2358 | ||
| 2359 | lemma butlast_rest_bl2bin: "bl_to_bin (butlast bl) = bin_rest (bl_to_bin bl)" | |
| 2360 | by (cases bl) (auto simp: bl_to_bin_def butlast_rest_bl2bin_aux) | |
| 2361 | ||
| 2362 | lemma trunc_bl2bin_aux: | |
| 2363 | "bintrunc m (bl_to_bin_aux bl w) = | |
| 2364 | bl_to_bin_aux (drop (length bl - m) bl) (bintrunc (m - length bl) w)" | |
| 2365 | proof (induct bl arbitrary: w) | |
| 2366 | case Nil | |
| 2367 | show ?case by simp | |
| 2368 | next | |
| 2369 | case (Cons b bl) | |
| 2370 | show ?case | |
| 2371 | proof (cases "m - length bl") | |
| 2372 | case 0 | |
| 2373 | then have "Suc (length bl) - m = Suc (length bl - m)" by simp | |
| 2374 | with Cons show ?thesis by simp | |
| 2375 | next | |
| 2376 | case (Suc n) | |
| 2377 | then have "m - Suc (length bl) = n" by simp | |
| 2378 | with Cons Suc show ?thesis by simp | |
| 2379 | qed | |
| 2380 | qed | |
| 2381 | ||
| 2382 | lemma trunc_bl2bin: "bintrunc m (bl_to_bin bl) = bl_to_bin (drop (length bl - m) bl)" | |
| 2383 | by (simp add: bl_to_bin_def trunc_bl2bin_aux) | |
| 2384 | ||
| 2385 | lemma trunc_bl2bin_len [simp]: "bintrunc (length bl) (bl_to_bin bl) = bl_to_bin bl" | |
| 2386 | by (simp add: trunc_bl2bin) | |
| 2387 | ||
| 2388 | lemma bl2bin_drop: "bl_to_bin (drop k bl) = bintrunc (length bl - k) (bl_to_bin bl)" | |
| 2389 | apply (rule trans) | |
| 2390 | prefer 2 | |
| 2391 | apply (rule trunc_bl2bin [symmetric]) | |
| 2392 | apply (cases "k \<le> length bl") | |
| 2393 | apply auto | |
| 2394 | done | |
| 2395 | ||
| 2396 | lemma take_rest_power_bin: "m \<le> n \<Longrightarrow> take m (bin_to_bl n w) = bin_to_bl m ((bin_rest ^^ (n - m)) w)" | |
| 2397 | apply (rule nth_equalityI) | |
| 2398 | apply simp | |
| 2399 | apply (clarsimp simp add: nth_bin_to_bl nth_rest_power_bin) | |
| 2400 | done | |
| 2401 | ||
| 2402 | lemma last_bin_last': "size xs > 0 \<Longrightarrow> last xs \<longleftrightarrow> bin_last (bl_to_bin_aux xs w)" | |
| 2403 | by (induct xs arbitrary: w) auto | |
| 2404 | ||
| 2405 | lemma last_bin_last: "size xs > 0 \<Longrightarrow> last xs \<longleftrightarrow> bin_last (bl_to_bin xs)" | |
| 2406 | unfolding bl_to_bin_def by (erule last_bin_last') | |
| 2407 | ||
| 2408 | lemma bin_last_last: "bin_last w \<longleftrightarrow> last (bin_to_bl (Suc n) w)" | |
| 2409 | by (simp add: bin_to_bl_def) (auto simp: bin_to_bl_aux_alt) | |
| 2410 | ||
| 2411 | lemma drop_bin2bl_aux: | |
| 2412 | "drop m (bin_to_bl_aux n bin bs) = | |
| 2413 | bin_to_bl_aux (n - m) bin (drop (m - n) bs)" | |
| 2414 | apply (induct n arbitrary: m bin bs, clarsimp) | |
| 2415 | apply clarsimp | |
| 2416 | apply (case_tac bin rule: bin_exhaust) | |
| 2417 | apply (case_tac "m \<le> n", simp) | |
| 2418 | apply (case_tac "m - n", simp) | |
| 2419 | apply simp | |
| 2420 | apply (rule_tac f = "\<lambda>nat. drop nat bs" in arg_cong) | |
| 2421 | apply simp | |
| 2422 | done | |
| 2423 | ||
| 2424 | lemma drop_bin2bl: "drop m (bin_to_bl n bin) = bin_to_bl (n - m) bin" | |
| 2425 | by (simp add: bin_to_bl_def drop_bin2bl_aux) | |
| 2426 | ||
| 2427 | lemma take_bin2bl_lem1: "take m (bin_to_bl_aux m w bs) = bin_to_bl m w" | |
| 2428 | apply (induct m arbitrary: w bs) | |
| 2429 | apply clarsimp | |
| 2430 | apply clarsimp | |
| 2431 | apply (simp add: bin_to_bl_aux_alt) | |
| 2432 | apply (simp add: bin_to_bl_def) | |
| 2433 | apply (simp add: bin_to_bl_aux_alt) | |
| 2434 | done | |
| 2435 | ||
| 2436 | lemma take_bin2bl_lem: "take m (bin_to_bl_aux (m + n) w bs) = take m (bin_to_bl (m + n) w)" | |
| 2437 | by (induct n arbitrary: w bs) (simp_all (no_asm) add: bin_to_bl_def take_bin2bl_lem1, simp) | |
| 2438 | ||
| 2439 | lemma bin_split_take: "bin_split n c = (a, b) \<Longrightarrow> bin_to_bl m a = take m (bin_to_bl (m + n) c)" | |
| 2440 | apply (induct n arbitrary: b c) | |
| 2441 | apply clarsimp | |
| 2442 | apply (clarsimp simp: Let_def split: prod.split_asm) | |
| 2443 | apply (simp add: bin_to_bl_def) | |
| 2444 | apply (simp add: take_bin2bl_lem) | |
| 2445 | done | |
| 2446 | ||
| 2447 | lemma bin_split_take1: | |
| 2448 | "k = m + n \<Longrightarrow> bin_split n c = (a, b) \<Longrightarrow> bin_to_bl m a = take m (bin_to_bl k c)" | |
| 2449 | by (auto elim: bin_split_take) | |
| 2450 | ||
| 2451 | lemma takefill_bintrunc: "takefill False n bl = rev (bin_to_bl n (bl_to_bin (rev bl)))" | |
| 2452 | apply (rule nth_equalityI) | |
| 2453 | apply simp | |
| 2454 | apply (clarsimp simp: nth_takefill nth_rev nth_bin_to_bl bin_nth_of_bl) | |
| 2455 | done | |
| 2456 | ||
| 2457 | lemma bl_bin_bl_rtf: "bin_to_bl n (bl_to_bin bl) = rev (takefill False n (rev bl))" | |
| 2458 | by (simp add: takefill_bintrunc) | |
| 2459 | ||
| 2460 | lemma bl_bin_bl_rep_drop: | |
| 2461 | "bin_to_bl n (bl_to_bin bl) = | |
| 2462 | replicate (n - length bl) False @ drop (length bl - n) bl" | |
| 2463 | by (simp add: bl_bin_bl_rtf takefill_alt rev_take) | |
| 2464 | ||
| 2465 | lemma bl_to_bin_aux_cat: | |
| 2466 | "\<And>nv v. bl_to_bin_aux bs (bin_cat w nv v) = | |
| 2467 | bin_cat w (nv + length bs) (bl_to_bin_aux bs v)" | |
| 2468 | by (induct bs) (simp, simp add: bin_cat_Suc_Bit [symmetric] del: bin_cat.simps) | |
| 2469 | ||
| 2470 | lemma bin_to_bl_aux_cat: | |
| 2471 | "\<And>w bs. bin_to_bl_aux (nv + nw) (bin_cat v nw w) bs = | |
| 2472 | bin_to_bl_aux nv v (bin_to_bl_aux nw w bs)" | |
| 2473 | by (induct nw) auto | |
| 2474 | ||
| 2475 | lemma bl_to_bin_aux_alt: "bl_to_bin_aux bs w = bin_cat w (length bs) (bl_to_bin bs)" | |
| 2476 | using bl_to_bin_aux_cat [where nv = "0" and v = "0"] | |
| 2477 | by (simp add: bl_to_bin_def [symmetric]) | |
| 2478 | ||
| 2479 | lemma bin_to_bl_cat: | |
| 2480 | "bin_to_bl (nv + nw) (bin_cat v nw w) = | |
| 2481 | bin_to_bl_aux nv v (bin_to_bl nw w)" | |
| 2482 | by (simp add: bin_to_bl_def bin_to_bl_aux_cat) | |
| 2483 | ||
| 2484 | lemmas bl_to_bin_aux_app_cat = | |
| 2485 | trans [OF bl_to_bin_aux_append bl_to_bin_aux_alt] | |
| 2486 | ||
| 2487 | lemmas bin_to_bl_aux_cat_app = | |
| 2488 | trans [OF bin_to_bl_aux_cat bin_to_bl_aux_alt] | |
| 2489 | ||
| 2490 | lemma bl_to_bin_app_cat: | |
| 2491 | "bl_to_bin (bsa @ bs) = bin_cat (bl_to_bin bsa) (length bs) (bl_to_bin bs)" | |
| 2492 | by (simp only: bl_to_bin_aux_app_cat bl_to_bin_def) | |
| 2493 | ||
| 2494 | lemma bin_to_bl_cat_app: | |
| 2495 | "bin_to_bl (n + nw) (bin_cat w nw wa) = bin_to_bl n w @ bin_to_bl nw wa" | |
| 2496 | by (simp only: bin_to_bl_def bin_to_bl_aux_cat_app) | |
| 2497 | ||
| 2498 | text \<open>\<open>bl_to_bin_app_cat_alt\<close> and \<open>bl_to_bin_app_cat\<close> are easily interderivable.\<close> | |
| 2499 | lemma bl_to_bin_app_cat_alt: "bin_cat (bl_to_bin cs) n w = bl_to_bin (cs @ bin_to_bl n w)" | |
| 2500 | by (simp add: bl_to_bin_app_cat) | |
| 2501 | ||
| 2502 | lemma mask_lem: "(bl_to_bin (True # replicate n False)) = bl_to_bin (replicate n True) + 1" | |
| 2503 | apply (unfold bl_to_bin_def) | |
| 2504 | apply (induct n) | |
| 2505 | apply simp | |
| 2506 | apply (simp only: Suc_eq_plus1 replicate_add append_Cons [symmetric] bl_to_bin_aux_append) | |
| 2507 | apply (simp add: Bit_B0_2t Bit_B1_2t) | |
| 2508 | done | |
| 2509 | ||
| 2510 | primrec rbl_succ :: "bool list \<Rightarrow> bool list" | |
| 2511 | where | |
| 2512 | Nil: "rbl_succ Nil = Nil" | |
| 2513 | | Cons: "rbl_succ (x # xs) = (if x then False # rbl_succ xs else True # xs)" | |
| 2514 | ||
| 2515 | primrec rbl_pred :: "bool list \<Rightarrow> bool list" | |
| 2516 | where | |
| 2517 | Nil: "rbl_pred Nil = Nil" | |
| 2518 | | Cons: "rbl_pred (x # xs) = (if x then False # xs else True # rbl_pred xs)" | |
| 2519 | ||
| 2520 | primrec rbl_add :: "bool list \<Rightarrow> bool list \<Rightarrow> bool list" | |
| 2521 | where \<comment> \<open>result is length of first arg, second arg may be longer\<close> | |
| 2522 | Nil: "rbl_add Nil x = Nil" | |
| 2523 | | Cons: "rbl_add (y # ys) x = | |
| 2524 | (let ws = rbl_add ys (tl x) | |
| 2525 | in (y \<noteq> hd x) # (if hd x \<and> y then rbl_succ ws else ws))" | |
| 2526 | ||
| 2527 | primrec rbl_mult :: "bool list \<Rightarrow> bool list \<Rightarrow> bool list" | |
| 2528 | where \<comment> \<open>result is length of first arg, second arg may be longer\<close> | |
| 2529 | Nil: "rbl_mult Nil x = Nil" | |
| 2530 | | Cons: "rbl_mult (y # ys) x = | |
| 2531 | (let ws = False # rbl_mult ys x | |
| 2532 | in if y then rbl_add ws x else ws)" | |
| 2533 | ||
| 2534 | lemma size_rbl_pred: "length (rbl_pred bl) = length bl" | |
| 2535 | by (induct bl) auto | |
| 2536 | ||
| 2537 | lemma size_rbl_succ: "length (rbl_succ bl) = length bl" | |
| 2538 | by (induct bl) auto | |
| 2539 | ||
| 2540 | lemma size_rbl_add: "length (rbl_add bl cl) = length bl" | |
| 2541 | by (induct bl arbitrary: cl) (auto simp: Let_def size_rbl_succ) | |
| 2542 | ||
| 2543 | lemma size_rbl_mult: "length (rbl_mult bl cl) = length bl" | |
| 2544 | by (induct bl arbitrary: cl) (auto simp add: Let_def size_rbl_add) | |
| 2545 | ||
| 2546 | lemmas rbl_sizes [simp] = | |
| 2547 | size_rbl_pred size_rbl_succ size_rbl_add size_rbl_mult | |
| 2548 | ||
| 2549 | lemmas rbl_Nils = | |
| 2550 | rbl_pred.Nil rbl_succ.Nil rbl_add.Nil rbl_mult.Nil | |
| 2551 | ||
| 2552 | lemma rbl_add_app2: "length blb \<ge> length bla \<Longrightarrow> rbl_add bla (blb @ blc) = rbl_add bla blb" | |
| 2553 | apply (induct bla arbitrary: blb) | |
| 2554 | apply simp | |
| 2555 | apply clarsimp | |
| 2556 | apply (case_tac blb, clarsimp) | |
| 2557 | apply (clarsimp simp: Let_def) | |
| 2558 | done | |
| 2559 | ||
| 2560 | lemma rbl_add_take2: | |
| 2561 | "length blb \<ge> length bla \<Longrightarrow> rbl_add bla (take (length bla) blb) = rbl_add bla blb" | |
| 2562 | apply (induct bla arbitrary: blb) | |
| 2563 | apply simp | |
| 2564 | apply clarsimp | |
| 2565 | apply (case_tac blb, clarsimp) | |
| 2566 | apply (clarsimp simp: Let_def) | |
| 2567 | done | |
| 2568 | ||
| 2569 | lemma rbl_mult_app2: "length blb \<ge> length bla \<Longrightarrow> rbl_mult bla (blb @ blc) = rbl_mult bla blb" | |
| 2570 | apply (induct bla arbitrary: blb) | |
| 2571 | apply simp | |
| 2572 | apply clarsimp | |
| 2573 | apply (case_tac blb, clarsimp) | |
| 2574 | apply (clarsimp simp: Let_def rbl_add_app2) | |
| 2575 | done | |
| 2576 | ||
| 2577 | lemma rbl_mult_take2: | |
| 2578 | "length blb \<ge> length bla \<Longrightarrow> rbl_mult bla (take (length bla) blb) = rbl_mult bla blb" | |
| 2579 | apply (rule trans) | |
| 2580 | apply (rule rbl_mult_app2 [symmetric]) | |
| 2581 | apply simp | |
| 2582 | apply (rule_tac f = "rbl_mult bla" in arg_cong) | |
| 2583 | apply (rule append_take_drop_id) | |
| 2584 | done | |
| 2585 | ||
| 2586 | lemma rbl_add_split: | |
| 2587 | "P (rbl_add (y # ys) (x # xs)) = | |
| 2588 | (\<forall>ws. length ws = length ys \<longrightarrow> ws = rbl_add ys xs \<longrightarrow> | |
| 2589 | (y \<longrightarrow> ((x \<longrightarrow> P (False # rbl_succ ws)) \<and> (\<not> x \<longrightarrow> P (True # ws)))) \<and> | |
| 2590 | (\<not> y \<longrightarrow> P (x # ws)))" | |
| 2591 | by (cases y) (auto simp: Let_def) | |
| 2592 | ||
| 2593 | lemma rbl_mult_split: | |
| 2594 | "P (rbl_mult (y # ys) xs) = | |
| 2595 | (\<forall>ws. length ws = Suc (length ys) \<longrightarrow> ws = False # rbl_mult ys xs \<longrightarrow> | |
| 2596 | (y \<longrightarrow> P (rbl_add ws xs)) \<and> (\<not> y \<longrightarrow> P ws))" | |
| 2597 | by (auto simp: Let_def) | |
| 2598 | ||
| 2599 | lemma rbl_pred: "rbl_pred (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin - 1))" | |
| 2600 | apply (unfold bin_to_bl_def) | |
| 2601 | apply (induct n arbitrary: bin) | |
| 2602 | apply simp | |
| 2603 | apply clarsimp | |
| 2604 | apply (case_tac bin rule: bin_exhaust) | |
| 2605 | apply (case_tac b) | |
| 2606 | apply (clarsimp simp: bin_to_bl_aux_alt)+ | |
| 2607 | done | |
| 2608 | ||
| 2609 | lemma rbl_succ: "rbl_succ (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin + 1))" | |
| 2610 | apply (unfold bin_to_bl_def) | |
| 2611 | apply (induct n arbitrary: bin) | |
| 2612 | apply simp | |
| 2613 | apply clarsimp | |
| 2614 | apply (case_tac bin rule: bin_exhaust) | |
| 2615 | apply (case_tac b) | |
| 2616 | apply (clarsimp simp: bin_to_bl_aux_alt)+ | |
| 2617 | done | |
| 2618 | ||
| 2619 | lemma rbl_add: | |
| 2620 | "\<And>bina binb. rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = | |
| 2621 | rev (bin_to_bl n (bina + binb))" | |
| 2622 | apply (unfold bin_to_bl_def) | |
| 2623 | apply (induct n) | |
| 2624 | apply simp | |
| 2625 | apply clarsimp | |
| 2626 | apply (case_tac bina rule: bin_exhaust) | |
| 2627 | apply (case_tac binb rule: bin_exhaust) | |
| 2628 | apply (case_tac b) | |
| 2629 | apply (case_tac [!] "ba") | |
| 2630 | apply (auto simp: rbl_succ bin_to_bl_aux_alt Let_def ac_simps) | |
| 2631 | done | |
| 2632 | ||
| 2633 | lemma rbl_add_long: | |
| 2634 | "m \<ge> n \<Longrightarrow> rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = | |
| 2635 | rev (bin_to_bl n (bina + binb))" | |
| 2636 | apply (rule box_equals [OF _ rbl_add_take2 rbl_add]) | |
| 2637 | apply (rule_tac f = "rbl_add (rev (bin_to_bl n bina))" in arg_cong) | |
| 2638 | apply (rule rev_swap [THEN iffD1]) | |
| 2639 | apply (simp add: rev_take drop_bin2bl) | |
| 2640 | apply simp | |
| 2641 | done | |
| 2642 | ||
| 2643 | lemma rbl_mult_gt1: | |
| 2644 | "m \<ge> length bl \<Longrightarrow> | |
| 2645 | rbl_mult bl (rev (bin_to_bl m binb)) = | |
| 2646 | rbl_mult bl (rev (bin_to_bl (length bl) binb))" | |
| 2647 | apply (rule trans) | |
| 2648 | apply (rule rbl_mult_take2 [symmetric]) | |
| 2649 | apply simp_all | |
| 2650 | apply (rule_tac f = "rbl_mult bl" in arg_cong) | |
| 2651 | apply (rule rev_swap [THEN iffD1]) | |
| 2652 | apply (simp add: rev_take drop_bin2bl) | |
| 2653 | done | |
| 2654 | ||
| 2655 | lemma rbl_mult_gt: | |
| 2656 | "m > n \<Longrightarrow> | |
| 2657 | rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = | |
| 2658 | rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb))" | |
| 2659 | by (auto intro: trans [OF rbl_mult_gt1]) | |
| 2660 | ||
| 2661 | lemmas rbl_mult_Suc = lessI [THEN rbl_mult_gt] | |
| 2662 | ||
| 2663 | lemma rbbl_Cons: "b # rev (bin_to_bl n x) = rev (bin_to_bl (Suc n) (x BIT b))" | |
| 2664 | by (simp add: bin_to_bl_def) (simp add: bin_to_bl_aux_alt) | |
| 2665 | ||
| 2666 | lemma rbl_mult: | |
| 2667 | "rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = | |
| 2668 | rev (bin_to_bl n (bina * binb))" | |
| 2669 | apply (induct n arbitrary: bina binb) | |
| 2670 | apply simp | |
| 2671 | apply (unfold bin_to_bl_def) | |
| 2672 | apply clarsimp | |
| 2673 | apply (case_tac bina rule: bin_exhaust) | |
| 2674 | apply (case_tac binb rule: bin_exhaust) | |
| 2675 | apply (case_tac b) | |
| 2676 | apply (case_tac [!] "ba") | |
| 2677 | apply (auto simp: bin_to_bl_aux_alt Let_def) | |
| 2678 | apply (auto simp: rbbl_Cons rbl_mult_Suc rbl_add) | |
| 2679 | done | |
| 2680 | ||
| 2681 | lemma sclem: "size (concat (map (bin_to_bl n) xs)) = length xs * n" | |
| 2682 | by (induct xs) auto | |
| 2683 | ||
| 2684 | lemma bin_cat_foldl_lem: | |
| 2685 | "foldl (\<lambda>u. bin_cat u n) x xs = | |
| 2686 | bin_cat x (size xs * n) (foldl (\<lambda>u. bin_cat u n) y xs)" | |
| 2687 | apply (induct xs arbitrary: x) | |
| 2688 | apply simp | |
| 2689 | apply (simp (no_asm)) | |
| 2690 | apply (frule asm_rl) | |
| 2691 | apply (drule meta_spec) | |
| 2692 | apply (erule trans) | |
| 2693 | apply (drule_tac x = "bin_cat y n a" in meta_spec) | |
| 2694 | apply (simp add: bin_cat_assoc_sym min.absorb2) | |
| 2695 | done | |
| 2696 | ||
| 2697 | lemma bin_rcat_bl: "bin_rcat n wl = bl_to_bin (concat (map (bin_to_bl n) wl))" | |
| 2698 | apply (unfold bin_rcat_def) | |
| 2699 | apply (rule sym) | |
| 2700 | apply (induct wl) | |
| 2701 | apply (auto simp add: bl_to_bin_append) | |
| 2702 | apply (simp add: bl_to_bin_aux_alt sclem) | |
| 2703 | apply (simp add: bin_cat_foldl_lem [symmetric]) | |
| 2704 | done | |
| 2705 | ||
| 2706 | lemma bin_last_bl_to_bin: "bin_last (bl_to_bin bs) \<longleftrightarrow> bs \<noteq> [] \<and> last bs" | |
| 2707 | by(cases "bs = []")(auto simp add: bl_to_bin_def last_bin_last'[where w=0]) | |
| 2708 | ||
| 2709 | lemma bin_rest_bl_to_bin: "bin_rest (bl_to_bin bs) = bl_to_bin (butlast bs)" | |
| 2710 | by(cases "bs = []")(simp_all add: bl_to_bin_def butlast_rest_bl2bin_aux) | |
| 2711 | ||
| 2712 | lemma bl_xor_aux_bin: | |
| 2713 | "map2 (\<lambda>x y. x \<noteq> y) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = | |
| 2714 | bin_to_bl_aux n (v XOR w) (map2 (\<lambda>x y. x \<noteq> y) bs cs)" | |
| 2715 | apply (induct n arbitrary: v w bs cs) | |
| 2716 | apply simp | |
| 2717 | apply (case_tac v rule: bin_exhaust) | |
| 2718 | apply (case_tac w rule: bin_exhaust) | |
| 2719 | apply clarsimp | |
| 2720 | apply (case_tac b) | |
| 2721 | apply auto | |
| 2722 | done | |
| 2723 | ||
| 2724 | lemma bl_or_aux_bin: | |
| 2725 | "map2 (\<or>) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = | |
| 2726 | bin_to_bl_aux n (v OR w) (map2 (\<or>) bs cs)" | |
| 2727 | apply (induct n arbitrary: v w bs cs) | |
| 2728 | apply simp | |
| 2729 | apply (case_tac v rule: bin_exhaust) | |
| 2730 | apply (case_tac w rule: bin_exhaust) | |
| 2731 | apply clarsimp | |
| 2732 | done | |
| 2733 | ||
| 2734 | lemma bl_and_aux_bin: | |
| 2735 | "map2 (\<and>) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = | |
| 2736 | bin_to_bl_aux n (v AND w) (map2 (\<and>) bs cs)" | |
| 2737 | apply (induct n arbitrary: v w bs cs) | |
| 2738 | apply simp | |
| 2739 | apply (case_tac v rule: bin_exhaust) | |
| 2740 | apply (case_tac w rule: bin_exhaust) | |
| 2741 | apply clarsimp | |
| 2742 | done | |
| 2743 | ||
| 2744 | lemma bl_not_aux_bin: "map Not (bin_to_bl_aux n w cs) = bin_to_bl_aux n (NOT w) (map Not cs)" | |
| 2745 | by (induct n arbitrary: w cs) auto | |
| 2746 | ||
| 2747 | lemma bl_not_bin: "map Not (bin_to_bl n w) = bin_to_bl n (NOT w)" | |
| 2748 | by (simp add: bin_to_bl_def bl_not_aux_bin) | |
| 2749 | ||
| 2750 | lemma bl_and_bin: "map2 (\<and>) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v AND w)" | |
| 2751 | by (simp add: bin_to_bl_def bl_and_aux_bin) | |
| 2752 | ||
| 2753 | lemma bl_or_bin: "map2 (\<or>) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v OR w)" | |
| 2754 | by (simp add: bin_to_bl_def bl_or_aux_bin) | |
| 2755 | ||
| 70193 | 2756 | lemma bl_xor_bin: "map2 (\<noteq>) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v XOR w)" | 
| 2757 | using bl_xor_aux_bin by (simp add: bin_to_bl_def) | |
| 70190 | 2758 | |
| 70169 | 2759 | end |