| author | wenzelm | 
| Tue, 16 Dec 2008 16:25:19 +0100 | |
| changeset 29120 | 8a904ff43f28 | 
| parent 28959 | 9d35303719b5 | 
| child 29631 | 3aa049e5f156 | 
| permissions | -rw-r--r-- | 
| 24333 | 1 | (* | 
| 2 | ID: $Id$ | |
| 3 | Author: Jeremy Dawson, NICTA | |
| 4 | ||
| 5 | contains basic definition to do with integers | |
| 6 | expressed using Pls, Min, BIT and important resulting theorems, | |
| 7 | in particular, bin_rec and related work | |
| 8 | *) | |
| 9 | ||
| 24350 | 10 | header {* Basic Definitions for Binary Integers *}
 | 
| 11 | ||
| 26557 | 12 | theory BinGeneral | 
| 13 | imports Num_Lemmas | |
| 24333 | 14 | begin | 
| 15 | ||
| 26557 | 16 | subsection {* Further properties of numerals *}
 | 
| 17 | ||
| 18 | datatype bit = B0 | B1 | |
| 19 | ||
| 20 | definition | |
| 21 | Bit :: "int \<Rightarrow> bit \<Rightarrow> int" (infixl "BIT" 90) where | |
| 22 | "k BIT b = (case b of B0 \<Rightarrow> 0 | B1 \<Rightarrow> 1) + k + k" | |
| 23 | ||
| 24 | lemma BIT_B0_eq_Bit0 [simp]: "w BIT B0 = Int.Bit0 w" | |
| 25 | unfolding Bit_def Bit0_def by simp | |
| 26 | ||
| 27 | lemma BIT_B1_eq_Bit1 [simp]: "w BIT B1 = Int.Bit1 w" | |
| 28 | unfolding Bit_def Bit1_def by simp | |
| 29 | ||
| 30 | lemmas BIT_simps = BIT_B0_eq_Bit0 BIT_B1_eq_Bit1 | |
| 24384 
0002537695df
move BIT datatype stuff from Num_Lemmas to BinGeneral
 huffman parents: 
24383diff
changeset | 31 | |
| 26557 | 32 | hide (open) const B0 B1 | 
| 33 | ||
| 34 | lemma Min_ne_Pls [iff]: | |
| 35 | "Int.Min ~= Int.Pls" | |
| 36 | unfolding Min_def Pls_def by auto | |
| 37 | ||
| 38 | lemmas Pls_ne_Min [iff] = Min_ne_Pls [symmetric] | |
| 39 | ||
| 40 | lemmas PlsMin_defs [intro!] = | |
| 41 | Pls_def Min_def Pls_def [symmetric] Min_def [symmetric] | |
| 42 | ||
| 43 | lemmas PlsMin_simps [simp] = PlsMin_defs [THEN Eq_TrueI] | |
| 44 | ||
| 45 | lemma number_of_False_cong: | |
| 46 | "False \<Longrightarrow> number_of x = number_of y" | |
| 47 | by (rule FalseE) | |
| 48 | ||
| 49 | (** ways in which type Bin resembles a datatype **) | |
| 24350 | 50 | |
| 26557 | 51 | lemma BIT_eq: "u BIT b = v BIT c ==> u = v & b = c" | 
| 52 | apply (unfold Bit_def) | |
| 53 | apply (simp (no_asm_use) split: bit.split_asm) | |
| 54 | apply simp_all | |
| 55 | apply (drule_tac f=even in arg_cong, clarsimp)+ | |
| 56 | done | |
| 57 | ||
| 58 | lemmas BIT_eqE [elim!] = BIT_eq [THEN conjE, standard] | |
| 59 | ||
| 60 | lemma BIT_eq_iff [simp]: | |
| 61 | "(u BIT b = v BIT c) = (u = v \<and> b = c)" | |
| 62 | by (rule iffI) auto | |
| 63 | ||
| 64 | lemmas BIT_eqI [intro!] = conjI [THEN BIT_eq_iff [THEN iffD2]] | |
| 65 | ||
| 66 | lemma less_Bits: | |
| 67 | "(v BIT b < w BIT c) = (v < w | v <= w & b = bit.B0 & c = bit.B1)" | |
| 68 | unfolding Bit_def by (auto split: bit.split) | |
| 24333 | 69 | |
| 26557 | 70 | lemma le_Bits: | 
| 71 | "(v BIT b <= w BIT c) = (v < w | v <= w & (b ~= bit.B1 | c ~= bit.B0))" | |
| 72 | unfolding Bit_def by (auto split: bit.split) | |
| 73 | ||
| 74 | lemma no_no [simp]: "number_of (number_of i) = i" | |
| 75 | unfolding number_of_eq by simp | |
| 76 | ||
| 77 | lemma Bit_B0: | |
| 78 | "k BIT bit.B0 = k + k" | |
| 79 | by (unfold Bit_def) simp | |
| 80 | ||
| 81 | lemma Bit_B1: | |
| 82 | "k BIT bit.B1 = k + k + 1" | |
| 83 | by (unfold Bit_def) simp | |
| 84 | ||
| 85 | lemma Bit_B0_2t: "k BIT bit.B0 = 2 * k" | |
| 86 | by (rule trans, rule Bit_B0) simp | |
| 87 | ||
| 88 | lemma Bit_B1_2t: "k BIT bit.B1 = 2 * k + 1" | |
| 89 | by (rule trans, rule Bit_B1) simp | |
| 90 | ||
| 91 | lemma B_mod_2': | |
| 92 | "X = 2 ==> (w BIT bit.B1) mod X = 1 & (w BIT bit.B0) mod X = 0" | |
| 93 | apply (simp (no_asm) only: Bit_B0 Bit_B1) | |
| 94 | apply (simp add: z1pmod2) | |
| 24465 | 95 | done | 
| 24333 | 96 | |
| 26557 | 97 | lemma B1_mod_2 [simp]: "(Int.Bit1 w) mod 2 = 1" | 
| 98 | unfolding numeral_simps number_of_is_id by (simp add: z1pmod2) | |
| 24333 | 99 | |
| 26557 | 100 | lemma B0_mod_2 [simp]: "(Int.Bit0 w) mod 2 = 0" | 
| 101 | unfolding numeral_simps number_of_is_id by simp | |
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 102 | |
| 26557 | 103 | lemma neB1E [elim!]: | 
| 104 | assumes ne: "y \<noteq> bit.B1" | |
| 105 | assumes y: "y = bit.B0 \<Longrightarrow> P" | |
| 106 | shows "P" | |
| 107 | apply (rule y) | |
| 108 | apply (cases y rule: bit.exhaust, simp) | |
| 109 | apply (simp add: ne) | |
| 110 | done | |
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 111 | |
| 26557 | 112 | lemma bin_ex_rl: "EX w b. w BIT b = bin" | 
| 113 | apply (unfold Bit_def) | |
| 114 | apply (cases "even bin") | |
| 115 | apply (clarsimp simp: even_equiv_def) | |
| 116 | apply (auto simp: odd_equiv_def split: bit.split) | |
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 117 | done | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 118 | |
| 26557 | 119 | lemma bin_exhaust: | 
| 120 | assumes Q: "\<And>x b. bin = x BIT b \<Longrightarrow> Q" | |
| 121 | shows "Q" | |
| 122 | apply (insert bin_ex_rl [of bin]) | |
| 123 | apply (erule exE)+ | |
| 124 | apply (rule Q) | |
| 125 | apply force | |
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 126 | done | 
| 24333 | 127 | |
| 128 | ||
| 24465 | 129 | subsection {* Destructors for binary integers *}
 | 
| 24364 | 130 | |
| 26557 | 131 | definition bin_rl :: "int \<Rightarrow> int \<times> bit" where | 
| 28562 | 132 | [code del]: "bin_rl w = (THE (r, l). w = r BIT l)" | 
| 26557 | 133 | |
| 134 | lemma bin_rl_char: "(bin_rl w = (r, l)) = (r BIT l = w)" | |
| 135 | apply (unfold bin_rl_def) | |
| 136 | apply safe | |
| 137 | apply (cases w rule: bin_exhaust) | |
| 138 | apply auto | |
| 139 | done | |
| 140 | ||
| 26514 | 141 | definition | 
| 28562 | 142 | bin_rest_def [code del]: "bin_rest w = fst (bin_rl w)" | 
| 26514 | 143 | |
| 144 | definition | |
| 28562 | 145 | bin_last_def [code del] : "bin_last w = snd (bin_rl w)" | 
| 24465 | 146 | |
| 26514 | 147 | primrec bin_nth where | 
| 26557 | 148 | Z: "bin_nth w 0 = (bin_last w = bit.B1)" | 
| 149 | | Suc: "bin_nth w (Suc n) = bin_nth (bin_rest w) n" | |
| 24364 | 150 | |
| 24465 | 151 | lemma bin_rl: "bin_rl w = (bin_rest w, bin_last w)" | 
| 152 | unfolding bin_rest_def bin_last_def by auto | |
| 153 | ||
| 26557 | 154 | lemma bin_rl_simps [simp]: | 
| 155 | "bin_rl Int.Pls = (Int.Pls, bit.B0)" | |
| 156 | "bin_rl Int.Min = (Int.Min, bit.B1)" | |
| 157 | "bin_rl (Int.Bit0 r) = (r, bit.B0)" | |
| 158 | "bin_rl (Int.Bit1 r) = (r, bit.B1)" | |
| 159 | "bin_rl (r BIT b) = (r, b)" | |
| 160 | unfolding bin_rl_char by simp_all | |
| 161 | ||
| 28562 | 162 | declare bin_rl_simps(1-4) [code] | 
| 26557 | 163 | |
| 24465 | 164 | lemmas bin_rl_simp [simp] = iffD1 [OF bin_rl_char bin_rl] | 
| 24364 | 165 | |
| 26557 | 166 | lemma bin_abs_lem: | 
| 167 | "bin = (w BIT b) ==> ~ bin = Int.Min --> ~ bin = Int.Pls --> | |
| 168 | nat (abs w) < nat (abs bin)" | |
| 169 | apply (clarsimp simp add: bin_rl_char) | |
| 170 | apply (unfold Pls_def Min_def Bit_def) | |
| 171 | apply (cases b) | |
| 172 | apply (clarsimp, arith) | |
| 173 | apply (clarsimp, arith) | |
| 174 | done | |
| 175 | ||
| 176 | lemma bin_induct: | |
| 177 | assumes PPls: "P Int.Pls" | |
| 178 | and PMin: "P Int.Min" | |
| 179 | and PBit: "!!bin bit. P bin ==> P (bin BIT bit)" | |
| 180 | shows "P bin" | |
| 181 | apply (rule_tac P=P and a=bin and f1="nat o abs" | |
| 182 | in wf_measure [THEN wf_induct]) | |
| 183 | apply (simp add: measure_def inv_image_def) | |
| 184 | apply (case_tac x rule: bin_exhaust) | |
| 185 | apply (frule bin_abs_lem) | |
| 186 | apply (auto simp add : PPls PMin PBit) | |
| 187 | done | |
| 188 | ||
| 189 | lemma numeral_induct: | |
| 190 | assumes Pls: "P Int.Pls" | |
| 191 | assumes Min: "P Int.Min" | |
| 192 | assumes Bit0: "\<And>w. \<lbrakk>P w; w \<noteq> Int.Pls\<rbrakk> \<Longrightarrow> P (Int.Bit0 w)" | |
| 193 | assumes Bit1: "\<And>w. \<lbrakk>P w; w \<noteq> Int.Min\<rbrakk> \<Longrightarrow> P (Int.Bit1 w)" | |
| 194 | shows "P x" | |
| 195 | apply (induct x rule: bin_induct) | |
| 196 | apply (rule Pls) | |
| 197 | apply (rule Min) | |
| 198 | apply (case_tac bit) | |
| 199 | apply (case_tac "bin = Int.Pls") | |
| 200 | apply simp | |
| 201 | apply (simp add: Bit0) | |
| 202 | apply (case_tac "bin = Int.Min") | |
| 203 | apply simp | |
| 204 | apply (simp add: Bit1) | |
| 205 | done | |
| 206 | ||
| 24465 | 207 | lemma bin_rest_simps [simp]: | 
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25349diff
changeset | 208 | "bin_rest Int.Pls = Int.Pls" | 
| 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25349diff
changeset | 209 | "bin_rest Int.Min = Int.Min" | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 210 | "bin_rest (Int.Bit0 w) = w" | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 211 | "bin_rest (Int.Bit1 w) = w" | 
| 26514 | 212 | "bin_rest (w BIT b) = w" | 
| 24465 | 213 | unfolding bin_rest_def by auto | 
| 214 | ||
| 28562 | 215 | declare bin_rest_simps(1-4) [code] | 
| 26514 | 216 | |
| 24465 | 217 | lemma bin_last_simps [simp]: | 
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25349diff
changeset | 218 | "bin_last Int.Pls = bit.B0" | 
| 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25349diff
changeset | 219 | "bin_last Int.Min = bit.B1" | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 220 | "bin_last (Int.Bit0 w) = bit.B0" | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 221 | "bin_last (Int.Bit1 w) = bit.B1" | 
| 26514 | 222 | "bin_last (w BIT b) = b" | 
| 24465 | 223 | unfolding bin_last_def by auto | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 224 | |
| 28562 | 225 | declare bin_last_simps(1-4) [code] | 
| 26514 | 226 | |
| 24333 | 227 | lemma bin_r_l_extras [simp]: | 
| 228 | "bin_last 0 = bit.B0" | |
| 229 | "bin_last (- 1) = bit.B1" | |
| 230 | "bin_last -1 = bit.B1" | |
| 231 | "bin_last 1 = bit.B1" | |
| 232 | "bin_rest 1 = 0" | |
| 233 | "bin_rest 0 = 0" | |
| 234 | "bin_rest (- 1) = - 1" | |
| 235 | "bin_rest -1 = -1" | |
| 236 | apply (unfold number_of_Min) | |
| 237 | apply (unfold Pls_def [symmetric] Min_def [symmetric]) | |
| 238 | apply (unfold numeral_1_eq_1 [symmetric]) | |
| 239 | apply (auto simp: number_of_eq) | |
| 240 | done | |
| 241 | ||
| 242 | lemma bin_last_mod: | |
| 243 | "bin_last w = (if w mod 2 = 0 then bit.B0 else bit.B1)" | |
| 24465 | 244 | apply (case_tac w rule: bin_exhaust) | 
| 245 | apply (case_tac b) | |
| 246 | apply auto | |
| 24333 | 247 | done | 
| 248 | ||
| 249 | lemma bin_rest_div: | |
| 250 | "bin_rest w = w div 2" | |
| 24465 | 251 | apply (case_tac w rule: bin_exhaust) | 
| 252 | apply (rule trans) | |
| 253 | apply clarsimp | |
| 254 | apply (rule refl) | |
| 255 | apply (drule trans) | |
| 256 | apply (rule Bit_def) | |
| 257 | apply (simp add: z1pdiv2 split: bit.split) | |
| 24333 | 258 | done | 
| 259 | ||
| 260 | lemma Bit_div2 [simp]: "(w BIT b) div 2 = w" | |
| 261 | unfolding bin_rest_div [symmetric] by auto | |
| 262 | ||
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 263 | lemma Bit0_div2 [simp]: "(Int.Bit0 w) div 2 = w" | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 264 | using Bit_div2 [where b=bit.B0] by simp | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 265 | |
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 266 | lemma Bit1_div2 [simp]: "(Int.Bit1 w) div 2 = w" | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 267 | using Bit_div2 [where b=bit.B1] by simp | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 268 | |
| 24333 | 269 | lemma bin_nth_lem [rule_format]: | 
| 270 | "ALL y. bin_nth x = bin_nth y --> x = y" | |
| 271 | apply (induct x rule: bin_induct) | |
| 272 | apply safe | |
| 273 | apply (erule rev_mp) | |
| 274 | apply (induct_tac y rule: bin_induct) | |
| 26827 
a62f8db42f4a
Deleted subset_antisym in a few proofs, because it is
 berghofe parents: 
26557diff
changeset | 275 | apply (safe del: subset_antisym) | 
| 24333 | 276 | apply (drule_tac x=0 in fun_cong, force) | 
| 277 | apply (erule notE, rule ext, | |
| 278 | drule_tac x="Suc x" in fun_cong, force) | |
| 279 | apply (drule_tac x=0 in fun_cong, force) | |
| 280 | apply (erule rev_mp) | |
| 281 | apply (induct_tac y rule: bin_induct) | |
| 26827 
a62f8db42f4a
Deleted subset_antisym in a few proofs, because it is
 berghofe parents: 
26557diff
changeset | 282 | apply (safe del: subset_antisym) | 
| 24333 | 283 | apply (drule_tac x=0 in fun_cong, force) | 
| 284 | apply (erule notE, rule ext, | |
| 285 | drule_tac x="Suc x" in fun_cong, force) | |
| 286 | apply (drule_tac x=0 in fun_cong, force) | |
| 287 | apply (case_tac y rule: bin_exhaust) | |
| 288 | apply clarify | |
| 289 | apply (erule allE) | |
| 290 | apply (erule impE) | |
| 291 | prefer 2 | |
| 292 | apply (erule BIT_eqI) | |
| 293 | apply (drule_tac x=0 in fun_cong, force) | |
| 294 | apply (rule ext) | |
| 295 | apply (drule_tac x="Suc ?x" in fun_cong, force) | |
| 296 | done | |
| 297 | ||
| 298 | lemma bin_nth_eq_iff: "(bin_nth x = bin_nth y) = (x = y)" | |
| 299 | by (auto elim: bin_nth_lem) | |
| 300 | ||
| 301 | lemmas bin_eqI = ext [THEN bin_nth_eq_iff [THEN iffD1], standard] | |
| 302 | ||
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25349diff
changeset | 303 | lemma bin_nth_Pls [simp]: "~ bin_nth Int.Pls n" | 
| 24333 | 304 | by (induct n) auto | 
| 305 | ||
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25349diff
changeset | 306 | lemma bin_nth_Min [simp]: "bin_nth Int.Min n" | 
| 24333 | 307 | by (induct n) auto | 
| 308 | ||
| 309 | lemma bin_nth_0_BIT: "bin_nth (w BIT b) 0 = (b = bit.B1)" | |
| 310 | by auto | |
| 311 | ||
| 312 | lemma bin_nth_Suc_BIT: "bin_nth (w BIT b) (Suc n) = bin_nth w n" | |
| 313 | by auto | |
| 314 | ||
| 315 | lemma bin_nth_minus [simp]: "0 < n ==> bin_nth (w BIT b) n = bin_nth w (n - 1)" | |
| 316 | by (cases n) auto | |
| 317 | ||
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 318 | lemma bin_nth_minus_Bit0 [simp]: | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 319 | "0 < n ==> bin_nth (Int.Bit0 w) n = bin_nth w (n - 1)" | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 320 | using bin_nth_minus [where b=bit.B0] by simp | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 321 | |
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 322 | lemma bin_nth_minus_Bit1 [simp]: | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 323 | "0 < n ==> bin_nth (Int.Bit1 w) n = bin_nth w (n - 1)" | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 324 | using bin_nth_minus [where b=bit.B1] by simp | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 325 | |
| 24333 | 326 | lemmas bin_nth_0 = bin_nth.simps(1) | 
| 327 | lemmas bin_nth_Suc = bin_nth.simps(2) | |
| 328 | ||
| 329 | lemmas bin_nth_simps = | |
| 330 | bin_nth_0 bin_nth_Suc bin_nth_Pls bin_nth_Min bin_nth_minus | |
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 331 | bin_nth_minus_Bit0 bin_nth_minus_Bit1 | 
| 24333 | 332 | |
| 26557 | 333 | |
| 334 | subsection {* Recursion combinator for binary integers *}
 | |
| 335 | ||
| 336 | lemma brlem: "(bin = Int.Min) = (- bin + Int.pred 0 = 0)" | |
| 337 | unfolding Min_def pred_def by arith | |
| 338 | ||
| 339 | function | |
| 340 | bin_rec :: "'a \<Rightarrow> 'a \<Rightarrow> (int \<Rightarrow> bit \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> int \<Rightarrow> 'a" | |
| 341 | where | |
| 342 | "bin_rec f1 f2 f3 bin = (if bin = Int.Pls then f1 | |
| 343 | else if bin = Int.Min then f2 | |
| 344 | else case bin_rl bin of (w, b) => f3 w b (bin_rec f1 f2 f3 w))" | |
| 345 | by pat_completeness auto | |
| 346 | ||
| 347 | termination | |
| 348 | apply (relation "measure (nat o abs o snd o snd o snd)") | |
| 349 | apply simp | |
| 350 | apply (simp add: Pls_def brlem) | |
| 351 | apply (clarsimp simp: bin_rl_char pred_def) | |
| 352 | apply (frule thin_rl [THEN refl [THEN bin_abs_lem [rule_format]]]) | |
| 353 | apply (unfold Pls_def Min_def number_of_eq) | |
| 354 | prefer 2 | |
| 355 | apply (erule asm_rl) | |
| 356 | apply auto | |
| 357 | done | |
| 358 | ||
| 359 | declare bin_rec.simps [simp del] | |
| 360 | ||
| 361 | lemma bin_rec_PM: | |
| 362 | "f = bin_rec f1 f2 f3 ==> f Int.Pls = f1 & f Int.Min = f2" | |
| 363 | by (auto simp add: bin_rec.simps) | |
| 364 | ||
| 365 | lemma bin_rec_Pls: "bin_rec f1 f2 f3 Int.Pls = f1" | |
| 366 | by (simp add: bin_rec.simps) | |
| 367 | ||
| 368 | lemma bin_rec_Min: "bin_rec f1 f2 f3 Int.Min = f2" | |
| 369 | by (simp add: bin_rec.simps) | |
| 370 | ||
| 371 | lemma bin_rec_Bit0: | |
| 372 | "f3 Int.Pls bit.B0 f1 = f1 \<Longrightarrow> | |
| 373 | bin_rec f1 f2 f3 (Int.Bit0 w) = f3 w bit.B0 (bin_rec f1 f2 f3 w)" | |
| 28959 | 374 | by (simp add: bin_rec_Pls bin_rec.simps [of _ _ _ "Int.Bit0 w"]) | 
| 26557 | 375 | |
| 376 | lemma bin_rec_Bit1: | |
| 377 | "f3 Int.Min bit.B1 f2 = f2 \<Longrightarrow> | |
| 378 | bin_rec f1 f2 f3 (Int.Bit1 w) = f3 w bit.B1 (bin_rec f1 f2 f3 w)" | |
| 28959 | 379 | by (simp add: bin_rec_Min bin_rec.simps [of _ _ _ "Int.Bit1 w"]) | 
| 26557 | 380 | |
| 381 | lemma bin_rec_Bit: | |
| 382 | "f = bin_rec f1 f2 f3 ==> f3 Int.Pls bit.B0 f1 = f1 ==> | |
| 383 | f3 Int.Min bit.B1 f2 = f2 ==> f (w BIT b) = f3 w b (f w)" | |
| 384 | by (cases b, simp add: bin_rec_Bit0, simp add: bin_rec_Bit1) | |
| 385 | ||
| 386 | lemmas bin_rec_simps = refl [THEN bin_rec_Bit] bin_rec_Pls bin_rec_Min | |
| 387 | bin_rec_Bit0 bin_rec_Bit1 | |
| 388 | ||
| 389 | ||
| 390 | subsection {* Truncating binary integers *}
 | |
| 391 | ||
| 392 | definition | |
| 28562 | 393 | bin_sign_def [code del] : "bin_sign = bin_rec Int.Pls Int.Min (%w b s. s)" | 
| 26557 | 394 | |
| 395 | lemma bin_sign_simps [simp]: | |
| 396 | "bin_sign Int.Pls = Int.Pls" | |
| 397 | "bin_sign Int.Min = Int.Min" | |
| 398 | "bin_sign (Int.Bit0 w) = bin_sign w" | |
| 399 | "bin_sign (Int.Bit1 w) = bin_sign w" | |
| 400 | "bin_sign (w BIT b) = bin_sign w" | |
| 401 | unfolding bin_sign_def by (auto simp: bin_rec_simps) | |
| 402 | ||
| 28562 | 403 | declare bin_sign_simps(1-4) [code] | 
| 26557 | 404 | |
| 24364 | 405 | lemma bin_sign_rest [simp]: | 
| 406 | "bin_sign (bin_rest w) = (bin_sign w)" | |
| 26557 | 407 | by (cases w rule: bin_exhaust) auto | 
| 24364 | 408 | |
| 409 | consts | |
| 410 | bintrunc :: "nat => int => int" | |
| 411 | primrec | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25349diff
changeset | 412 | Z : "bintrunc 0 bin = Int.Pls" | 
| 24364 | 413 | Suc : "bintrunc (Suc n) bin = bintrunc n (bin_rest bin) BIT (bin_last bin)" | 
| 414 | ||
| 415 | consts | |
| 416 | sbintrunc :: "nat => int => int" | |
| 417 | primrec | |
| 418 | Z : "sbintrunc 0 bin = | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25349diff
changeset | 419 | (case bin_last bin of bit.B1 => Int.Min | bit.B0 => Int.Pls)" | 
| 24364 | 420 | Suc : "sbintrunc (Suc n) bin = sbintrunc n (bin_rest bin) BIT (bin_last bin)" | 
| 421 | ||
| 24333 | 422 | lemma sign_bintr: | 
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25349diff
changeset | 423 | "!!w. bin_sign (bintrunc n w) = Int.Pls" | 
| 24333 | 424 | by (induct n) auto | 
| 425 | ||
| 426 | lemma bintrunc_mod2p: | |
| 427 | "!!w. bintrunc n w = (w mod 2 ^ n :: int)" | |
| 428 | apply (induct n, clarsimp) | |
| 429 | apply (simp add: bin_last_mod bin_rest_div Bit_def zmod_zmult2_eq | |
| 430 | cong: number_of_False_cong) | |
| 431 | done | |
| 432 | ||
| 433 | lemma sbintrunc_mod2p: | |
| 434 | "!!w. sbintrunc n w = ((w + 2 ^ n) mod 2 ^ (Suc n) - 2 ^ n :: int)" | |
| 435 | apply (induct n) | |
| 436 | apply clarsimp | |
| 437 | apply (subst zmod_zadd_left_eq) | |
| 438 | apply (simp add: bin_last_mod) | |
| 439 | apply (simp add: number_of_eq) | |
| 440 | apply clarsimp | |
| 441 | apply (simp add: bin_last_mod bin_rest_div Bit_def | |
| 442 | cong: number_of_False_cong) | |
| 443 | apply (clarsimp simp: zmod_zmult_zmult1 [symmetric] | |
| 444 | zmod_zdiv_equality [THEN diff_eq_eq [THEN iffD2 [THEN sym]]]) | |
| 445 | apply (rule trans [symmetric, OF _ emep1]) | |
| 446 | apply auto | |
| 447 | apply (auto simp: even_def) | |
| 448 | done | |
| 449 | ||
| 24465 | 450 | subsection "Simplifications for (s)bintrunc" | 
| 451 | ||
| 452 | lemma bit_bool: | |
| 453 | "(b = (b' = bit.B1)) = (b' = (if b then bit.B1 else bit.B0))" | |
| 454 | by (cases b') auto | |
| 455 | ||
| 456 | lemmas bit_bool1 [simp] = refl [THEN bit_bool [THEN iffD1], symmetric] | |
| 24333 | 457 | |
| 458 | lemma bin_sign_lem: | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25349diff
changeset | 459 | "!!bin. (bin_sign (sbintrunc n bin) = Int.Min) = bin_nth bin n" | 
| 24333 | 460 | apply (induct n) | 
| 461 | apply (case_tac bin rule: bin_exhaust, case_tac b, auto)+ | |
| 462 | done | |
| 463 | ||
| 464 | lemma nth_bintr: | |
| 465 | "!!w m. bin_nth (bintrunc m w) n = (n < m & bin_nth w n)" | |
| 466 | apply (induct n) | |
| 467 | apply (case_tac m, auto)[1] | |
| 468 | apply (case_tac m, auto)[1] | |
| 469 | done | |
| 470 | ||
| 471 | lemma nth_sbintr: | |
| 472 | "!!w m. bin_nth (sbintrunc m w) n = | |
| 473 | (if n < m then bin_nth w n else bin_nth w m)" | |
| 474 | apply (induct n) | |
| 475 | apply (case_tac m, simp_all split: bit.splits)[1] | |
| 476 | apply (case_tac m, simp_all split: bit.splits)[1] | |
| 477 | done | |
| 478 | ||
| 479 | lemma bin_nth_Bit: | |
| 480 | "bin_nth (w BIT b) n = (n = 0 & b = bit.B1 | (EX m. n = Suc m & bin_nth w m))" | |
| 481 | by (cases n) auto | |
| 482 | ||
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 483 | lemma bin_nth_Bit0: | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 484 | "bin_nth (Int.Bit0 w) n = (EX m. n = Suc m & bin_nth w m)" | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 485 | using bin_nth_Bit [where b=bit.B0] by simp | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 486 | |
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 487 | lemma bin_nth_Bit1: | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 488 | "bin_nth (Int.Bit1 w) n = (n = 0 | (EX m. n = Suc m & bin_nth w m))" | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 489 | using bin_nth_Bit [where b=bit.B1] by simp | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 490 | |
| 24333 | 491 | lemma bintrunc_bintrunc_l: | 
| 492 | "n <= m ==> (bintrunc m (bintrunc n w) = bintrunc n w)" | |
| 493 | by (rule bin_eqI) (auto simp add : nth_bintr) | |
| 494 | ||
| 495 | lemma sbintrunc_sbintrunc_l: | |
| 496 | "n <= m ==> (sbintrunc m (sbintrunc n w) = sbintrunc n w)" | |
| 497 | by (rule bin_eqI) (auto simp: nth_sbintr min_def) | |
| 498 | ||
| 499 | lemma bintrunc_bintrunc_ge: | |
| 500 | "n <= m ==> (bintrunc n (bintrunc m w) = bintrunc n w)" | |
| 501 | by (rule bin_eqI) (auto simp: nth_bintr) | |
| 502 | ||
| 503 | lemma bintrunc_bintrunc_min [simp]: | |
| 504 | "bintrunc m (bintrunc n w) = bintrunc (min m n) w" | |
| 505 | apply (unfold min_def) | |
| 506 | apply (rule bin_eqI) | |
| 507 | apply (auto simp: nth_bintr) | |
| 508 | done | |
| 509 | ||
| 510 | lemma sbintrunc_sbintrunc_min [simp]: | |
| 511 | "sbintrunc m (sbintrunc n w) = sbintrunc (min m n) w" | |
| 512 | apply (unfold min_def) | |
| 513 | apply (rule bin_eqI) | |
| 514 | apply (auto simp: nth_sbintr) | |
| 515 | done | |
| 516 | ||
| 517 | lemmas bintrunc_Pls = | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25349diff
changeset | 518 | bintrunc.Suc [where bin="Int.Pls", simplified bin_last_simps bin_rest_simps, standard] | 
| 24333 | 519 | |
| 520 | lemmas bintrunc_Min [simp] = | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25349diff
changeset | 521 | bintrunc.Suc [where bin="Int.Min", simplified bin_last_simps bin_rest_simps, standard] | 
| 24333 | 522 | |
| 523 | lemmas bintrunc_BIT [simp] = | |
| 25349 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 524 | bintrunc.Suc [where bin="w BIT b", simplified bin_last_simps bin_rest_simps, standard] | 
| 24333 | 525 | |
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 526 | lemma bintrunc_Bit0 [simp]: | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 527 | "bintrunc (Suc n) (Int.Bit0 w) = Int.Bit0 (bintrunc n w)" | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 528 | using bintrunc_BIT [where b=bit.B0] by simp | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 529 | |
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 530 | lemma bintrunc_Bit1 [simp]: | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 531 | "bintrunc (Suc n) (Int.Bit1 w) = Int.Bit1 (bintrunc n w)" | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 532 | using bintrunc_BIT [where b=bit.B1] by simp | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 533 | |
| 24333 | 534 | lemmas bintrunc_Sucs = bintrunc_Pls bintrunc_Min bintrunc_BIT | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 535 | bintrunc_Bit0 bintrunc_Bit1 | 
| 24333 | 536 | |
| 537 | lemmas sbintrunc_Suc_Pls = | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25349diff
changeset | 538 | sbintrunc.Suc [where bin="Int.Pls", simplified bin_last_simps bin_rest_simps, standard] | 
| 24333 | 539 | |
| 540 | lemmas sbintrunc_Suc_Min = | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25349diff
changeset | 541 | sbintrunc.Suc [where bin="Int.Min", simplified bin_last_simps bin_rest_simps, standard] | 
| 24333 | 542 | |
| 543 | lemmas sbintrunc_Suc_BIT [simp] = | |
| 25349 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 544 | sbintrunc.Suc [where bin="w BIT b", simplified bin_last_simps bin_rest_simps, standard] | 
| 24333 | 545 | |
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 546 | lemma sbintrunc_Suc_Bit0 [simp]: | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 547 | "sbintrunc (Suc n) (Int.Bit0 w) = Int.Bit0 (sbintrunc n w)" | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 548 | using sbintrunc_Suc_BIT [where b=bit.B0] by simp | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 549 | |
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 550 | lemma sbintrunc_Suc_Bit1 [simp]: | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 551 | "sbintrunc (Suc n) (Int.Bit1 w) = Int.Bit1 (sbintrunc n w)" | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 552 | using sbintrunc_Suc_BIT [where b=bit.B1] by simp | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 553 | |
| 24333 | 554 | lemmas sbintrunc_Sucs = sbintrunc_Suc_Pls sbintrunc_Suc_Min sbintrunc_Suc_BIT | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 555 | sbintrunc_Suc_Bit0 sbintrunc_Suc_Bit1 | 
| 24333 | 556 | |
| 557 | lemmas sbintrunc_Pls = | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25349diff
changeset | 558 | sbintrunc.Z [where bin="Int.Pls", | 
| 25349 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 559 | simplified bin_last_simps bin_rest_simps bit.simps, standard] | 
| 24333 | 560 | |
| 561 | lemmas sbintrunc_Min = | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25349diff
changeset | 562 | sbintrunc.Z [where bin="Int.Min", | 
| 25349 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 563 | simplified bin_last_simps bin_rest_simps bit.simps, standard] | 
| 24333 | 564 | |
| 565 | lemmas sbintrunc_0_BIT_B0 [simp] = | |
| 25349 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 566 | sbintrunc.Z [where bin="w BIT bit.B0", | 
| 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 567 | simplified bin_last_simps bin_rest_simps bit.simps, standard] | 
| 24333 | 568 | |
| 569 | lemmas sbintrunc_0_BIT_B1 [simp] = | |
| 25349 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 570 | sbintrunc.Z [where bin="w BIT bit.B1", | 
| 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 571 | simplified bin_last_simps bin_rest_simps bit.simps, standard] | 
| 24333 | 572 | |
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 573 | lemma sbintrunc_0_Bit0 [simp]: "sbintrunc 0 (Int.Bit0 w) = Int.Pls" | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 574 | using sbintrunc_0_BIT_B0 by simp | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 575 | |
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 576 | lemma sbintrunc_0_Bit1 [simp]: "sbintrunc 0 (Int.Bit1 w) = Int.Min" | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 577 | using sbintrunc_0_BIT_B1 by simp | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 578 | |
| 24333 | 579 | lemmas sbintrunc_0_simps = | 
| 580 | sbintrunc_Pls sbintrunc_Min sbintrunc_0_BIT_B0 sbintrunc_0_BIT_B1 | |
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 581 | sbintrunc_0_Bit0 sbintrunc_0_Bit1 | 
| 24333 | 582 | |
| 583 | lemmas bintrunc_simps = bintrunc.Z bintrunc_Sucs | |
| 584 | lemmas sbintrunc_simps = sbintrunc_0_simps sbintrunc_Sucs | |
| 585 | ||
| 586 | lemma bintrunc_minus: | |
| 587 | "0 < n ==> bintrunc (Suc (n - 1)) w = bintrunc n w" | |
| 588 | by auto | |
| 589 | ||
| 590 | lemma sbintrunc_minus: | |
| 591 | "0 < n ==> sbintrunc (Suc (n - 1)) w = sbintrunc n w" | |
| 592 | by auto | |
| 593 | ||
| 594 | lemmas bintrunc_minus_simps = | |
| 595 | bintrunc_Sucs [THEN [2] bintrunc_minus [symmetric, THEN trans], standard] | |
| 596 | lemmas sbintrunc_minus_simps = | |
| 597 | sbintrunc_Sucs [THEN [2] sbintrunc_minus [symmetric, THEN trans], standard] | |
| 598 | ||
| 599 | lemma bintrunc_n_Pls [simp]: | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25349diff
changeset | 600 | "bintrunc n Int.Pls = Int.Pls" | 
| 24333 | 601 | by (induct n) auto | 
| 602 | ||
| 603 | lemma sbintrunc_n_PM [simp]: | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25349diff
changeset | 604 | "sbintrunc n Int.Pls = Int.Pls" | 
| 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25349diff
changeset | 605 | "sbintrunc n Int.Min = Int.Min" | 
| 24333 | 606 | by (induct n) auto | 
| 607 | ||
| 25349 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 608 | lemmas thobini1 = arg_cong [where f = "%w. w BIT b", standard] | 
| 24333 | 609 | |
| 610 | lemmas bintrunc_BIT_I = trans [OF bintrunc_BIT thobini1] | |
| 611 | lemmas bintrunc_Min_I = trans [OF bintrunc_Min thobini1] | |
| 612 | ||
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 613 | lemmas bmsts = bintrunc_minus_simps(1-3) [THEN thobini1 [THEN [2] trans], standard] | 
| 24333 | 614 | lemmas bintrunc_Pls_minus_I = bmsts(1) | 
| 615 | lemmas bintrunc_Min_minus_I = bmsts(2) | |
| 616 | lemmas bintrunc_BIT_minus_I = bmsts(3) | |
| 617 | ||
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25349diff
changeset | 618 | lemma bintrunc_0_Min: "bintrunc 0 Int.Min = Int.Pls" | 
| 24333 | 619 | by auto | 
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25349diff
changeset | 620 | lemma bintrunc_0_BIT: "bintrunc 0 (w BIT b) = Int.Pls" | 
| 24333 | 621 | by auto | 
| 622 | ||
| 623 | lemma bintrunc_Suc_lem: | |
| 624 | "bintrunc (Suc n) x = y ==> m = Suc n ==> bintrunc m x = y" | |
| 625 | by auto | |
| 626 | ||
| 627 | lemmas bintrunc_Suc_Ialts = | |
| 26294 | 628 | bintrunc_Min_I [THEN bintrunc_Suc_lem, standard] | 
| 629 | bintrunc_BIT_I [THEN bintrunc_Suc_lem, standard] | |
| 24333 | 630 | |
| 631 | lemmas sbintrunc_BIT_I = trans [OF sbintrunc_Suc_BIT thobini1] | |
| 632 | ||
| 633 | lemmas sbintrunc_Suc_Is = | |
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 634 | sbintrunc_Sucs(1-3) [THEN thobini1 [THEN [2] trans], standard] | 
| 24333 | 635 | |
| 636 | lemmas sbintrunc_Suc_minus_Is = | |
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 637 | sbintrunc_minus_simps(1-3) [THEN thobini1 [THEN [2] trans], standard] | 
| 24333 | 638 | |
| 639 | lemma sbintrunc_Suc_lem: | |
| 640 | "sbintrunc (Suc n) x = y ==> m = Suc n ==> sbintrunc m x = y" | |
| 641 | by auto | |
| 642 | ||
| 643 | lemmas sbintrunc_Suc_Ialts = | |
| 644 | sbintrunc_Suc_Is [THEN sbintrunc_Suc_lem, standard] | |
| 645 | ||
| 646 | lemma sbintrunc_bintrunc_lt: | |
| 647 | "m > n ==> sbintrunc n (bintrunc m w) = sbintrunc n w" | |
| 648 | by (rule bin_eqI) (auto simp: nth_sbintr nth_bintr) | |
| 649 | ||
| 650 | lemma bintrunc_sbintrunc_le: | |
| 651 | "m <= Suc n ==> bintrunc m (sbintrunc n w) = bintrunc m w" | |
| 652 | apply (rule bin_eqI) | |
| 653 | apply (auto simp: nth_sbintr nth_bintr) | |
| 654 | apply (subgoal_tac "x=n", safe, arith+)[1] | |
| 655 | apply (subgoal_tac "x=n", safe, arith+)[1] | |
| 656 | done | |
| 657 | ||
| 658 | lemmas bintrunc_sbintrunc [simp] = order_refl [THEN bintrunc_sbintrunc_le] | |
| 659 | lemmas sbintrunc_bintrunc [simp] = lessI [THEN sbintrunc_bintrunc_lt] | |
| 660 | lemmas bintrunc_bintrunc [simp] = order_refl [THEN bintrunc_bintrunc_l] | |
| 661 | lemmas sbintrunc_sbintrunc [simp] = order_refl [THEN sbintrunc_sbintrunc_l] | |
| 662 | ||
| 663 | lemma bintrunc_sbintrunc' [simp]: | |
| 664 | "0 < n \<Longrightarrow> bintrunc n (sbintrunc (n - 1) w) = bintrunc n w" | |
| 665 | by (cases n) (auto simp del: bintrunc.Suc) | |
| 666 | ||
| 667 | lemma sbintrunc_bintrunc' [simp]: | |
| 668 | "0 < n \<Longrightarrow> sbintrunc (n - 1) (bintrunc n w) = sbintrunc (n - 1) w" | |
| 669 | by (cases n) (auto simp del: bintrunc.Suc) | |
| 670 | ||
| 671 | lemma bin_sbin_eq_iff: | |
| 672 | "bintrunc (Suc n) x = bintrunc (Suc n) y <-> | |
| 673 | sbintrunc n x = sbintrunc n y" | |
| 674 | apply (rule iffI) | |
| 675 | apply (rule box_equals [OF _ sbintrunc_bintrunc sbintrunc_bintrunc]) | |
| 676 | apply simp | |
| 677 | apply (rule box_equals [OF _ bintrunc_sbintrunc bintrunc_sbintrunc]) | |
| 678 | apply simp | |
| 679 | done | |
| 680 | ||
| 681 | lemma bin_sbin_eq_iff': | |
| 682 | "0 < n \<Longrightarrow> bintrunc n x = bintrunc n y <-> | |
| 683 | sbintrunc (n - 1) x = sbintrunc (n - 1) y" | |
| 684 | by (cases n) (simp_all add: bin_sbin_eq_iff del: bintrunc.Suc) | |
| 685 | ||
| 686 | lemmas bintrunc_sbintruncS0 [simp] = bintrunc_sbintrunc' [unfolded One_nat_def] | |
| 687 | lemmas sbintrunc_bintruncS0 [simp] = sbintrunc_bintrunc' [unfolded One_nat_def] | |
| 688 | ||
| 689 | lemmas bintrunc_bintrunc_l' = le_add1 [THEN bintrunc_bintrunc_l] | |
| 690 | lemmas sbintrunc_sbintrunc_l' = le_add1 [THEN sbintrunc_sbintrunc_l] | |
| 691 | ||
| 692 | (* although bintrunc_minus_simps, if added to default simpset, | |
| 693 | tends to get applied where it's not wanted in developing the theories, | |
| 694 | we get a version for when the word length is given literally *) | |
| 695 | ||
| 696 | lemmas nat_non0_gr = | |
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
24465diff
changeset | 697 | trans [OF iszero_def [THEN Not_eq_iff [THEN iffD2]] refl, standard] | 
| 24333 | 698 | |
| 699 | lemmas bintrunc_pred_simps [simp] = | |
| 700 | bintrunc_minus_simps [of "number_of bin", simplified nobm1, standard] | |
| 701 | ||
| 702 | lemmas sbintrunc_pred_simps [simp] = | |
| 703 | sbintrunc_minus_simps [of "number_of bin", simplified nobm1, standard] | |
| 704 | ||
| 705 | lemma no_bintr_alt: | |
| 706 | "number_of (bintrunc n w) = w mod 2 ^ n" | |
| 707 | by (simp add: number_of_eq bintrunc_mod2p) | |
| 708 | ||
| 709 | lemma no_bintr_alt1: "bintrunc n = (%w. w mod 2 ^ n :: int)" | |
| 710 | by (rule ext) (rule bintrunc_mod2p) | |
| 711 | ||
| 712 | lemma range_bintrunc: "range (bintrunc n) = {i. 0 <= i & i < 2 ^ n}"
 | |
| 713 | apply (unfold no_bintr_alt1) | |
| 714 | apply (auto simp add: image_iff) | |
| 715 | apply (rule exI) | |
| 716 | apply (auto intro: int_mod_lem [THEN iffD1, symmetric]) | |
| 717 | done | |
| 718 | ||
| 719 | lemma no_bintr: | |
| 720 | "number_of (bintrunc n w) = (number_of w mod 2 ^ n :: int)" | |
| 721 | by (simp add : bintrunc_mod2p number_of_eq) | |
| 722 | ||
| 723 | lemma no_sbintr_alt2: | |
| 724 | "sbintrunc n = (%w. (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n :: int)" | |
| 725 | by (rule ext) (simp add : sbintrunc_mod2p) | |
| 726 | ||
| 727 | lemma no_sbintr: | |
| 728 | "number_of (sbintrunc n w) = | |
| 729 | ((number_of w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n :: int)" | |
| 730 | by (simp add : no_sbintr_alt2 number_of_eq) | |
| 731 | ||
| 732 | lemma range_sbintrunc: | |
| 733 |   "range (sbintrunc n) = {i. - (2 ^ n) <= i & i < 2 ^ n}"
 | |
| 734 | apply (unfold no_sbintr_alt2) | |
| 735 | apply (auto simp add: image_iff eq_diff_eq) | |
| 736 | apply (rule exI) | |
| 737 | apply (auto intro: int_mod_lem [THEN iffD1, symmetric]) | |
| 738 | done | |
| 739 | ||
| 25349 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 740 | lemma sb_inc_lem: | 
| 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 741 | "(a::int) + 2^k < 0 \<Longrightarrow> a + 2^k + 2^(Suc k) <= (a + 2^k) mod 2^(Suc k)" | 
| 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 742 | apply (erule int_mod_ge' [where n = "2 ^ (Suc k)" and b = "a + 2 ^ k", simplified zless2p]) | 
| 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 743 | apply (rule TrueI) | 
| 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 744 | done | 
| 24333 | 745 | |
| 25349 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 746 | lemma sb_inc_lem': | 
| 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 747 | "(a::int) < - (2^k) \<Longrightarrow> a + 2^k + 2^(Suc k) <= (a + 2^k) mod 2^(Suc k)" | 
| 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 748 | by (rule iffD1 [OF less_diff_eq, THEN sb_inc_lem, simplified OrderedGroup.diff_0]) | 
| 24333 | 749 | |
| 750 | lemma sbintrunc_inc: | |
| 25349 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 751 | "x < - (2^n) ==> x + 2^(Suc n) <= sbintrunc n x" | 
| 24333 | 752 | unfolding no_sbintr_alt2 by (drule sb_inc_lem') simp | 
| 753 | ||
| 25349 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 754 | lemma sb_dec_lem: | 
| 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 755 | "(0::int) <= - (2^k) + a ==> (a + 2^k) mod (2 * 2 ^ k) <= - (2 ^ k) + a" | 
| 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 756 | by (rule int_mod_le' [where n = "2 ^ (Suc k)" and b = "a + 2 ^ k", | 
| 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 757 | simplified zless2p, OF _ TrueI, simplified]) | 
| 24333 | 758 | |
| 25349 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 759 | lemma sb_dec_lem': | 
| 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 760 | "(2::int) ^ k <= a ==> (a + 2 ^ k) mod (2 * 2 ^ k) <= - (2 ^ k) + a" | 
| 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 761 | by (rule iffD1 [OF diff_le_eq', THEN sb_dec_lem, simplified]) | 
| 24333 | 762 | |
| 763 | lemma sbintrunc_dec: | |
| 764 | "x >= (2 ^ n) ==> x - 2 ^ (Suc n) >= sbintrunc n x" | |
| 765 | unfolding no_sbintr_alt2 by (drule sb_dec_lem') simp | |
| 766 | ||
| 25349 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 767 | lemmas zmod_uminus' = zmod_uminus [where b="c", standard] | 
| 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 768 | lemmas zpower_zmod' = zpower_zmod [where m="c" and y="k", standard] | 
| 24333 | 769 | |
| 770 | lemmas brdmod1s' [symmetric] = | |
| 771 | zmod_zadd_left_eq zmod_zadd_right_eq | |
| 772 | zmod_zsub_left_eq zmod_zsub_right_eq | |
| 773 | zmod_zmult1_eq zmod_zmult1_eq_rev | |
| 774 | ||
| 775 | lemmas brdmods' [symmetric] = | |
| 776 | zpower_zmod' [symmetric] | |
| 777 | trans [OF zmod_zadd_left_eq zmod_zadd_right_eq] | |
| 778 | trans [OF zmod_zsub_left_eq zmod_zsub_right_eq] | |
| 779 | trans [OF zmod_zmult1_eq zmod_zmult1_eq_rev] | |
| 780 | zmod_uminus' [symmetric] | |
| 781 | zmod_zadd_left_eq [where b = "1"] | |
| 782 | zmod_zsub_left_eq [where b = "1"] | |
| 783 | ||
| 784 | lemmas bintr_arith1s = | |
| 25349 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 785 | brdmod1s' [where c="2^n", folded pred_def succ_def bintrunc_mod2p, standard] | 
| 24333 | 786 | lemmas bintr_ariths = | 
| 25349 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 787 | brdmods' [where c="2^n", folded pred_def succ_def bintrunc_mod2p, standard] | 
| 24333 | 788 | |
| 25349 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 789 | lemmas m2pths = pos_mod_sign pos_mod_bound [OF zless2p, standard] | 
| 24364 | 790 | |
| 24333 | 791 | lemma bintr_ge0: "(0 :: int) <= number_of (bintrunc n w)" | 
| 792 | by (simp add : no_bintr m2pths) | |
| 793 | ||
| 794 | lemma bintr_lt2p: "number_of (bintrunc n w) < (2 ^ n :: int)" | |
| 795 | by (simp add : no_bintr m2pths) | |
| 796 | ||
| 797 | lemma bintr_Min: | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25349diff
changeset | 798 | "number_of (bintrunc n Int.Min) = (2 ^ n :: int) - 1" | 
| 24333 | 799 | by (simp add : no_bintr m1mod2k) | 
| 800 | ||
| 801 | lemma sbintr_ge: "(- (2 ^ n) :: int) <= number_of (sbintrunc n w)" | |
| 802 | by (simp add : no_sbintr m2pths) | |
| 803 | ||
| 804 | lemma sbintr_lt: "number_of (sbintrunc n w) < (2 ^ n :: int)" | |
| 805 | by (simp add : no_sbintr m2pths) | |
| 806 | ||
| 807 | lemma bintrunc_Suc: | |
| 808 | "bintrunc (Suc n) bin = bintrunc n (bin_rest bin) BIT bin_last bin" | |
| 809 | by (case_tac bin rule: bin_exhaust) auto | |
| 810 | ||
| 811 | lemma sign_Pls_ge_0: | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25349diff
changeset | 812 | "(bin_sign bin = Int.Pls) = (number_of bin >= (0 :: int))" | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 813 | by (induct bin rule: numeral_induct) auto | 
| 24333 | 814 | |
| 815 | lemma sign_Min_lt_0: | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25349diff
changeset | 816 | "(bin_sign bin = Int.Min) = (number_of bin < (0 :: int))" | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 817 | by (induct bin rule: numeral_induct) auto | 
| 24333 | 818 | |
| 819 | lemmas sign_Min_neg = trans [OF sign_Min_lt_0 neg_def [symmetric]] | |
| 820 | ||
| 821 | lemma bin_rest_trunc: | |
| 822 | "!!bin. (bin_rest (bintrunc n bin)) = bintrunc (n - 1) (bin_rest bin)" | |
| 823 | by (induct n) auto | |
| 824 | ||
| 825 | lemma bin_rest_power_trunc [rule_format] : | |
| 826 | "(bin_rest ^ k) (bintrunc n bin) = | |
| 827 | bintrunc (n - k) ((bin_rest ^ k) bin)" | |
| 828 | by (induct k) (auto simp: bin_rest_trunc) | |
| 829 | ||
| 830 | lemma bin_rest_trunc_i: | |
| 831 | "bintrunc n (bin_rest bin) = bin_rest (bintrunc (Suc n) bin)" | |
| 832 | by auto | |
| 833 | ||
| 834 | lemma bin_rest_strunc: | |
| 835 | "!!bin. bin_rest (sbintrunc (Suc n) bin) = sbintrunc n (bin_rest bin)" | |
| 836 | by (induct n) auto | |
| 837 | ||
| 838 | lemma bintrunc_rest [simp]: | |
| 839 | "!!bin. bintrunc n (bin_rest (bintrunc n bin)) = bin_rest (bintrunc n bin)" | |
| 840 | apply (induct n, simp) | |
| 841 | apply (case_tac bin rule: bin_exhaust) | |
| 842 | apply (auto simp: bintrunc_bintrunc_l) | |
| 843 | done | |
| 844 | ||
| 845 | lemma sbintrunc_rest [simp]: | |
| 846 | "!!bin. sbintrunc n (bin_rest (sbintrunc n bin)) = bin_rest (sbintrunc n bin)" | |
| 847 | apply (induct n, simp) | |
| 848 | apply (case_tac bin rule: bin_exhaust) | |
| 849 | apply (auto simp: bintrunc_bintrunc_l split: bit.splits) | |
| 850 | done | |
| 851 | ||
| 852 | lemma bintrunc_rest': | |
| 853 | "bintrunc n o bin_rest o bintrunc n = bin_rest o bintrunc n" | |
| 854 | by (rule ext) auto | |
| 855 | ||
| 856 | lemma sbintrunc_rest' : | |
| 857 | "sbintrunc n o bin_rest o sbintrunc n = bin_rest o sbintrunc n" | |
| 858 | by (rule ext) auto | |
| 859 | ||
| 860 | lemma rco_lem: | |
| 861 | "f o g o f = g o f ==> f o (g o f) ^ n = g ^ n o f" | |
| 862 | apply (rule ext) | |
| 863 | apply (induct_tac n) | |
| 864 | apply (simp_all (no_asm)) | |
| 865 | apply (drule fun_cong) | |
| 866 | apply (unfold o_def) | |
| 867 | apply (erule trans) | |
| 868 | apply simp | |
| 869 | done | |
| 870 | ||
| 871 | lemma rco_alt: "(f o g) ^ n o f = f o (g o f) ^ n" | |
| 872 | apply (rule ext) | |
| 873 | apply (induct n) | |
| 874 | apply (simp_all add: o_def) | |
| 875 | done | |
| 876 | ||
| 877 | lemmas rco_bintr = bintrunc_rest' | |
| 878 | [THEN rco_lem [THEN fun_cong], unfolded o_def] | |
| 879 | lemmas rco_sbintr = sbintrunc_rest' | |
| 880 | [THEN rco_lem [THEN fun_cong], unfolded o_def] | |
| 881 | ||
| 24364 | 882 | subsection {* Splitting and concatenation *}
 | 
| 883 | ||
| 26557 | 884 | primrec bin_split :: "nat \<Rightarrow> int \<Rightarrow> int \<times> int" where | 
| 885 | Z: "bin_split 0 w = (w, Int.Pls)" | |
| 886 | | Suc: "bin_split (Suc n) w = (let (w1, w2) = bin_split n (bin_rest w) | |
| 887 | in (w1, w2 BIT bin_last w))" | |
| 24364 | 888 | |
| 26557 | 889 | primrec bin_cat :: "int \<Rightarrow> nat \<Rightarrow> int \<Rightarrow> int" where | 
| 890 | Z: "bin_cat w 0 v = w" | |
| 891 | | Suc: "bin_cat w (Suc n) v = bin_cat w n (bin_rest v) BIT bin_last v" | |
| 24364 | 892 | |
| 893 | subsection {* Miscellaneous lemmas *}
 | |
| 894 | ||
| 895 | lemmas funpow_minus_simp = | |
| 896 | trans [OF gen_minus [where f = "power f"] funpow_Suc, standard] | |
| 897 | ||
| 898 | lemmas funpow_pred_simp [simp] = | |
| 899 | funpow_minus_simp [of "number_of bin", simplified nobm1, standard] | |
| 900 | ||
| 901 | lemmas replicate_minus_simp = | |
| 902 | trans [OF gen_minus [where f = "%n. replicate n x"] replicate.replicate_Suc, | |
| 903 | standard] | |
| 904 | ||
| 905 | lemmas replicate_pred_simp [simp] = | |
| 906 | replicate_minus_simp [of "number_of bin", simplified nobm1, standard] | |
| 907 | ||
| 25349 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 908 | lemmas power_Suc_no [simp] = power_Suc [of "number_of a", standard] | 
| 24364 | 909 | |
| 910 | lemmas power_minus_simp = | |
| 911 | trans [OF gen_minus [where f = "power f"] power_Suc, standard] | |
| 912 | ||
| 913 | lemmas power_pred_simp = | |
| 914 | power_minus_simp [of "number_of bin", simplified nobm1, standard] | |
| 25349 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 915 | lemmas power_pred_simp_no [simp] = power_pred_simp [where f= "number_of f", standard] | 
| 24364 | 916 | |
| 917 | lemma list_exhaust_size_gt0: | |
| 918 | assumes y: "\<And>a list. y = a # list \<Longrightarrow> P" | |
| 919 | shows "0 < length y \<Longrightarrow> P" | |
| 920 | apply (cases y, simp) | |
| 921 | apply (rule y) | |
| 922 | apply fastsimp | |
| 923 | done | |
| 924 | ||
| 925 | lemma list_exhaust_size_eq0: | |
| 926 | assumes y: "y = [] \<Longrightarrow> P" | |
| 927 | shows "length y = 0 \<Longrightarrow> P" | |
| 928 | apply (cases y) | |
| 929 | apply (rule y, simp) | |
| 930 | apply simp | |
| 931 | done | |
| 932 | ||
| 933 | lemma size_Cons_lem_eq: | |
| 934 | "y = xa # list ==> size y = Suc k ==> size list = k" | |
| 935 | by auto | |
| 936 | ||
| 937 | lemma size_Cons_lem_eq_bin: | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25349diff
changeset | 938 | "y = xa # list ==> size y = number_of (Int.succ k) ==> | 
| 24364 | 939 | size list = number_of k" | 
| 940 | by (auto simp: pred_def succ_def split add : split_if_asm) | |
| 941 | ||
| 24333 | 942 | lemmas ls_splits = | 
| 943 | prod.split split_split prod.split_asm split_split_asm split_if_asm | |
| 944 | ||
| 945 | lemma not_B1_is_B0: "y \<noteq> bit.B1 \<Longrightarrow> y = bit.B0" | |
| 946 | by (cases y) auto | |
| 947 | ||
| 948 | lemma B1_ass_B0: | |
| 949 | assumes y: "y = bit.B0 \<Longrightarrow> y = bit.B1" | |
| 950 | shows "y = bit.B1" | |
| 951 | apply (rule classical) | |
| 952 | apply (drule not_B1_is_B0) | |
| 953 | apply (erule y) | |
| 954 | done | |
| 955 | ||
| 956 | -- "simplifications for specific word lengths" | |
| 957 | lemmas n2s_ths [THEN eq_reflection] = add_2_eq_Suc add_2_eq_Suc' | |
| 958 | ||
| 959 | lemmas s2n_ths = n2s_ths [symmetric] | |
| 960 | ||
| 961 | end |