| author | wenzelm | 
| Sun, 23 Jun 2013 22:31:50 +0200 | |
| changeset 52432 | c03090937c3b | 
| parent 47219 | 172c031ad743 | 
| child 53062 | 3af1a6020014 | 
| permissions | -rw-r--r-- | 
| 24333 | 1 | (* | 
| 2 | Author: Jeremy Dawson, NICTA | |
| 3 | ||
| 44939 
5930d35c976d
removed unused legacy lemma names, some comment cleanup.
 kleing parents: 
44890diff
changeset | 4 | Basic definitions to do with integers, expressed using Pls, Min, BIT. | 
| 24333 | 5 | *) | 
| 6 | ||
| 24350 | 7 | header {* Basic Definitions for Binary Integers *}
 | 
| 8 | ||
| 37658 | 9 | theory Bit_Representation | 
| 41413 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
37667diff
changeset | 10 | imports Misc_Numeric "~~/src/HOL/Library/Bit" | 
| 24333 | 11 | begin | 
| 12 | ||
| 26557 | 13 | subsection {* Further properties of numerals *}
 | 
| 14 | ||
| 37667 | 15 | definition bitval :: "bit \<Rightarrow> 'a\<Colon>zero_neq_one" where | 
| 16 | "bitval = bit_case 0 1" | |
| 17 | ||
| 18 | lemma bitval_simps [simp]: | |
| 19 | "bitval 0 = 0" | |
| 20 | "bitval 1 = 1" | |
| 21 | by (simp_all add: bitval_def) | |
| 22 | ||
| 37654 | 23 | definition Bit :: "int \<Rightarrow> bit \<Rightarrow> int" (infixl "BIT" 90) where | 
| 37667 | 24 | "k BIT b = bitval b + k + k" | 
| 26557 | 25 | |
| 45843 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 26 | definition bin_last :: "int \<Rightarrow> bit" where | 
| 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 27 | "bin_last w = (if w mod 2 = 0 then (0::bit) else (1::bit))" | 
| 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 28 | |
| 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 29 | definition bin_rest :: "int \<Rightarrow> int" where | 
| 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 30 | "bin_rest w = w div 2" | 
| 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 31 | |
| 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 32 | lemma bin_rl_simp [simp]: | 
| 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 33 | "bin_rest w BIT bin_last w = w" | 
| 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 34 | unfolding bin_rest_def bin_last_def Bit_def | 
| 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 35 | using mod_div_equality [of w 2] | 
| 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 36 | by (cases "w mod 2 = 0", simp_all) | 
| 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 37 | |
| 46600 | 38 | lemma bin_rest_BIT [simp]: "bin_rest (x BIT b) = x" | 
| 45843 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 39 | unfolding bin_rest_def Bit_def | 
| 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 40 | by (cases b, simp_all) | 
| 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 41 | |
| 46600 | 42 | lemma bin_last_BIT [simp]: "bin_last (x BIT b) = b" | 
| 45843 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 43 | unfolding bin_last_def Bit_def | 
| 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 44 | by (cases b, simp_all add: z1pmod2) | 
| 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 45 | |
| 45848 | 46 | lemma BIT_eq_iff [iff]: "u BIT b = v BIT c \<longleftrightarrow> u = v \<and> b = c" | 
| 45843 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 47 | by (metis bin_rest_BIT bin_last_BIT) | 
| 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 48 | |
| 45849 | 49 | lemma BIT_bin_simps [simp]: | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 50 | "numeral k BIT 0 = numeral (Num.Bit0 k)" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 51 | "numeral k BIT 1 = numeral (Num.Bit1 k)" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 52 | "neg_numeral k BIT 0 = neg_numeral (Num.Bit0 k)" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 53 | "neg_numeral k BIT 1 = neg_numeral (Num.BitM k)" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 54 | unfolding neg_numeral_def numeral.simps numeral_BitM | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 55 | unfolding Bit_def bitval_simps | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 56 | by (simp_all del: arith_simps add_numeral_special diff_numeral_special) | 
| 45849 | 57 | |
| 58 | lemma BIT_special_simps [simp]: | |
| 59 | shows "0 BIT 0 = 0" and "0 BIT 1 = 1" and "1 BIT 0 = 2" and "1 BIT 1 = 3" | |
| 60 | unfolding Bit_def by simp_all | |
| 61 | ||
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 62 | lemma BitM_inc: "Num.BitM (Num.inc w) = Num.Bit1 w" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 63 | by (induct w, simp_all) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 64 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 65 | lemma expand_BIT: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 66 | "numeral (Num.Bit0 w) = numeral w BIT 0" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 67 | "numeral (Num.Bit1 w) = numeral w BIT 1" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 68 | "neg_numeral (Num.Bit0 w) = neg_numeral w BIT 0" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 69 | "neg_numeral (Num.Bit1 w) = neg_numeral (w + Num.One) BIT 1" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 70 | unfolding add_One by (simp_all add: BitM_inc) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 71 | |
| 45851 
19f7ac6cf3cc
add simp rules for bin_rest and bin_last applied to numerals
 huffman parents: 
45850diff
changeset | 72 | lemma bin_last_numeral_simps [simp]: | 
| 
19f7ac6cf3cc
add simp rules for bin_rest and bin_last applied to numerals
 huffman parents: 
45850diff
changeset | 73 | "bin_last 0 = 0" | 
| 
19f7ac6cf3cc
add simp rules for bin_rest and bin_last applied to numerals
 huffman parents: 
45850diff
changeset | 74 | "bin_last 1 = 1" | 
| 
19f7ac6cf3cc
add simp rules for bin_rest and bin_last applied to numerals
 huffman parents: 
45850diff
changeset | 75 | "bin_last -1 = 1" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 76 | "bin_last Numeral1 = 1" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 77 | "bin_last (numeral (Num.Bit0 w)) = 0" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 78 | "bin_last (numeral (Num.Bit1 w)) = 1" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 79 | "bin_last (neg_numeral (Num.Bit0 w)) = 0" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 80 | "bin_last (neg_numeral (Num.Bit1 w)) = 1" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 81 | unfolding expand_BIT bin_last_BIT by (simp_all add: bin_last_def) | 
| 45851 
19f7ac6cf3cc
add simp rules for bin_rest and bin_last applied to numerals
 huffman parents: 
45850diff
changeset | 82 | |
| 
19f7ac6cf3cc
add simp rules for bin_rest and bin_last applied to numerals
 huffman parents: 
45850diff
changeset | 83 | lemma bin_rest_numeral_simps [simp]: | 
| 
19f7ac6cf3cc
add simp rules for bin_rest and bin_last applied to numerals
 huffman parents: 
45850diff
changeset | 84 | "bin_rest 0 = 0" | 
| 
19f7ac6cf3cc
add simp rules for bin_rest and bin_last applied to numerals
 huffman parents: 
45850diff
changeset | 85 | "bin_rest 1 = 0" | 
| 
19f7ac6cf3cc
add simp rules for bin_rest and bin_last applied to numerals
 huffman parents: 
45850diff
changeset | 86 | "bin_rest -1 = -1" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 87 | "bin_rest Numeral1 = 0" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 88 | "bin_rest (numeral (Num.Bit0 w)) = numeral w" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 89 | "bin_rest (numeral (Num.Bit1 w)) = numeral w" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 90 | "bin_rest (neg_numeral (Num.Bit0 w)) = neg_numeral w" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 91 | "bin_rest (neg_numeral (Num.Bit1 w)) = neg_numeral (w + Num.One)" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 92 | unfolding expand_BIT bin_rest_BIT by (simp_all add: bin_rest_def) | 
| 26557 | 93 | |
| 94 | lemma less_Bits: | |
| 37654 | 95 | "(v BIT b < w BIT c) = (v < w | v <= w & b = (0::bit) & c = (1::bit))" | 
| 37667 | 96 | unfolding Bit_def by (auto simp add: bitval_def split: bit.split) | 
| 24333 | 97 | |
| 26557 | 98 | lemma le_Bits: | 
| 37654 | 99 | "(v BIT b <= w BIT c) = (v < w | v <= w & (b ~= (1::bit) | c ~= (0::bit)))" | 
| 37667 | 100 | unfolding Bit_def by (auto simp add: bitval_def split: bit.split) | 
| 26557 | 101 | |
| 102 | lemma Bit_B0: | |
| 37654 | 103 | "k BIT (0::bit) = k + k" | 
| 26557 | 104 | by (unfold Bit_def) simp | 
| 105 | ||
| 106 | lemma Bit_B1: | |
| 37654 | 107 | "k BIT (1::bit) = k + k + 1" | 
| 26557 | 108 | by (unfold Bit_def) simp | 
| 109 | ||
| 37654 | 110 | lemma Bit_B0_2t: "k BIT (0::bit) = 2 * k" | 
| 26557 | 111 | by (rule trans, rule Bit_B0) simp | 
| 112 | ||
| 37654 | 113 | lemma Bit_B1_2t: "k BIT (1::bit) = 2 * k + 1" | 
| 26557 | 114 | by (rule trans, rule Bit_B1) simp | 
| 115 | ||
| 116 | lemma B_mod_2': | |
| 37654 | 117 | "X = 2 ==> (w BIT (1::bit)) mod X = 1 & (w BIT (0::bit)) mod X = 0" | 
| 26557 | 118 | apply (simp (no_asm) only: Bit_B0 Bit_B1) | 
| 119 | apply (simp add: z1pmod2) | |
| 24465 | 120 | done | 
| 24333 | 121 | |
| 26557 | 122 | lemma neB1E [elim!]: | 
| 37654 | 123 | assumes ne: "y \<noteq> (1::bit)" | 
| 124 | assumes y: "y = (0::bit) \<Longrightarrow> P" | |
| 26557 | 125 | shows "P" | 
| 126 | apply (rule y) | |
| 127 | apply (cases y rule: bit.exhaust, simp) | |
| 128 | apply (simp add: ne) | |
| 129 | done | |
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 130 | |
| 26557 | 131 | lemma bin_ex_rl: "EX w b. w BIT b = bin" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 132 | by (metis bin_rl_simp) | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 133 | |
| 26557 | 134 | lemma bin_exhaust: | 
| 135 | assumes Q: "\<And>x b. bin = x BIT b \<Longrightarrow> Q" | |
| 136 | shows "Q" | |
| 137 | apply (insert bin_ex_rl [of bin]) | |
| 138 | apply (erule exE)+ | |
| 139 | apply (rule Q) | |
| 140 | apply force | |
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 141 | done | 
| 24333 | 142 | |
| 143 | ||
| 24465 | 144 | subsection {* Destructors for binary integers *}
 | 
| 24364 | 145 | |
| 26514 | 146 | primrec bin_nth where | 
| 37654 | 147 | Z: "bin_nth w 0 = (bin_last w = (1::bit))" | 
| 26557 | 148 | | Suc: "bin_nth w (Suc n) = bin_nth (bin_rest w) n" | 
| 24364 | 149 | |
| 26557 | 150 | lemma bin_abs_lem: | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 151 | "bin = (w BIT b) ==> bin ~= -1 --> bin ~= 0 --> | 
| 26557 | 152 | nat (abs w) < nat (abs bin)" | 
| 46598 | 153 | apply clarsimp | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 154 | apply (unfold Bit_def) | 
| 26557 | 155 | apply (cases b) | 
| 156 | apply (clarsimp, arith) | |
| 157 | apply (clarsimp, arith) | |
| 158 | done | |
| 159 | ||
| 160 | lemma bin_induct: | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 161 | assumes PPls: "P 0" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 162 | and PMin: "P -1" | 
| 26557 | 163 | and PBit: "!!bin bit. P bin ==> P (bin BIT bit)" | 
| 164 | shows "P bin" | |
| 165 | apply (rule_tac P=P and a=bin and f1="nat o abs" | |
| 166 | in wf_measure [THEN wf_induct]) | |
| 167 | apply (simp add: measure_def inv_image_def) | |
| 168 | apply (case_tac x rule: bin_exhaust) | |
| 169 | apply (frule bin_abs_lem) | |
| 170 | apply (auto simp add : PPls PMin PBit) | |
| 171 | done | |
| 172 | ||
| 24333 | 173 | lemma Bit_div2 [simp]: "(w BIT b) div 2 = w" | 
| 46600 | 174 | unfolding bin_rest_def [symmetric] by (rule bin_rest_BIT) | 
| 24333 | 175 | |
| 176 | lemma bin_nth_lem [rule_format]: | |
| 177 | "ALL y. bin_nth x = bin_nth y --> x = y" | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 178 | apply (induct x rule: bin_induct) | 
| 24333 | 179 | apply safe | 
| 180 | apply (erule rev_mp) | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 181 | apply (induct_tac y rule: bin_induct) | 
| 46008 
c296c75f4cf4
reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
 wenzelm parents: 
46001diff
changeset | 182 | apply safe | 
| 24333 | 183 | apply (drule_tac x=0 in fun_cong, force) | 
| 184 | apply (erule notE, rule ext, | |
| 185 | drule_tac x="Suc x" in fun_cong, force) | |
| 46600 | 186 | apply (drule_tac x=0 in fun_cong, force) | 
| 24333 | 187 | apply (erule rev_mp) | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 188 | apply (induct_tac y rule: bin_induct) | 
| 46008 
c296c75f4cf4
reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
 wenzelm parents: 
46001diff
changeset | 189 | apply safe | 
| 24333 | 190 | apply (drule_tac x=0 in fun_cong, force) | 
| 191 | apply (erule notE, rule ext, | |
| 192 | drule_tac x="Suc x" in fun_cong, force) | |
| 46600 | 193 | apply (drule_tac x=0 in fun_cong, force) | 
| 24333 | 194 | apply (case_tac y rule: bin_exhaust) | 
| 195 | apply clarify | |
| 196 | apply (erule allE) | |
| 197 | apply (erule impE) | |
| 198 | prefer 2 | |
| 45848 | 199 | apply (erule conjI) | 
| 24333 | 200 | apply (drule_tac x=0 in fun_cong, force) | 
| 201 | apply (rule ext) | |
| 202 | apply (drule_tac x="Suc ?x" in fun_cong, force) | |
| 203 | done | |
| 204 | ||
| 205 | lemma bin_nth_eq_iff: "(bin_nth x = bin_nth y) = (x = y)" | |
| 206 | by (auto elim: bin_nth_lem) | |
| 207 | ||
| 45604 | 208 | lemmas bin_eqI = ext [THEN bin_nth_eq_iff [THEN iffD1]] | 
| 24333 | 209 | |
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 210 | lemma bin_eq_iff: "x = y \<longleftrightarrow> (\<forall>n. bin_nth x n = bin_nth y n)" | 
| 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 211 | by (auto intro!: bin_nth_lem del: equalityI) | 
| 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 212 | |
| 45853 | 213 | lemma bin_nth_zero [simp]: "\<not> bin_nth 0 n" | 
| 214 | by (induct n) auto | |
| 215 | ||
| 46023 
fad87bb608fc
restate some lemmas to respect int/bin distinction
 huffman parents: 
46001diff
changeset | 216 | lemma bin_nth_1 [simp]: "bin_nth 1 n \<longleftrightarrow> n = 0" | 
| 
fad87bb608fc
restate some lemmas to respect int/bin distinction
 huffman parents: 
46001diff
changeset | 217 | by (cases n) simp_all | 
| 
fad87bb608fc
restate some lemmas to respect int/bin distinction
 huffman parents: 
46001diff
changeset | 218 | |
| 45952 | 219 | lemma bin_nth_minus1 [simp]: "bin_nth -1 n" | 
| 220 | by (induct n) auto | |
| 221 | ||
| 37654 | 222 | lemma bin_nth_0_BIT: "bin_nth (w BIT b) 0 = (b = (1::bit))" | 
| 24333 | 223 | by auto | 
| 224 | ||
| 225 | lemma bin_nth_Suc_BIT: "bin_nth (w BIT b) (Suc n) = bin_nth w n" | |
| 226 | by auto | |
| 227 | ||
| 228 | lemma bin_nth_minus [simp]: "0 < n ==> bin_nth (w BIT b) n = bin_nth w (n - 1)" | |
| 229 | by (cases n) auto | |
| 230 | ||
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 231 | lemma bin_nth_numeral: | 
| 47219 
172c031ad743
restate various simp rules for word operations using pred_numeral
 huffman parents: 
47170diff
changeset | 232 | "bin_rest x = y \<Longrightarrow> bin_nth x (numeral n) = bin_nth y (pred_numeral n)" | 
| 
172c031ad743
restate various simp rules for word operations using pred_numeral
 huffman parents: 
47170diff
changeset | 233 | by (simp add: numeral_eq_Suc) | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 234 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 235 | lemmas bin_nth_numeral_simps [simp] = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 236 | bin_nth_numeral [OF bin_rest_numeral_simps(2)] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 237 | bin_nth_numeral [OF bin_rest_numeral_simps(5)] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 238 | bin_nth_numeral [OF bin_rest_numeral_simps(6)] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 239 | bin_nth_numeral [OF bin_rest_numeral_simps(7)] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 240 | bin_nth_numeral [OF bin_rest_numeral_simps(8)] | 
| 24333 | 241 | |
| 242 | lemmas bin_nth_simps = | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 243 | bin_nth.Z bin_nth.Suc bin_nth_zero bin_nth_minus1 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 244 | bin_nth_numeral_simps | 
| 24333 | 245 | |
| 26557 | 246 | |
| 247 | subsection {* Truncating binary integers *}
 | |
| 248 | ||
| 45846 | 249 | definition bin_sign :: "int \<Rightarrow> int" where | 
| 37667 | 250 | bin_sign_def: "bin_sign k = (if k \<ge> 0 then 0 else - 1)" | 
| 26557 | 251 | |
| 252 | lemma bin_sign_simps [simp]: | |
| 45850 | 253 | "bin_sign 0 = 0" | 
| 46604 
9f9e85264e4d
make uses of bin_sign respect int/bin distinction
 huffman parents: 
46601diff
changeset | 254 | "bin_sign 1 = 0" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 255 | "bin_sign (numeral k) = 0" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 256 | "bin_sign (neg_numeral k) = -1" | 
| 26557 | 257 | "bin_sign (w BIT b) = bin_sign w" | 
| 46604 
9f9e85264e4d
make uses of bin_sign respect int/bin distinction
 huffman parents: 
46601diff
changeset | 258 | unfolding bin_sign_def Bit_def bitval_def | 
| 45850 | 259 | by (simp_all split: bit.split) | 
| 26557 | 260 | |
| 24364 | 261 | lemma bin_sign_rest [simp]: | 
| 37667 | 262 | "bin_sign (bin_rest w) = bin_sign w" | 
| 26557 | 263 | by (cases w rule: bin_exhaust) auto | 
| 24364 | 264 | |
| 37667 | 265 | primrec bintrunc :: "nat \<Rightarrow> int \<Rightarrow> int" where | 
| 46001 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
46000diff
changeset | 266 | Z : "bintrunc 0 bin = 0" | 
| 37667 | 267 | | Suc : "bintrunc (Suc n) bin = bintrunc n (bin_rest bin) BIT (bin_last bin)" | 
| 24364 | 268 | |
| 37667 | 269 | primrec sbintrunc :: "nat => int => int" where | 
| 46001 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
46000diff
changeset | 270 | Z : "sbintrunc 0 bin = (case bin_last bin of (1::bit) \<Rightarrow> -1 | (0::bit) \<Rightarrow> 0)" | 
| 37667 | 271 | | Suc : "sbintrunc (Suc n) bin = sbintrunc n (bin_rest bin) BIT (bin_last bin)" | 
| 272 | ||
| 46001 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
46000diff
changeset | 273 | lemma sign_bintr: "bin_sign (bintrunc n w) = 0" | 
| 45954 
f67d3bb5f09c
use 'induct arbitrary' instead of universal quantifiers
 huffman parents: 
45953diff
changeset | 274 | by (induct n arbitrary: w) auto | 
| 24333 | 275 | |
| 45954 
f67d3bb5f09c
use 'induct arbitrary' instead of universal quantifiers
 huffman parents: 
45953diff
changeset | 276 | lemma bintrunc_mod2p: "bintrunc n w = (w mod 2 ^ n)" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 277 | apply (induct n arbitrary: w, clarsimp) | 
| 45954 
f67d3bb5f09c
use 'induct arbitrary' instead of universal quantifiers
 huffman parents: 
45953diff
changeset | 278 | apply (simp add: bin_last_def bin_rest_def Bit_def zmod_zmult2_eq) | 
| 24333 | 279 | done | 
| 280 | ||
| 45954 
f67d3bb5f09c
use 'induct arbitrary' instead of universal quantifiers
 huffman parents: 
45953diff
changeset | 281 | lemma sbintrunc_mod2p: "sbintrunc n w = (w + 2 ^ n) mod 2 ^ (Suc n) - 2 ^ n" | 
| 
f67d3bb5f09c
use 'induct arbitrary' instead of universal quantifiers
 huffman parents: 
45953diff
changeset | 282 | apply (induct n arbitrary: w) | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 283 | apply simp | 
| 30034 | 284 | apply (subst mod_add_left_eq) | 
| 45529 
0e1037d4e049
remove redundant lemmas bin_last_mod and bin_rest_div, use bin_last_def and bin_rest_def instead
 huffman parents: 
44939diff
changeset | 285 | apply (simp add: bin_last_def) | 
| 46001 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
46000diff
changeset | 286 | apply (simp add: bin_last_def bin_rest_def Bit_def) | 
| 30940 
663af91c0720
zmod_zmult_zmult1 now subsumed by mod_mult_mult1
 haftmann parents: 
30034diff
changeset | 287 | apply (clarsimp simp: mod_mult_mult1 [symmetric] | 
| 24333 | 288 | zmod_zdiv_equality [THEN diff_eq_eq [THEN iffD2 [THEN sym]]]) | 
| 289 | apply (rule trans [symmetric, OF _ emep1]) | |
| 290 | apply auto | |
| 291 | apply (auto simp: even_def) | |
| 292 | done | |
| 293 | ||
| 24465 | 294 | subsection "Simplifications for (s)bintrunc" | 
| 295 | ||
| 45852 | 296 | lemma bintrunc_n_0 [simp]: "bintrunc n 0 = 0" | 
| 46001 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
46000diff
changeset | 297 | by (induct n) auto | 
| 45852 | 298 | |
| 45855 | 299 | lemma sbintrunc_n_0 [simp]: "sbintrunc n 0 = 0" | 
| 46001 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
46000diff
changeset | 300 | by (induct n) auto | 
| 45855 | 301 | |
| 45856 | 302 | lemma sbintrunc_n_minus1 [simp]: "sbintrunc n -1 = -1" | 
| 46001 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
46000diff
changeset | 303 | by (induct n) auto | 
| 45856 | 304 | |
| 45852 | 305 | lemma bintrunc_Suc_numeral: | 
| 306 | "bintrunc (Suc n) 1 = 1" | |
| 307 | "bintrunc (Suc n) -1 = bintrunc n -1 BIT 1" | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 308 | "bintrunc (Suc n) (numeral (Num.Bit0 w)) = bintrunc n (numeral w) BIT 0" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 309 | "bintrunc (Suc n) (numeral (Num.Bit1 w)) = bintrunc n (numeral w) BIT 1" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 310 | "bintrunc (Suc n) (neg_numeral (Num.Bit0 w)) = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 311 | bintrunc n (neg_numeral w) BIT 0" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 312 | "bintrunc (Suc n) (neg_numeral (Num.Bit1 w)) = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 313 | bintrunc n (neg_numeral (w + Num.One)) BIT 1" | 
| 45852 | 314 | by simp_all | 
| 315 | ||
| 45856 | 316 | lemma sbintrunc_0_numeral [simp]: | 
| 317 | "sbintrunc 0 1 = -1" | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 318 | "sbintrunc 0 (numeral (Num.Bit0 w)) = 0" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 319 | "sbintrunc 0 (numeral (Num.Bit1 w)) = -1" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 320 | "sbintrunc 0 (neg_numeral (Num.Bit0 w)) = 0" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 321 | "sbintrunc 0 (neg_numeral (Num.Bit1 w)) = -1" | 
| 46001 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
46000diff
changeset | 322 | by simp_all | 
| 45856 | 323 | |
| 45855 | 324 | lemma sbintrunc_Suc_numeral: | 
| 325 | "sbintrunc (Suc n) 1 = 1" | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 326 | "sbintrunc (Suc n) (numeral (Num.Bit0 w)) = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 327 | sbintrunc n (numeral w) BIT 0" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 328 | "sbintrunc (Suc n) (numeral (Num.Bit1 w)) = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 329 | sbintrunc n (numeral w) BIT 1" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 330 | "sbintrunc (Suc n) (neg_numeral (Num.Bit0 w)) = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 331 | sbintrunc n (neg_numeral w) BIT 0" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 332 | "sbintrunc (Suc n) (neg_numeral (Num.Bit1 w)) = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 333 | sbintrunc n (neg_numeral (w + Num.One)) BIT 1" | 
| 45855 | 334 | by simp_all | 
| 335 | ||
| 24465 | 336 | lemma bit_bool: | 
| 37654 | 337 | "(b = (b' = (1::bit))) = (b' = (if b then (1::bit) else (0::bit)))" | 
| 24465 | 338 | by (cases b') auto | 
| 339 | ||
| 340 | lemmas bit_bool1 [simp] = refl [THEN bit_bool [THEN iffD1], symmetric] | |
| 24333 | 341 | |
| 46001 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
46000diff
changeset | 342 | lemma bin_sign_lem: "(bin_sign (sbintrunc n bin) = -1) = bin_nth bin n" | 
| 45954 
f67d3bb5f09c
use 'induct arbitrary' instead of universal quantifiers
 huffman parents: 
45953diff
changeset | 343 | apply (induct n arbitrary: bin) | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 344 | apply (case_tac bin rule: bin_exhaust, case_tac b, auto) | 
| 24333 | 345 | done | 
| 346 | ||
| 45954 
f67d3bb5f09c
use 'induct arbitrary' instead of universal quantifiers
 huffman parents: 
45953diff
changeset | 347 | lemma nth_bintr: "bin_nth (bintrunc m w) n = (n < m & bin_nth w n)" | 
| 
f67d3bb5f09c
use 'induct arbitrary' instead of universal quantifiers
 huffman parents: 
45953diff
changeset | 348 | apply (induct n arbitrary: w m) | 
| 24333 | 349 | apply (case_tac m, auto)[1] | 
| 350 | apply (case_tac m, auto)[1] | |
| 351 | done | |
| 352 | ||
| 353 | lemma nth_sbintr: | |
| 45954 
f67d3bb5f09c
use 'induct arbitrary' instead of universal quantifiers
 huffman parents: 
45953diff
changeset | 354 | "bin_nth (sbintrunc m w) n = | 
| 24333 | 355 | (if n < m then bin_nth w n else bin_nth w m)" | 
| 45954 
f67d3bb5f09c
use 'induct arbitrary' instead of universal quantifiers
 huffman parents: 
45953diff
changeset | 356 | apply (induct n arbitrary: w m) | 
| 24333 | 357 | apply (case_tac m, simp_all split: bit.splits)[1] | 
| 358 | apply (case_tac m, simp_all split: bit.splits)[1] | |
| 359 | done | |
| 360 | ||
| 361 | lemma bin_nth_Bit: | |
| 37654 | 362 | "bin_nth (w BIT b) n = (n = 0 & b = (1::bit) | (EX m. n = Suc m & bin_nth w m))" | 
| 24333 | 363 | by (cases n) auto | 
| 364 | ||
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 365 | lemma bin_nth_Bit0: | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 366 | "bin_nth (numeral (Num.Bit0 w)) n \<longleftrightarrow> | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 367 | (\<exists>m. n = Suc m \<and> bin_nth (numeral w) m)" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 368 | using bin_nth_Bit [where w="numeral w" and b="(0::bit)"] by simp | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 369 | |
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 370 | lemma bin_nth_Bit1: | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 371 | "bin_nth (numeral (Num.Bit1 w)) n \<longleftrightarrow> | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 372 | n = 0 \<or> (\<exists>m. n = Suc m \<and> bin_nth (numeral w) m)" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 373 | using bin_nth_Bit [where w="numeral w" and b="(1::bit)"] by simp | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 374 | |
| 24333 | 375 | lemma bintrunc_bintrunc_l: | 
| 376 | "n <= m ==> (bintrunc m (bintrunc n w) = bintrunc n w)" | |
| 377 | by (rule bin_eqI) (auto simp add : nth_bintr) | |
| 378 | ||
| 379 | lemma sbintrunc_sbintrunc_l: | |
| 380 | "n <= m ==> (sbintrunc m (sbintrunc n w) = sbintrunc n w)" | |
| 32439 | 381 | by (rule bin_eqI) (auto simp: nth_sbintr) | 
| 24333 | 382 | |
| 383 | lemma bintrunc_bintrunc_ge: | |
| 384 | "n <= m ==> (bintrunc n (bintrunc m w) = bintrunc n w)" | |
| 385 | by (rule bin_eqI) (auto simp: nth_bintr) | |
| 386 | ||
| 387 | lemma bintrunc_bintrunc_min [simp]: | |
| 388 | "bintrunc m (bintrunc n w) = bintrunc (min m n) w" | |
| 389 | apply (rule bin_eqI) | |
| 390 | apply (auto simp: nth_bintr) | |
| 391 | done | |
| 392 | ||
| 393 | lemma sbintrunc_sbintrunc_min [simp]: | |
| 394 | "sbintrunc m (sbintrunc n w) = sbintrunc (min m n) w" | |
| 395 | 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 | 396 | apply (auto simp: nth_sbintr min_max.inf_absorb1 min_max.inf_absorb2) | 
| 24333 | 397 | done | 
| 398 | ||
| 399 | lemmas bintrunc_Pls = | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 400 | bintrunc.Suc [where bin="0", simplified bin_last_numeral_simps bin_rest_numeral_simps] | 
| 24333 | 401 | |
| 402 | lemmas bintrunc_Min [simp] = | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 403 | bintrunc.Suc [where bin="-1", simplified bin_last_numeral_simps bin_rest_numeral_simps] | 
| 24333 | 404 | |
| 405 | lemmas bintrunc_BIT [simp] = | |
| 46600 | 406 | bintrunc.Suc [where bin="w BIT b", simplified bin_last_BIT bin_rest_BIT] for w b | 
| 24333 | 407 | |
| 408 | lemmas bintrunc_Sucs = bintrunc_Pls bintrunc_Min bintrunc_BIT | |
| 45852 | 409 | bintrunc_Suc_numeral | 
| 24333 | 410 | |
| 411 | lemmas sbintrunc_Suc_Pls = | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 412 | sbintrunc.Suc [where bin="0", simplified bin_last_numeral_simps bin_rest_numeral_simps] | 
| 24333 | 413 | |
| 414 | lemmas sbintrunc_Suc_Min = | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 415 | sbintrunc.Suc [where bin="-1", simplified bin_last_numeral_simps bin_rest_numeral_simps] | 
| 24333 | 416 | |
| 417 | lemmas sbintrunc_Suc_BIT [simp] = | |
| 46600 | 418 | sbintrunc.Suc [where bin="w BIT b", simplified bin_last_BIT bin_rest_BIT] for w b | 
| 24333 | 419 | |
| 420 | lemmas sbintrunc_Sucs = sbintrunc_Suc_Pls sbintrunc_Suc_Min sbintrunc_Suc_BIT | |
| 45855 | 421 | sbintrunc_Suc_numeral | 
| 24333 | 422 | |
| 423 | lemmas sbintrunc_Pls = | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 424 | sbintrunc.Z [where bin="0", | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 425 | simplified bin_last_numeral_simps bin_rest_numeral_simps bit.simps] | 
| 24333 | 426 | |
| 427 | lemmas sbintrunc_Min = | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 428 | sbintrunc.Z [where bin="-1", | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 429 | simplified bin_last_numeral_simps bin_rest_numeral_simps bit.simps] | 
| 24333 | 430 | |
| 431 | lemmas sbintrunc_0_BIT_B0 [simp] = | |
| 37654 | 432 | sbintrunc.Z [where bin="w BIT (0::bit)", | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 433 | simplified bin_last_numeral_simps bin_rest_numeral_simps bit.simps] for w | 
| 24333 | 434 | |
| 435 | lemmas sbintrunc_0_BIT_B1 [simp] = | |
| 37654 | 436 | sbintrunc.Z [where bin="w BIT (1::bit)", | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 437 | simplified bin_last_BIT bin_rest_numeral_simps bit.simps] for w | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 438 | |
| 24333 | 439 | lemmas sbintrunc_0_simps = | 
| 440 | sbintrunc_Pls sbintrunc_Min sbintrunc_0_BIT_B0 sbintrunc_0_BIT_B1 | |
| 441 | ||
| 442 | lemmas bintrunc_simps = bintrunc.Z bintrunc_Sucs | |
| 443 | lemmas sbintrunc_simps = sbintrunc_0_simps sbintrunc_Sucs | |
| 444 | ||
| 445 | lemma bintrunc_minus: | |
| 446 | "0 < n ==> bintrunc (Suc (n - 1)) w = bintrunc n w" | |
| 447 | by auto | |
| 448 | ||
| 449 | lemma sbintrunc_minus: | |
| 450 | "0 < n ==> sbintrunc (Suc (n - 1)) w = sbintrunc n w" | |
| 451 | by auto | |
| 452 | ||
| 453 | lemmas bintrunc_minus_simps = | |
| 45604 | 454 | bintrunc_Sucs [THEN [2] bintrunc_minus [symmetric, THEN trans]] | 
| 24333 | 455 | lemmas sbintrunc_minus_simps = | 
| 45604 | 456 | sbintrunc_Sucs [THEN [2] sbintrunc_minus [symmetric, THEN trans]] | 
| 24333 | 457 | |
| 45604 | 458 | lemmas thobini1 = arg_cong [where f = "%w. w BIT b"] for b | 
| 24333 | 459 | |
| 460 | lemmas bintrunc_BIT_I = trans [OF bintrunc_BIT thobini1] | |
| 461 | lemmas bintrunc_Min_I = trans [OF bintrunc_Min thobini1] | |
| 462 | ||
| 45604 | 463 | lemmas bmsts = bintrunc_minus_simps(1-3) [THEN thobini1 [THEN [2] trans]] | 
| 24333 | 464 | lemmas bintrunc_Pls_minus_I = bmsts(1) | 
| 465 | lemmas bintrunc_Min_minus_I = bmsts(2) | |
| 466 | lemmas bintrunc_BIT_minus_I = bmsts(3) | |
| 467 | ||
| 468 | lemma bintrunc_Suc_lem: | |
| 469 | "bintrunc (Suc n) x = y ==> m = Suc n ==> bintrunc m x = y" | |
| 470 | by auto | |
| 471 | ||
| 472 | lemmas bintrunc_Suc_Ialts = | |
| 45604 | 473 | bintrunc_Min_I [THEN bintrunc_Suc_lem] | 
| 474 | bintrunc_BIT_I [THEN bintrunc_Suc_lem] | |
| 24333 | 475 | |
| 476 | lemmas sbintrunc_BIT_I = trans [OF sbintrunc_Suc_BIT thobini1] | |
| 477 | ||
| 478 | lemmas sbintrunc_Suc_Is = | |
| 45604 | 479 | sbintrunc_Sucs(1-3) [THEN thobini1 [THEN [2] trans]] | 
| 24333 | 480 | |
| 481 | lemmas sbintrunc_Suc_minus_Is = | |
| 45604 | 482 | sbintrunc_minus_simps(1-3) [THEN thobini1 [THEN [2] trans]] | 
| 24333 | 483 | |
| 484 | lemma sbintrunc_Suc_lem: | |
| 485 | "sbintrunc (Suc n) x = y ==> m = Suc n ==> sbintrunc m x = y" | |
| 486 | by auto | |
| 487 | ||
| 488 | lemmas sbintrunc_Suc_Ialts = | |
| 45604 | 489 | sbintrunc_Suc_Is [THEN sbintrunc_Suc_lem] | 
| 24333 | 490 | |
| 491 | lemma sbintrunc_bintrunc_lt: | |
| 492 | "m > n ==> sbintrunc n (bintrunc m w) = sbintrunc n w" | |
| 493 | by (rule bin_eqI) (auto simp: nth_sbintr nth_bintr) | |
| 494 | ||
| 495 | lemma bintrunc_sbintrunc_le: | |
| 496 | "m <= Suc n ==> bintrunc m (sbintrunc n w) = bintrunc m w" | |
| 497 | apply (rule bin_eqI) | |
| 498 | apply (auto simp: nth_sbintr nth_bintr) | |
| 499 | apply (subgoal_tac "x=n", safe, arith+)[1] | |
| 500 | apply (subgoal_tac "x=n", safe, arith+)[1] | |
| 501 | done | |
| 502 | ||
| 503 | lemmas bintrunc_sbintrunc [simp] = order_refl [THEN bintrunc_sbintrunc_le] | |
| 504 | lemmas sbintrunc_bintrunc [simp] = lessI [THEN sbintrunc_bintrunc_lt] | |
| 505 | lemmas bintrunc_bintrunc [simp] = order_refl [THEN bintrunc_bintrunc_l] | |
| 506 | lemmas sbintrunc_sbintrunc [simp] = order_refl [THEN sbintrunc_sbintrunc_l] | |
| 507 | ||
| 508 | lemma bintrunc_sbintrunc' [simp]: | |
| 509 | "0 < n \<Longrightarrow> bintrunc n (sbintrunc (n - 1) w) = bintrunc n w" | |
| 510 | by (cases n) (auto simp del: bintrunc.Suc) | |
| 511 | ||
| 512 | lemma sbintrunc_bintrunc' [simp]: | |
| 513 | "0 < n \<Longrightarrow> sbintrunc (n - 1) (bintrunc n w) = sbintrunc (n - 1) w" | |
| 514 | by (cases n) (auto simp del: bintrunc.Suc) | |
| 515 | ||
| 516 | lemma bin_sbin_eq_iff: | |
| 517 | "bintrunc (Suc n) x = bintrunc (Suc n) y <-> | |
| 518 | sbintrunc n x = sbintrunc n y" | |
| 519 | apply (rule iffI) | |
| 520 | apply (rule box_equals [OF _ sbintrunc_bintrunc sbintrunc_bintrunc]) | |
| 521 | apply simp | |
| 522 | apply (rule box_equals [OF _ bintrunc_sbintrunc bintrunc_sbintrunc]) | |
| 523 | apply simp | |
| 524 | done | |
| 525 | ||
| 526 | lemma bin_sbin_eq_iff': | |
| 527 | "0 < n \<Longrightarrow> bintrunc n x = bintrunc n y <-> | |
| 528 | sbintrunc (n - 1) x = sbintrunc (n - 1) y" | |
| 529 | by (cases n) (simp_all add: bin_sbin_eq_iff del: bintrunc.Suc) | |
| 530 | ||
| 531 | lemmas bintrunc_sbintruncS0 [simp] = bintrunc_sbintrunc' [unfolded One_nat_def] | |
| 532 | lemmas sbintrunc_bintruncS0 [simp] = sbintrunc_bintrunc' [unfolded One_nat_def] | |
| 533 | ||
| 534 | lemmas bintrunc_bintrunc_l' = le_add1 [THEN bintrunc_bintrunc_l] | |
| 535 | lemmas sbintrunc_sbintrunc_l' = le_add1 [THEN sbintrunc_sbintrunc_l] | |
| 536 | ||
| 537 | (* although bintrunc_minus_simps, if added to default simpset, | |
| 538 | tends to get applied where it's not wanted in developing the theories, | |
| 539 | we get a version for when the word length is given literally *) | |
| 540 | ||
| 541 | lemmas nat_non0_gr = | |
| 45604 | 542 | trans [OF iszero_def [THEN Not_eq_iff [THEN iffD2]] refl] | 
| 24333 | 543 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 544 | lemma bintrunc_numeral: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 545 | "bintrunc (numeral k) x = | 
| 47219 
172c031ad743
restate various simp rules for word operations using pred_numeral
 huffman parents: 
47170diff
changeset | 546 | bintrunc (pred_numeral k) (bin_rest x) BIT bin_last x" | 
| 
172c031ad743
restate various simp rules for word operations using pred_numeral
 huffman parents: 
47170diff
changeset | 547 | by (simp add: numeral_eq_Suc) | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 548 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 549 | lemma sbintrunc_numeral: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 550 | "sbintrunc (numeral k) x = | 
| 47219 
172c031ad743
restate various simp rules for word operations using pred_numeral
 huffman parents: 
47170diff
changeset | 551 | sbintrunc (pred_numeral k) (bin_rest x) BIT bin_last x" | 
| 
172c031ad743
restate various simp rules for word operations using pred_numeral
 huffman parents: 
47170diff
changeset | 552 | by (simp add: numeral_eq_Suc) | 
| 24333 | 553 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 554 | lemma bintrunc_numeral_simps [simp]: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 555 | "bintrunc (numeral k) (numeral (Num.Bit0 w)) = | 
| 47219 
172c031ad743
restate various simp rules for word operations using pred_numeral
 huffman parents: 
47170diff
changeset | 556 | bintrunc (pred_numeral k) (numeral w) BIT 0" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 557 | "bintrunc (numeral k) (numeral (Num.Bit1 w)) = | 
| 47219 
172c031ad743
restate various simp rules for word operations using pred_numeral
 huffman parents: 
47170diff
changeset | 558 | bintrunc (pred_numeral k) (numeral w) BIT 1" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 559 | "bintrunc (numeral k) (neg_numeral (Num.Bit0 w)) = | 
| 47219 
172c031ad743
restate various simp rules for word operations using pred_numeral
 huffman parents: 
47170diff
changeset | 560 | bintrunc (pred_numeral k) (neg_numeral w) BIT 0" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 561 | "bintrunc (numeral k) (neg_numeral (Num.Bit1 w)) = | 
| 47219 
172c031ad743
restate various simp rules for word operations using pred_numeral
 huffman parents: 
47170diff
changeset | 562 | bintrunc (pred_numeral k) (neg_numeral (w + Num.One)) BIT 1" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 563 | "bintrunc (numeral k) 1 = 1" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 564 | by (simp_all add: bintrunc_numeral) | 
| 24333 | 565 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 566 | lemma sbintrunc_numeral_simps [simp]: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 567 | "sbintrunc (numeral k) (numeral (Num.Bit0 w)) = | 
| 47219 
172c031ad743
restate various simp rules for word operations using pred_numeral
 huffman parents: 
47170diff
changeset | 568 | sbintrunc (pred_numeral k) (numeral w) BIT 0" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 569 | "sbintrunc (numeral k) (numeral (Num.Bit1 w)) = | 
| 47219 
172c031ad743
restate various simp rules for word operations using pred_numeral
 huffman parents: 
47170diff
changeset | 570 | sbintrunc (pred_numeral k) (numeral w) BIT 1" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 571 | "sbintrunc (numeral k) (neg_numeral (Num.Bit0 w)) = | 
| 47219 
172c031ad743
restate various simp rules for word operations using pred_numeral
 huffman parents: 
47170diff
changeset | 572 | sbintrunc (pred_numeral k) (neg_numeral w) BIT 0" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 573 | "sbintrunc (numeral k) (neg_numeral (Num.Bit1 w)) = | 
| 47219 
172c031ad743
restate various simp rules for word operations using pred_numeral
 huffman parents: 
47170diff
changeset | 574 | sbintrunc (pred_numeral k) (neg_numeral (w + Num.One)) BIT 1" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 575 | "sbintrunc (numeral k) 1 = 1" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 576 | by (simp_all add: sbintrunc_numeral) | 
| 24333 | 577 | |
| 578 | lemma no_bintr_alt1: "bintrunc n = (%w. w mod 2 ^ n :: int)" | |
| 579 | by (rule ext) (rule bintrunc_mod2p) | |
| 580 | ||
| 581 | lemma range_bintrunc: "range (bintrunc n) = {i. 0 <= i & i < 2 ^ n}"
 | |
| 582 | apply (unfold no_bintr_alt1) | |
| 583 | apply (auto simp add: image_iff) | |
| 584 | apply (rule exI) | |
| 585 | apply (auto intro: int_mod_lem [THEN iffD1, symmetric]) | |
| 586 | done | |
| 587 | ||
| 588 | lemma no_sbintr_alt2: | |
| 589 | "sbintrunc n = (%w. (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n :: int)" | |
| 590 | by (rule ext) (simp add : sbintrunc_mod2p) | |
| 591 | ||
| 592 | lemma range_sbintrunc: | |
| 593 |   "range (sbintrunc n) = {i. - (2 ^ n) <= i & i < 2 ^ n}"
 | |
| 594 | apply (unfold no_sbintr_alt2) | |
| 595 | apply (auto simp add: image_iff eq_diff_eq) | |
| 596 | apply (rule exI) | |
| 597 | apply (auto intro: int_mod_lem [THEN iffD1, symmetric]) | |
| 598 | done | |
| 599 | ||
| 25349 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 600 | lemma sb_inc_lem: | 
| 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 601 | "(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 | 602 | 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 | 603 | apply (rule TrueI) | 
| 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 604 | done | 
| 24333 | 605 | |
| 25349 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 606 | lemma sb_inc_lem': | 
| 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 607 | "(a::int) < - (2^k) \<Longrightarrow> a + 2^k + 2^(Suc k) <= (a + 2^k) mod 2^(Suc k)" | 
| 35048 | 608 | by (rule sb_inc_lem) simp | 
| 24333 | 609 | |
| 610 | lemma sbintrunc_inc: | |
| 25349 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 611 | "x < - (2^n) ==> x + 2^(Suc n) <= sbintrunc n x" | 
| 24333 | 612 | unfolding no_sbintr_alt2 by (drule sb_inc_lem') simp | 
| 613 | ||
| 25349 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 614 | lemma sb_dec_lem: | 
| 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 615 | "(0::int) <= - (2^k) + a ==> (a + 2^k) mod (2 * 2 ^ k) <= - (2 ^ k) + a" | 
| 47170 | 616 | by (rule int_mod_le' [where n = "2 ^ (Suc k)" and b = "a + 2 ^ k", simplified]) | 
| 24333 | 617 | |
| 25349 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 618 | lemma sb_dec_lem': | 
| 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 619 | "(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 | 620 | by (rule iffD1 [OF diff_le_eq', THEN sb_dec_lem, simplified]) | 
| 24333 | 621 | |
| 622 | lemma sbintrunc_dec: | |
| 623 | "x >= (2 ^ n) ==> x - 2 ^ (Suc n) >= sbintrunc n x" | |
| 624 | unfolding no_sbintr_alt2 by (drule sb_dec_lem') simp | |
| 625 | ||
| 47168 | 626 | lemmas zmod_uminus' = zminus_zmod [where m=c] for c | 
| 47164 | 627 | lemmas zpower_zmod' = power_mod [where b=c and n=k] for c k | 
| 24333 | 628 | |
| 47168 | 629 | lemmas brdmod1s' [symmetric] = | 
| 630 | mod_add_left_eq mod_add_right_eq | |
| 631 | mod_diff_left_eq mod_diff_right_eq | |
| 632 | mod_mult_left_eq mod_mult_right_eq | |
| 24333 | 633 | |
| 634 | lemmas brdmods' [symmetric] = | |
| 635 | zpower_zmod' [symmetric] | |
| 30034 | 636 | trans [OF mod_add_left_eq mod_add_right_eq] | 
| 47168 | 637 | trans [OF mod_diff_left_eq mod_diff_right_eq] | 
| 638 | trans [OF mod_mult_right_eq mod_mult_left_eq] | |
| 24333 | 639 | zmod_uminus' [symmetric] | 
| 30034 | 640 | mod_add_left_eq [where b = "1::int"] | 
| 47168 | 641 | mod_diff_left_eq [where b = "1::int"] | 
| 24333 | 642 | |
| 643 | lemmas bintr_arith1s = | |
| 46000 | 644 | brdmod1s' [where c="2^n::int", folded bintrunc_mod2p] for n | 
| 24333 | 645 | lemmas bintr_ariths = | 
| 46000 | 646 | brdmods' [where c="2^n::int", folded bintrunc_mod2p] for n | 
| 24333 | 647 | |
| 45604 | 648 | lemmas m2pths = pos_mod_sign pos_mod_bound [OF zless2p] | 
| 24364 | 649 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 650 | lemma bintr_ge0: "0 \<le> bintrunc n w" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 651 | by (simp add: bintrunc_mod2p) | 
| 24333 | 652 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 653 | lemma bintr_lt2p: "bintrunc n w < 2 ^ n" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 654 | by (simp add: bintrunc_mod2p) | 
| 24333 | 655 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 656 | lemma bintr_Min: "bintrunc n -1 = 2 ^ n - 1" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 657 | by (simp add: bintrunc_mod2p m1mod2k) | 
| 24333 | 658 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 659 | lemma sbintr_ge: "- (2 ^ n) \<le> sbintrunc n w" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 660 | by (simp add: sbintrunc_mod2p) | 
| 24333 | 661 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 662 | lemma sbintr_lt: "sbintrunc n w < 2 ^ n" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 663 | by (simp add: sbintrunc_mod2p) | 
| 24333 | 664 | |
| 665 | lemma sign_Pls_ge_0: | |
| 46604 
9f9e85264e4d
make uses of bin_sign respect int/bin distinction
 huffman parents: 
46601diff
changeset | 666 | "(bin_sign bin = 0) = (bin >= (0 :: int))" | 
| 
9f9e85264e4d
make uses of bin_sign respect int/bin distinction
 huffman parents: 
46601diff
changeset | 667 | unfolding bin_sign_def by simp | 
| 24333 | 668 | |
| 669 | lemma sign_Min_lt_0: | |
| 46604 
9f9e85264e4d
make uses of bin_sign respect int/bin distinction
 huffman parents: 
46601diff
changeset | 670 | "(bin_sign bin = -1) = (bin < (0 :: int))" | 
| 
9f9e85264e4d
make uses of bin_sign respect int/bin distinction
 huffman parents: 
46601diff
changeset | 671 | unfolding bin_sign_def by simp | 
| 24333 | 672 | |
| 673 | lemma bin_rest_trunc: | |
| 45954 
f67d3bb5f09c
use 'induct arbitrary' instead of universal quantifiers
 huffman parents: 
45953diff
changeset | 674 | "(bin_rest (bintrunc n bin)) = bintrunc (n - 1) (bin_rest bin)" | 
| 
f67d3bb5f09c
use 'induct arbitrary' instead of universal quantifiers
 huffman parents: 
45953diff
changeset | 675 | by (induct n arbitrary: bin) auto | 
| 24333 | 676 | |
| 677 | lemma bin_rest_power_trunc [rule_format] : | |
| 30971 | 678 | "(bin_rest ^^ k) (bintrunc n bin) = | 
| 679 | bintrunc (n - k) ((bin_rest ^^ k) bin)" | |
| 24333 | 680 | by (induct k) (auto simp: bin_rest_trunc) | 
| 681 | ||
| 682 | lemma bin_rest_trunc_i: | |
| 683 | "bintrunc n (bin_rest bin) = bin_rest (bintrunc (Suc n) bin)" | |
| 684 | by auto | |
| 685 | ||
| 686 | lemma bin_rest_strunc: | |
| 45954 
f67d3bb5f09c
use 'induct arbitrary' instead of universal quantifiers
 huffman parents: 
45953diff
changeset | 687 | "bin_rest (sbintrunc (Suc n) bin) = sbintrunc n (bin_rest bin)" | 
| 
f67d3bb5f09c
use 'induct arbitrary' instead of universal quantifiers
 huffman parents: 
45953diff
changeset | 688 | by (induct n arbitrary: bin) auto | 
| 24333 | 689 | |
| 690 | lemma bintrunc_rest [simp]: | |
| 45954 
f67d3bb5f09c
use 'induct arbitrary' instead of universal quantifiers
 huffman parents: 
45953diff
changeset | 691 | "bintrunc n (bin_rest (bintrunc n bin)) = bin_rest (bintrunc n bin)" | 
| 
f67d3bb5f09c
use 'induct arbitrary' instead of universal quantifiers
 huffman parents: 
45953diff
changeset | 692 | apply (induct n arbitrary: bin, simp) | 
| 24333 | 693 | apply (case_tac bin rule: bin_exhaust) | 
| 694 | apply (auto simp: bintrunc_bintrunc_l) | |
| 695 | done | |
| 696 | ||
| 697 | lemma sbintrunc_rest [simp]: | |
| 45954 
f67d3bb5f09c
use 'induct arbitrary' instead of universal quantifiers
 huffman parents: 
45953diff
changeset | 698 | "sbintrunc n (bin_rest (sbintrunc n bin)) = bin_rest (sbintrunc n bin)" | 
| 
f67d3bb5f09c
use 'induct arbitrary' instead of universal quantifiers
 huffman parents: 
45953diff
changeset | 699 | apply (induct n arbitrary: bin, simp) | 
| 24333 | 700 | apply (case_tac bin rule: bin_exhaust) | 
| 701 | apply (auto simp: bintrunc_bintrunc_l split: bit.splits) | |
| 702 | done | |
| 703 | ||
| 704 | lemma bintrunc_rest': | |
| 705 | "bintrunc n o bin_rest o bintrunc n = bin_rest o bintrunc n" | |
| 706 | by (rule ext) auto | |
| 707 | ||
| 708 | lemma sbintrunc_rest' : | |
| 709 | "sbintrunc n o bin_rest o sbintrunc n = bin_rest o sbintrunc n" | |
| 710 | by (rule ext) auto | |
| 711 | ||
| 712 | lemma rco_lem: | |
| 30971 | 713 | "f o g o f = g o f ==> f o (g o f) ^^ n = g ^^ n o f" | 
| 24333 | 714 | apply (rule ext) | 
| 715 | apply (induct_tac n) | |
| 716 | apply (simp_all (no_asm)) | |
| 717 | apply (drule fun_cong) | |
| 718 | apply (unfold o_def) | |
| 719 | apply (erule trans) | |
| 720 | apply simp | |
| 721 | done | |
| 722 | ||
| 30971 | 723 | lemma rco_alt: "(f o g) ^^ n o f = f o (g o f) ^^ n" | 
| 24333 | 724 | apply (rule ext) | 
| 725 | apply (induct n) | |
| 726 | apply (simp_all add: o_def) | |
| 727 | done | |
| 728 | ||
| 729 | lemmas rco_bintr = bintrunc_rest' | |
| 730 | [THEN rco_lem [THEN fun_cong], unfolded o_def] | |
| 731 | lemmas rco_sbintr = sbintrunc_rest' | |
| 732 | [THEN rco_lem [THEN fun_cong], unfolded o_def] | |
| 733 | ||
| 24364 | 734 | subsection {* Splitting and concatenation *}
 | 
| 735 | ||
| 26557 | 736 | primrec bin_split :: "nat \<Rightarrow> int \<Rightarrow> int \<times> int" where | 
| 46001 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
46000diff
changeset | 737 | Z: "bin_split 0 w = (w, 0)" | 
| 26557 | 738 | | Suc: "bin_split (Suc n) w = (let (w1, w2) = bin_split n (bin_rest w) | 
| 739 | in (w1, w2 BIT bin_last w))" | |
| 24364 | 740 | |
| 37667 | 741 | lemma [code]: | 
| 742 | "bin_split (Suc n) w = (let (w1, w2) = bin_split n (bin_rest w) in (w1, w2 BIT bin_last w))" | |
| 743 | "bin_split 0 w = (w, 0)" | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 744 | by simp_all | 
| 37667 | 745 | |
| 26557 | 746 | primrec bin_cat :: "int \<Rightarrow> nat \<Rightarrow> int \<Rightarrow> int" where | 
| 747 | Z: "bin_cat w 0 v = w" | |
| 748 | | Suc: "bin_cat w (Suc n) v = bin_cat w n (bin_rest v) BIT bin_last v" | |
| 24364 | 749 | |
| 750 | subsection {* Miscellaneous lemmas *}
 | |
| 751 | ||
| 30952 
7ab2716dd93b
power operation on functions with syntax o^; power operation on relations with syntax ^^
 haftmann parents: 
30940diff
changeset | 752 | lemma funpow_minus_simp: | 
| 30971 | 753 | "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 | 754 | by (cases n) simp_all | 
| 24364 | 755 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 756 | lemma funpow_numeral [simp]: | 
| 47219 
172c031ad743
restate various simp rules for word operations using pred_numeral
 huffman parents: 
47170diff
changeset | 757 | "f ^^ numeral k = f \<circ> f ^^ (pred_numeral k)" | 
| 
172c031ad743
restate various simp rules for word operations using pred_numeral
 huffman parents: 
47170diff
changeset | 758 | by (simp add: numeral_eq_Suc) | 
| 24364 | 759 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 760 | lemma replicate_numeral [simp]: (* TODO: move to List.thy *) | 
| 47219 
172c031ad743
restate various simp rules for word operations using pred_numeral
 huffman parents: 
47170diff
changeset | 761 | "replicate (numeral k) x = x # replicate (pred_numeral k) x" | 
| 
172c031ad743
restate various simp rules for word operations using pred_numeral
 huffman parents: 
47170diff
changeset | 762 | by (simp add: numeral_eq_Suc) | 
| 24364 | 763 | |
| 764 | lemmas power_minus_simp = | |
| 45604 | 765 | trans [OF gen_minus [where f = "power f"] power_Suc] for f | 
| 24364 | 766 | |
| 767 | lemma list_exhaust_size_gt0: | |
| 768 | assumes y: "\<And>a list. y = a # list \<Longrightarrow> P" | |
| 769 | shows "0 < length y \<Longrightarrow> P" | |
| 770 | apply (cases y, simp) | |
| 771 | apply (rule y) | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
41413diff
changeset | 772 | apply fastforce | 
| 24364 | 773 | done | 
| 774 | ||
| 775 | lemma list_exhaust_size_eq0: | |
| 776 | assumes y: "y = [] \<Longrightarrow> P" | |
| 777 | shows "length y = 0 \<Longrightarrow> P" | |
| 778 | apply (cases y) | |
| 779 | apply (rule y, simp) | |
| 780 | apply simp | |
| 781 | done | |
| 782 | ||
| 783 | lemma size_Cons_lem_eq: | |
| 784 | "y = xa # list ==> size y = Suc k ==> size list = k" | |
| 785 | by auto | |
| 786 | ||
| 44939 
5930d35c976d
removed unused legacy lemma names, some comment cleanup.
 kleing parents: 
44890diff
changeset | 787 | lemmas ls_splits = prod.split prod.split_asm split_if_asm | 
| 24333 | 788 | |
| 37654 | 789 | lemma not_B1_is_B0: "y \<noteq> (1::bit) \<Longrightarrow> y = (0::bit)" | 
| 24333 | 790 | by (cases y) auto | 
| 791 | ||
| 792 | lemma B1_ass_B0: | |
| 37654 | 793 | assumes y: "y = (0::bit) \<Longrightarrow> y = (1::bit)" | 
| 794 | shows "y = (1::bit)" | |
| 24333 | 795 | apply (rule classical) | 
| 796 | apply (drule not_B1_is_B0) | |
| 797 | apply (erule y) | |
| 798 | done | |
| 799 | ||
| 800 | -- "simplifications for specific word lengths" | |
| 801 | lemmas n2s_ths [THEN eq_reflection] = add_2_eq_Suc add_2_eq_Suc' | |
| 802 | ||
| 803 | lemmas s2n_ths = n2s_ths [symmetric] | |
| 804 | ||
| 805 | end |