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