| author | nipkow | 
| Wed, 02 Aug 2017 18:22:02 +0200 | |
| changeset 66308 | b6a0d95b94be | 
| parent 65363 | 5eb619751b14 | 
| child 67142 | fa1173288322 | 
| permissions | -rw-r--r-- | 
| 65363 | 1 | (* Title: HOL/Word/Bit_Representation.thy | 
| 2 | Author: Jeremy Dawson, NICTA | |
| 3 | *) | |
| 24333 | 4 | |
| 61799 | 5 | section \<open>Integers as implict bit strings\<close> | 
| 24350 | 6 | |
| 37658 | 7 | theory Bit_Representation | 
| 65363 | 8 | imports Misc_Numeric | 
| 24333 | 9 | begin | 
| 10 | ||
| 61799 | 11 | subsection \<open>Constructors and destructors for binary integers\<close> | 
| 26557 | 12 | |
| 65363 | 13 | definition Bit :: "int \<Rightarrow> bool \<Rightarrow> int" (infixl "BIT" 90) | 
| 14 | where "k BIT b = (if b then 1 else 0) + k + k" | |
| 26557 | 15 | |
| 65363 | 16 | lemma Bit_B0: "k BIT False = k + k" | 
| 17 | by (simp add: Bit_def) | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 18 | |
| 65363 | 19 | lemma Bit_B1: "k BIT True = k + k + 1" | 
| 20 | by (simp add: Bit_def) | |
| 21 | ||
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 22 | lemma Bit_B0_2t: "k BIT False = 2 * k" | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 23 | by (rule trans, rule Bit_B0) simp | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 24 | |
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 25 | lemma Bit_B1_2t: "k BIT True = 2 * k + 1" | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 26 | by (rule trans, rule Bit_B1) simp | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 27 | |
| 54848 | 28 | definition bin_last :: "int \<Rightarrow> bool" | 
| 65363 | 29 | where "bin_last w \<longleftrightarrow> w mod 2 = 1" | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 30 | |
| 65363 | 31 | lemma bin_last_odd: "bin_last = odd" | 
| 58645 | 32 | by (rule ext) (simp add: bin_last_def even_iff_mod_2_eq_zero) | 
| 45843 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 33 | |
| 54848 | 34 | definition bin_rest :: "int \<Rightarrow> int" | 
| 65363 | 35 | where "bin_rest w = w div 2" | 
| 45843 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 36 | |
| 65363 | 37 | lemma bin_rl_simp [simp]: "bin_rest w BIT bin_last w = w" | 
| 45843 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 38 | unfolding bin_rest_def bin_last_def Bit_def | 
| 65363 | 39 | by (cases "w mod 2 = 0") (use div_mult_mod_eq [of w 2] in simp_all) | 
| 45843 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 40 | |
| 46600 | 41 | 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 | 42 | unfolding bin_rest_def Bit_def | 
| 65363 | 43 | by (cases b) simp_all | 
| 45843 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 44 | |
| 46600 | 45 | 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 | 46 | unfolding bin_last_def Bit_def | 
| 54873 | 47 | by (cases b) simp_all | 
| 45843 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 48 | |
| 45848 | 49 | lemma BIT_eq_iff [iff]: "u BIT b = v BIT c \<longleftrightarrow> u = v \<and> b = c" | 
| 65363 | 50 | by (auto simp: Bit_def) arith+ | 
| 45843 
c58ce659ce2a
reorder some definitions and proofs, in preparation for new numeral representation
 huffman parents: 
45604diff
changeset | 51 | |
| 45849 | 52 | lemma BIT_bin_simps [simp]: | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 53 | "numeral k BIT False = numeral (Num.Bit0 k)" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 54 | "numeral k BIT True = numeral (Num.Bit1 k)" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 55 | "(- numeral k) BIT False = - numeral (Num.Bit0 k)" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 56 | "(- numeral k) BIT True = - numeral (Num.BitM k)" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 57 | unfolding numeral.simps numeral_BitM | 
| 65363 | 58 | by (simp_all add: Bit_def del: arith_simps add_numeral_special diff_numeral_special) | 
| 45849 | 59 | |
| 60 | lemma BIT_special_simps [simp]: | |
| 65363 | 61 | shows "0 BIT False = 0" | 
| 62 | and "0 BIT True = 1" | |
| 63 | and "1 BIT False = 2" | |
| 64 | and "1 BIT True = 3" | |
| 65 | and "(- 1) BIT False = - 2" | |
| 66 | and "(- 1) BIT True = - 1" | |
| 67 | by (simp_all add: Bit_def) | |
| 45849 | 68 | |
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 69 | lemma Bit_eq_0_iff: "w BIT b = 0 \<longleftrightarrow> w = 0 \<and> \<not> b" | 
| 65363 | 70 | by (auto simp: Bit_def) arith | 
| 53438 | 71 | |
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 72 | lemma Bit_eq_m1_iff: "w BIT b = -1 \<longleftrightarrow> w = -1 \<and> b" | 
| 65363 | 73 | by (auto simp: Bit_def) arith | 
| 53438 | 74 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 75 | lemma BitM_inc: "Num.BitM (Num.inc w) = Num.Bit1 w" | 
| 65363 | 76 | by (induct w) simp_all | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 77 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 78 | lemma expand_BIT: | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 79 | "numeral (Num.Bit0 w) = numeral w BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 80 | "numeral (Num.Bit1 w) = numeral w BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 81 | "- numeral (Num.Bit0 w) = (- numeral w) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 82 | "- numeral (Num.Bit1 w) = (- numeral (w + Num.One)) BIT True" | 
| 65363 | 83 | by (simp_all add: add_One BitM_inc) | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 84 | |
| 45851 
19f7ac6cf3cc
add simp rules for bin_rest and bin_last applied to numerals
 huffman parents: 
45850diff
changeset | 85 | lemma bin_last_numeral_simps [simp]: | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 86 | "\<not> bin_last 0" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 87 | "bin_last 1" | 
| 58410 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
54873diff
changeset | 88 | "bin_last (- 1)" | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 89 | "bin_last Numeral1" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 90 | "\<not> bin_last (numeral (Num.Bit0 w))" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 91 | "bin_last (numeral (Num.Bit1 w))" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 92 | "\<not> bin_last (- numeral (Num.Bit0 w))" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 93 | "bin_last (- numeral (Num.Bit1 w))" | 
| 60867 | 94 | by (simp_all add: bin_last_def zmod_zminus1_eq_if) (auto simp add: divmod_def) | 
| 45851 
19f7ac6cf3cc
add simp rules for bin_rest and bin_last applied to numerals
 huffman parents: 
45850diff
changeset | 95 | |
| 
19f7ac6cf3cc
add simp rules for bin_rest and bin_last applied to numerals
 huffman parents: 
45850diff
changeset | 96 | lemma bin_rest_numeral_simps [simp]: | 
| 
19f7ac6cf3cc
add simp rules for bin_rest and bin_last applied to numerals
 huffman parents: 
45850diff
changeset | 97 | "bin_rest 0 = 0" | 
| 
19f7ac6cf3cc
add simp rules for bin_rest and bin_last applied to numerals
 huffman parents: 
45850diff
changeset | 98 | "bin_rest 1 = 0" | 
| 58410 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
54873diff
changeset | 99 | "bin_rest (- 1) = - 1" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 100 | "bin_rest Numeral1 = 0" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 101 | "bin_rest (numeral (Num.Bit0 w)) = numeral w" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 102 | "bin_rest (numeral (Num.Bit1 w)) = numeral w" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 103 | "bin_rest (- numeral (Num.Bit0 w)) = - numeral w" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 104 | "bin_rest (- numeral (Num.Bit1 w)) = - numeral (w + Num.One)" | 
| 60867 | 105 | by (simp_all add: bin_rest_def zdiv_zminus1_eq_if) (auto simp add: divmod_def) | 
| 26557 | 106 | |
| 65363 | 107 | lemma less_Bits: "v BIT b < w BIT c \<longleftrightarrow> v < w \<or> v \<le> w \<and> \<not> b \<and> c" | 
| 108 | by (auto simp: Bit_def) | |
| 24333 | 109 | |
| 65363 | 110 | lemma le_Bits: "v BIT b \<le> w BIT c \<longleftrightarrow> v < w \<or> v \<le> w \<and> (\<not> b \<or> c)" | 
| 111 | by (auto simp: Bit_def) | |
| 26557 | 112 | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 113 | lemma pred_BIT_simps [simp]: | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 114 | "x BIT False - 1 = (x - 1) BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 115 | "x BIT True - 1 = x BIT False" | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 116 | by (simp_all add: Bit_B0_2t Bit_B1_2t) | 
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 117 | |
| 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 118 | lemma succ_BIT_simps [simp]: | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 119 | "x BIT False + 1 = x BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 120 | "x BIT True + 1 = (x + 1) BIT False" | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 121 | by (simp_all add: Bit_B0_2t Bit_B1_2t) | 
| 26557 | 122 | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 123 | lemma add_BIT_simps [simp]: | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 124 | "x BIT False + y BIT False = (x + y) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 125 | "x BIT False + y BIT True = (x + y) BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 126 | "x BIT True + y BIT False = (x + y) BIT True" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 127 | "x BIT True + y BIT True = (x + y + 1) BIT False" | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 128 | by (simp_all add: Bit_B0_2t Bit_B1_2t) | 
| 26557 | 129 | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 130 | lemma mult_BIT_simps [simp]: | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 131 | "x BIT False * y = (x * y) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 132 | "x * y BIT False = (x * y) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 133 | "x BIT True * y = (x * y) BIT False + y" | 
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 134 | by (simp_all add: Bit_B0_2t Bit_B1_2t algebra_simps) | 
| 26557 | 135 | |
| 65363 | 136 | lemma B_mod_2': "X = 2 \<Longrightarrow> (w BIT True) mod X = 1 \<and> (w BIT False) mod X = 0" | 
| 137 | by (simp add: Bit_B0 Bit_B1) | |
| 24333 | 138 | |
| 65363 | 139 | lemma bin_ex_rl: "\<exists>w b. w BIT b = bin" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 140 | 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 | 141 | |
| 26557 | 142 | lemma bin_exhaust: | 
| 65363 | 143 | assumes that: "\<And>x b. bin = x BIT b \<Longrightarrow> Q" | 
| 26557 | 144 | shows "Q" | 
| 65363 | 145 | apply (insert bin_ex_rl [of bin]) | 
| 26557 | 146 | apply (erule exE)+ | 
| 65363 | 147 | apply (rule that) | 
| 26557 | 148 | apply force | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 149 | done | 
| 24333 | 150 | |
| 65363 | 151 | primrec bin_nth | 
| 152 | where | |
| 153 | Z: "bin_nth w 0 \<longleftrightarrow> bin_last w" | |
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 154 | | Suc: "bin_nth w (Suc n) \<longleftrightarrow> bin_nth (bin_rest w) n" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 155 | |
| 65363 | 156 | lemma bin_abs_lem: "bin = (w BIT b) \<Longrightarrow> bin \<noteq> -1 \<longrightarrow> bin \<noteq> 0 \<longrightarrow> nat \<bar>w\<bar> < nat \<bar>bin\<bar>" | 
| 46598 | 157 | apply clarsimp | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 158 | apply (unfold Bit_def) | 
| 26557 | 159 | apply (cases b) | 
| 160 | apply (clarsimp, arith) | |
| 161 | apply (clarsimp, arith) | |
| 162 | done | |
| 163 | ||
| 164 | lemma bin_induct: | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 165 | assumes PPls: "P 0" | 
| 58410 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
54873diff
changeset | 166 | and PMin: "P (- 1)" | 
| 65363 | 167 | and PBit: "\<And>bin bit. P bin \<Longrightarrow> P (bin BIT bit)" | 
| 26557 | 168 | shows "P bin" | 
| 65363 | 169 | apply (rule_tac P=P and a=bin and f1="nat \<circ> abs" in wf_measure [THEN wf_induct]) | 
| 26557 | 170 | apply (simp add: measure_def inv_image_def) | 
| 171 | apply (case_tac x rule: bin_exhaust) | |
| 172 | apply (frule bin_abs_lem) | |
| 173 | apply (auto simp add : PPls PMin PBit) | |
| 174 | done | |
| 175 | ||
| 24333 | 176 | lemma Bit_div2 [simp]: "(w BIT b) div 2 = w" | 
| 46600 | 177 | unfolding bin_rest_def [symmetric] by (rule bin_rest_BIT) | 
| 24333 | 178 | |
| 65363 | 179 | lemma bin_nth_eq_iff: "bin_nth x = bin_nth y \<longleftrightarrow> x = y" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 180 | proof - | 
| 65363 | 181 | have bin_nth_lem [rule_format]: "\<forall>y. bin_nth x = bin_nth y \<longrightarrow> x = y" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 182 | apply (induct x 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 | 183 | apply safe | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 184 | apply (erule rev_mp) | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 185 | apply (induct_tac y rule: bin_induct) | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 186 | apply safe | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 187 | apply (drule_tac x=0 in fun_cong, force) | 
| 65363 | 188 | apply (erule notE, rule ext, drule_tac x="Suc x" in fun_cong, force) | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 189 | apply (drule_tac x=0 in fun_cong, force) | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 190 | apply (erule rev_mp) | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 191 | apply (induct_tac y rule: bin_induct) | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 192 | apply safe | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 193 | apply (drule_tac x=0 in fun_cong, force) | 
| 65363 | 194 | apply (erule notE, rule ext, drule_tac x="Suc x" in fun_cong, force) | 
| 195 | apply (metis Bit_eq_m1_iff Z bin_last_BIT) | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 196 | apply (case_tac y rule: bin_exhaust) | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 197 | apply clarify | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 198 | apply (erule allE) | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 199 | apply (erule impE) | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 200 | prefer 2 | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 201 | apply (erule conjI) | 
| 24333 | 202 | apply (drule_tac x=0 in fun_cong, force) | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 203 | apply (rule ext) | 
| 59807 | 204 | apply (drule_tac x="Suc x" for x in fun_cong, force) | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 205 | done | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 206 | show ?thesis | 
| 65363 | 207 | by (auto elim: bin_nth_lem) | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 208 | qed | 
| 24333 | 209 | |
| 45604 | 210 | lemmas bin_eqI = ext [THEN bin_nth_eq_iff [THEN iffD1]] | 
| 24333 | 211 | |
| 65363 | 212 | lemma bin_eq_iff: "x = y \<longleftrightarrow> (\<forall>n. bin_nth x n = bin_nth y n)" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 213 | using bin_nth_eq_iff by auto | 
| 45543 
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
 huffman parents: 
45529diff
changeset | 214 | |
| 45853 | 215 | lemma bin_nth_zero [simp]: "\<not> bin_nth 0 n" | 
| 216 | by (induct n) auto | |
| 217 | ||
| 46023 
fad87bb608fc
restate some lemmas to respect int/bin distinction
 huffman parents: 
46001diff
changeset | 218 | 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 | 219 | by (cases n) simp_all | 
| 
fad87bb608fc
restate some lemmas to respect int/bin distinction
 huffman parents: 
46001diff
changeset | 220 | |
| 58410 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
54873diff
changeset | 221 | lemma bin_nth_minus1 [simp]: "bin_nth (- 1) n" | 
| 45952 | 222 | by (induct n) auto | 
| 223 | ||
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 224 | lemma bin_nth_0_BIT: "bin_nth (w BIT b) 0 \<longleftrightarrow> b" | 
| 24333 | 225 | by auto | 
| 226 | ||
| 227 | lemma bin_nth_Suc_BIT: "bin_nth (w BIT b) (Suc n) = bin_nth w n" | |
| 228 | by auto | |
| 229 | ||
| 65363 | 230 | lemma bin_nth_minus [simp]: "0 < n \<Longrightarrow> bin_nth (w BIT b) n = bin_nth w (n - 1)" | 
| 24333 | 231 | by (cases n) auto | 
| 232 | ||
| 65363 | 233 | lemma bin_nth_numeral: "bin_rest x = y \<Longrightarrow> bin_nth x (numeral n) = bin_nth y (pred_numeral n)" | 
| 47219 
172c031ad743
restate various simp rules for word operations using pred_numeral
 huffman parents: 
47170diff
changeset | 234 | 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 | 235 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 236 | lemmas bin_nth_numeral_simps [simp] = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 237 | bin_nth_numeral [OF bin_rest_numeral_simps(2)] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 238 | bin_nth_numeral [OF bin_rest_numeral_simps(5)] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 239 | bin_nth_numeral [OF bin_rest_numeral_simps(6)] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 240 | bin_nth_numeral [OF bin_rest_numeral_simps(7)] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 241 | bin_nth_numeral [OF bin_rest_numeral_simps(8)] | 
| 24333 | 242 | |
| 65363 | 243 | lemmas bin_nth_simps = | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 244 | 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 | 245 | bin_nth_numeral_simps | 
| 24333 | 246 | |
| 26557 | 247 | |
| 61799 | 248 | subsection \<open>Truncating binary integers\<close> | 
| 26557 | 249 | |
| 54848 | 250 | definition bin_sign :: "int \<Rightarrow> int" | 
| 65363 | 251 | where "bin_sign k = (if k \<ge> 0 then 0 else - 1)" | 
| 26557 | 252 | |
| 253 | lemma bin_sign_simps [simp]: | |
| 45850 | 254 | "bin_sign 0 = 0" | 
| 46604 
9f9e85264e4d
make uses of bin_sign respect int/bin distinction
 huffman parents: 
46601diff
changeset | 255 | "bin_sign 1 = 0" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 256 | "bin_sign (- 1) = - 1" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 257 | "bin_sign (numeral k) = 0" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 258 | "bin_sign (- numeral k) = -1" | 
| 26557 | 259 | "bin_sign (w BIT b) = bin_sign w" | 
| 65363 | 260 | by (simp_all add: bin_sign_def Bit_def) | 
| 26557 | 261 | |
| 65363 | 262 | lemma bin_sign_rest [simp]: "bin_sign (bin_rest w) = bin_sign w" | 
| 26557 | 263 | by (cases w rule: bin_exhaust) auto | 
| 24364 | 264 | |
| 65363 | 265 | primrec bintrunc :: "nat \<Rightarrow> int \<Rightarrow> int" | 
| 266 | where | |
| 267 | Z : "bintrunc 0 bin = 0" | |
| 268 | | Suc : "bintrunc (Suc n) bin = bintrunc n (bin_rest bin) BIT (bin_last bin)" | |
| 24364 | 269 | |
| 65363 | 270 | primrec sbintrunc :: "nat \<Rightarrow> int \<Rightarrow> int" | 
| 271 | where | |
| 272 | Z : "sbintrunc 0 bin = (if bin_last bin then -1 else 0)" | |
| 273 | | Suc : "sbintrunc (Suc n) bin = sbintrunc n (bin_rest bin) BIT (bin_last bin)" | |
| 37667 | 274 | |
| 46001 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
46000diff
changeset | 275 | lemma sign_bintr: "bin_sign (bintrunc n w) = 0" | 
| 45954 
f67d3bb5f09c
use 'induct arbitrary' instead of universal quantifiers
 huffman parents: 
45953diff
changeset | 276 | by (induct n arbitrary: w) auto | 
| 24333 | 277 | |
| 65363 | 278 | lemma bintrunc_mod2p: "bintrunc n w = w mod 2 ^ n" | 
| 279 | by (induct n arbitrary: w) (auto simp add: bin_last_def bin_rest_def Bit_def zmod_zmult2_eq) | |
| 24333 | 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) | 
| 64593 
50c715579715
reoriented congruence rules in non-explosive direction
 haftmann parents: 
64246diff
changeset | 283 | apply (auto simp add: bin_last_odd bin_rest_def Bit_def elim!: evenE oddE) | 
| 
50c715579715
reoriented congruence rules in non-explosive direction
 haftmann parents: 
64246diff
changeset | 284 | apply (smt pos_zmod_mult_2 zle2p) | 
| 
50c715579715
reoriented congruence rules in non-explosive direction
 haftmann parents: 
64246diff
changeset | 285 | apply (simp add: mult_mod_right) | 
| 24333 | 286 | done | 
| 287 | ||
| 64593 
50c715579715
reoriented congruence rules in non-explosive direction
 haftmann parents: 
64246diff
changeset | 288 | |
| 24465 | 289 | subsection "Simplifications for (s)bintrunc" | 
| 290 | ||
| 45852 | 291 | 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 | 292 | by (induct n) auto | 
| 45852 | 293 | |
| 45855 | 294 | 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 | 295 | by (induct n) auto | 
| 45855 | 296 | |
| 58410 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
54873diff
changeset | 297 | 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 | 298 | by (induct n) auto | 
| 45856 | 299 | |
| 45852 | 300 | lemma bintrunc_Suc_numeral: | 
| 301 | "bintrunc (Suc n) 1 = 1" | |
| 58410 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
54873diff
changeset | 302 | "bintrunc (Suc n) (- 1) = bintrunc n (- 1) BIT True" | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 303 | "bintrunc (Suc n) (numeral (Num.Bit0 w)) = bintrunc n (numeral w) BIT False" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 304 | "bintrunc (Suc n) (numeral (Num.Bit1 w)) = bintrunc n (numeral w) BIT True" | 
| 65363 | 305 | "bintrunc (Suc n) (- numeral (Num.Bit0 w)) = bintrunc n (- numeral w) BIT False" | 
| 306 | "bintrunc (Suc n) (- numeral (Num.Bit1 w)) = bintrunc n (- numeral (w + Num.One)) BIT True" | |
| 45852 | 307 | by simp_all | 
| 308 | ||
| 45856 | 309 | lemma sbintrunc_0_numeral [simp]: | 
| 310 | "sbintrunc 0 1 = -1" | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 311 | "sbintrunc 0 (numeral (Num.Bit0 w)) = 0" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 312 | "sbintrunc 0 (numeral (Num.Bit1 w)) = -1" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 313 | "sbintrunc 0 (- numeral (Num.Bit0 w)) = 0" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 314 | "sbintrunc 0 (- 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 | 315 | by simp_all | 
| 45856 | 316 | |
| 45855 | 317 | lemma sbintrunc_Suc_numeral: | 
| 318 | "sbintrunc (Suc n) 1 = 1" | |
| 65363 | 319 | "sbintrunc (Suc n) (numeral (Num.Bit0 w)) = sbintrunc n (numeral w) BIT False" | 
| 320 | "sbintrunc (Suc n) (numeral (Num.Bit1 w)) = sbintrunc n (numeral w) BIT True" | |
| 321 | "sbintrunc (Suc n) (- numeral (Num.Bit0 w)) = sbintrunc n (- numeral w) BIT False" | |
| 322 | "sbintrunc (Suc n) (- numeral (Num.Bit1 w)) = sbintrunc n (- numeral (w + Num.One)) BIT True" | |
| 45855 | 323 | by simp_all | 
| 324 | ||
| 46001 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 huffman parents: 
46000diff
changeset | 325 | 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 | 326 | apply (induct n arbitrary: bin) | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 327 | apply (case_tac bin rule: bin_exhaust, case_tac b, auto) | 
| 24333 | 328 | done | 
| 329 | ||
| 65363 | 330 | lemma nth_bintr: "bin_nth (bintrunc m w) n \<longleftrightarrow> n < m \<and> bin_nth w n" | 
| 45954 
f67d3bb5f09c
use 'induct arbitrary' instead of universal quantifiers
 huffman parents: 
45953diff
changeset | 331 | apply (induct n arbitrary: w m) | 
| 24333 | 332 | apply (case_tac m, auto)[1] | 
| 333 | apply (case_tac m, auto)[1] | |
| 334 | done | |
| 335 | ||
| 65363 | 336 | lemma nth_sbintr: "bin_nth (sbintrunc m w) n = (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 | 337 | apply (induct n arbitrary: w m) | 
| 65363 | 338 | apply (case_tac m) | 
| 339 | apply simp_all | |
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 340 | apply (case_tac m) | 
| 65363 | 341 | apply simp_all | 
| 24333 | 342 | done | 
| 343 | ||
| 65363 | 344 | lemma bin_nth_Bit: "bin_nth (w BIT b) n \<longleftrightarrow> n = 0 \<and> b \<or> (\<exists>m. n = Suc m \<and> bin_nth w m)" | 
| 24333 | 345 | by (cases n) auto | 
| 346 | ||
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 347 | lemma bin_nth_Bit0: | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 348 | "bin_nth (numeral (Num.Bit0 w)) n \<longleftrightarrow> | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 349 | (\<exists>m. n = Suc m \<and> bin_nth (numeral w) m)" | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 350 | using bin_nth_Bit [where w="numeral w" and b="False"] by simp | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 351 | |
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 352 | lemma bin_nth_Bit1: | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 353 | "bin_nth (numeral (Num.Bit1 w)) n \<longleftrightarrow> | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 354 | n = 0 \<or> (\<exists>m. n = Suc m \<and> bin_nth (numeral w) m)" | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 355 | using bin_nth_Bit [where w="numeral w" and b="True"] by simp | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 356 | |
| 65363 | 357 | lemma bintrunc_bintrunc_l: "n \<le> m \<Longrightarrow> bintrunc m (bintrunc n w) = bintrunc n w" | 
| 358 | by (rule bin_eqI) (auto simp: nth_bintr) | |
| 24333 | 359 | |
| 65363 | 360 | lemma sbintrunc_sbintrunc_l: "n \<le> m \<Longrightarrow> sbintrunc m (sbintrunc n w) = sbintrunc n w" | 
| 32439 | 361 | by (rule bin_eqI) (auto simp: nth_sbintr) | 
| 24333 | 362 | |
| 65363 | 363 | lemma bintrunc_bintrunc_ge: "n \<le> m \<Longrightarrow> bintrunc n (bintrunc m w) = bintrunc n w" | 
| 24333 | 364 | by (rule bin_eqI) (auto simp: nth_bintr) | 
| 365 | ||
| 65363 | 366 | lemma bintrunc_bintrunc_min [simp]: "bintrunc m (bintrunc n w) = bintrunc (min m n) w" | 
| 367 | by (rule bin_eqI) (auto simp: nth_bintr) | |
| 24333 | 368 | |
| 65363 | 369 | lemma sbintrunc_sbintrunc_min [simp]: "sbintrunc m (sbintrunc n w) = sbintrunc (min m n) w" | 
| 370 | by (rule bin_eqI) (auto simp: nth_sbintr min.absorb1 min.absorb2) | |
| 24333 | 371 | |
| 65363 | 372 | lemmas bintrunc_Pls = | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 373 | bintrunc.Suc [where bin="0", simplified bin_last_numeral_simps bin_rest_numeral_simps] | 
| 24333 | 374 | |
| 65363 | 375 | lemmas bintrunc_Min [simp] = | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 376 | bintrunc.Suc [where bin="-1", simplified bin_last_numeral_simps bin_rest_numeral_simps] | 
| 24333 | 377 | |
| 65363 | 378 | lemmas bintrunc_BIT [simp] = | 
| 46600 | 379 | bintrunc.Suc [where bin="w BIT b", simplified bin_last_BIT bin_rest_BIT] for w b | 
| 24333 | 380 | |
| 381 | lemmas bintrunc_Sucs = bintrunc_Pls bintrunc_Min bintrunc_BIT | |
| 45852 | 382 | bintrunc_Suc_numeral | 
| 24333 | 383 | |
| 65363 | 384 | lemmas sbintrunc_Suc_Pls = | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 385 | sbintrunc.Suc [where bin="0", simplified bin_last_numeral_simps bin_rest_numeral_simps] | 
| 24333 | 386 | |
| 65363 | 387 | lemmas sbintrunc_Suc_Min = | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 388 | sbintrunc.Suc [where bin="-1", simplified bin_last_numeral_simps bin_rest_numeral_simps] | 
| 24333 | 389 | |
| 65363 | 390 | lemmas sbintrunc_Suc_BIT [simp] = | 
| 46600 | 391 | sbintrunc.Suc [where bin="w BIT b", simplified bin_last_BIT bin_rest_BIT] for w b | 
| 24333 | 392 | |
| 393 | lemmas sbintrunc_Sucs = sbintrunc_Suc_Pls sbintrunc_Suc_Min sbintrunc_Suc_BIT | |
| 45855 | 394 | sbintrunc_Suc_numeral | 
| 24333 | 395 | |
| 65363 | 396 | lemmas sbintrunc_Pls = | 
| 397 | sbintrunc.Z [where bin="0", simplified bin_last_numeral_simps bin_rest_numeral_simps] | |
| 24333 | 398 | |
| 65363 | 399 | lemmas sbintrunc_Min = | 
| 400 | sbintrunc.Z [where bin="-1", simplified bin_last_numeral_simps bin_rest_numeral_simps] | |
| 24333 | 401 | |
| 65363 | 402 | lemmas sbintrunc_0_BIT_B0 [simp] = | 
| 403 | sbintrunc.Z [where bin="w BIT False", simplified bin_last_numeral_simps bin_rest_numeral_simps] | |
| 404 | for w | |
| 24333 | 405 | |
| 65363 | 406 | lemmas sbintrunc_0_BIT_B1 [simp] = | 
| 407 | sbintrunc.Z [where bin="w BIT True", simplified bin_last_BIT bin_rest_numeral_simps] | |
| 408 | for w | |
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25919diff
changeset | 409 | |
| 24333 | 410 | lemmas sbintrunc_0_simps = | 
| 411 | sbintrunc_Pls sbintrunc_Min sbintrunc_0_BIT_B0 sbintrunc_0_BIT_B1 | |
| 412 | ||
| 413 | lemmas bintrunc_simps = bintrunc.Z bintrunc_Sucs | |
| 414 | lemmas sbintrunc_simps = sbintrunc_0_simps sbintrunc_Sucs | |
| 415 | ||
| 65363 | 416 | lemma bintrunc_minus: "0 < n \<Longrightarrow> bintrunc (Suc (n - 1)) w = bintrunc n w" | 
| 417 | by auto | |
| 418 | ||
| 419 | lemma sbintrunc_minus: "0 < n \<Longrightarrow> sbintrunc (Suc (n - 1)) w = sbintrunc n w" | |
| 24333 | 420 | by auto | 
| 421 | ||
| 65363 | 422 | lemmas bintrunc_minus_simps = | 
| 45604 | 423 | bintrunc_Sucs [THEN [2] bintrunc_minus [symmetric, THEN trans]] | 
| 65363 | 424 | lemmas sbintrunc_minus_simps = | 
| 45604 | 425 | sbintrunc_Sucs [THEN [2] sbintrunc_minus [symmetric, THEN trans]] | 
| 24333 | 426 | |
| 65363 | 427 | lemmas thobini1 = arg_cong [where f = "\<lambda>w. w BIT b"] for b | 
| 24333 | 428 | |
| 429 | lemmas bintrunc_BIT_I = trans [OF bintrunc_BIT thobini1] | |
| 430 | lemmas bintrunc_Min_I = trans [OF bintrunc_Min thobini1] | |
| 431 | ||
| 45604 | 432 | lemmas bmsts = bintrunc_minus_simps(1-3) [THEN thobini1 [THEN [2] trans]] | 
| 24333 | 433 | lemmas bintrunc_Pls_minus_I = bmsts(1) | 
| 434 | lemmas bintrunc_Min_minus_I = bmsts(2) | |
| 435 | lemmas bintrunc_BIT_minus_I = bmsts(3) | |
| 436 | ||
| 65363 | 437 | lemma bintrunc_Suc_lem: "bintrunc (Suc n) x = y \<Longrightarrow> m = Suc n \<Longrightarrow> bintrunc m x = y" | 
| 24333 | 438 | by auto | 
| 439 | ||
| 65363 | 440 | lemmas bintrunc_Suc_Ialts = | 
| 45604 | 441 | bintrunc_Min_I [THEN bintrunc_Suc_lem] | 
| 442 | bintrunc_BIT_I [THEN bintrunc_Suc_lem] | |
| 24333 | 443 | |
| 444 | lemmas sbintrunc_BIT_I = trans [OF sbintrunc_Suc_BIT thobini1] | |
| 445 | ||
| 65363 | 446 | lemmas sbintrunc_Suc_Is = | 
| 45604 | 447 | sbintrunc_Sucs(1-3) [THEN thobini1 [THEN [2] trans]] | 
| 24333 | 448 | |
| 65363 | 449 | lemmas sbintrunc_Suc_minus_Is = | 
| 45604 | 450 | sbintrunc_minus_simps(1-3) [THEN thobini1 [THEN [2] trans]] | 
| 24333 | 451 | |
| 65363 | 452 | lemma sbintrunc_Suc_lem: "sbintrunc (Suc n) x = y \<Longrightarrow> m = Suc n \<Longrightarrow> sbintrunc m x = y" | 
| 24333 | 453 | by auto | 
| 454 | ||
| 65363 | 455 | lemmas sbintrunc_Suc_Ialts = | 
| 45604 | 456 | sbintrunc_Suc_Is [THEN sbintrunc_Suc_lem] | 
| 24333 | 457 | |
| 65363 | 458 | lemma sbintrunc_bintrunc_lt: "m > n \<Longrightarrow> sbintrunc n (bintrunc m w) = sbintrunc n w" | 
| 24333 | 459 | by (rule bin_eqI) (auto simp: nth_sbintr nth_bintr) | 
| 460 | ||
| 65363 | 461 | lemma bintrunc_sbintrunc_le: "m \<le> Suc n \<Longrightarrow> bintrunc m (sbintrunc n w) = bintrunc m w" | 
| 24333 | 462 | apply (rule bin_eqI) | 
| 463 | apply (auto simp: nth_sbintr nth_bintr) | |
| 464 | apply (subgoal_tac "x=n", safe, arith+)[1] | |
| 465 | apply (subgoal_tac "x=n", safe, arith+)[1] | |
| 466 | done | |
| 467 | ||
| 468 | lemmas bintrunc_sbintrunc [simp] = order_refl [THEN bintrunc_sbintrunc_le] | |
| 469 | lemmas sbintrunc_bintrunc [simp] = lessI [THEN sbintrunc_bintrunc_lt] | |
| 470 | lemmas bintrunc_bintrunc [simp] = order_refl [THEN bintrunc_bintrunc_l] | |
| 65363 | 471 | lemmas sbintrunc_sbintrunc [simp] = order_refl [THEN sbintrunc_sbintrunc_l] | 
| 24333 | 472 | |
| 65363 | 473 | lemma bintrunc_sbintrunc' [simp]: "0 < n \<Longrightarrow> bintrunc n (sbintrunc (n - 1) w) = bintrunc n w" | 
| 24333 | 474 | by (cases n) (auto simp del: bintrunc.Suc) | 
| 475 | ||
| 65363 | 476 | lemma sbintrunc_bintrunc' [simp]: "0 < n \<Longrightarrow> sbintrunc (n - 1) (bintrunc n w) = sbintrunc (n - 1) w" | 
| 24333 | 477 | by (cases n) (auto simp del: bintrunc.Suc) | 
| 478 | ||
| 65363 | 479 | lemma bin_sbin_eq_iff: "bintrunc (Suc n) x = bintrunc (Suc n) y \<longleftrightarrow> sbintrunc n x = sbintrunc n y" | 
| 24333 | 480 | apply (rule iffI) | 
| 481 | apply (rule box_equals [OF _ sbintrunc_bintrunc sbintrunc_bintrunc]) | |
| 482 | apply simp | |
| 483 | apply (rule box_equals [OF _ bintrunc_sbintrunc bintrunc_sbintrunc]) | |
| 484 | apply simp | |
| 485 | done | |
| 486 | ||
| 487 | lemma bin_sbin_eq_iff': | |
| 65363 | 488 | "0 < n \<Longrightarrow> bintrunc n x = bintrunc n y \<longleftrightarrow> sbintrunc (n - 1) x = sbintrunc (n - 1) y" | 
| 24333 | 489 | by (cases n) (simp_all add: bin_sbin_eq_iff del: bintrunc.Suc) | 
| 490 | ||
| 491 | lemmas bintrunc_sbintruncS0 [simp] = bintrunc_sbintrunc' [unfolded One_nat_def] | |
| 492 | lemmas sbintrunc_bintruncS0 [simp] = sbintrunc_bintrunc' [unfolded One_nat_def] | |
| 493 | ||
| 494 | lemmas bintrunc_bintrunc_l' = le_add1 [THEN bintrunc_bintrunc_l] | |
| 495 | lemmas sbintrunc_sbintrunc_l' = le_add1 [THEN sbintrunc_sbintrunc_l] | |
| 496 | ||
| 497 | (* although bintrunc_minus_simps, if added to default simpset, | |
| 498 | tends to get applied where it's not wanted in developing the theories, | |
| 499 | we get a version for when the word length is given literally *) | |
| 500 | ||
| 65363 | 501 | lemmas nat_non0_gr = | 
| 45604 | 502 | trans [OF iszero_def [THEN Not_eq_iff [THEN iffD2]] refl] | 
| 24333 | 503 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 504 | lemma bintrunc_numeral: | 
| 65363 | 505 | "bintrunc (numeral k) x = bintrunc (pred_numeral k) (bin_rest x) BIT bin_last x" | 
| 47219 
172c031ad743
restate various simp rules for word operations using pred_numeral
 huffman parents: 
47170diff
changeset | 506 | by (simp add: numeral_eq_Suc) | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 507 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 508 | lemma sbintrunc_numeral: | 
| 65363 | 509 | "sbintrunc (numeral k) x = sbintrunc (pred_numeral k) (bin_rest x) BIT bin_last x" | 
| 47219 
172c031ad743
restate various simp rules for word operations using pred_numeral
 huffman parents: 
47170diff
changeset | 510 | by (simp add: numeral_eq_Suc) | 
| 24333 | 511 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 512 | lemma bintrunc_numeral_simps [simp]: | 
| 65363 | 513 | "bintrunc (numeral k) (numeral (Num.Bit0 w)) = bintrunc (pred_numeral k) (numeral w) BIT False" | 
| 514 | "bintrunc (numeral k) (numeral (Num.Bit1 w)) = bintrunc (pred_numeral k) (numeral w) BIT True" | |
| 515 | "bintrunc (numeral k) (- numeral (Num.Bit0 w)) = bintrunc (pred_numeral k) (- numeral w) BIT False" | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 516 | "bintrunc (numeral k) (- numeral (Num.Bit1 w)) = | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 517 | bintrunc (pred_numeral k) (- numeral (w + Num.One)) BIT True" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 518 | "bintrunc (numeral k) 1 = 1" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 519 | by (simp_all add: bintrunc_numeral) | 
| 24333 | 520 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 521 | lemma sbintrunc_numeral_simps [simp]: | 
| 65363 | 522 | "sbintrunc (numeral k) (numeral (Num.Bit0 w)) = sbintrunc (pred_numeral k) (numeral w) BIT False" | 
| 523 | "sbintrunc (numeral k) (numeral (Num.Bit1 w)) = sbintrunc (pred_numeral k) (numeral w) BIT True" | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 524 | "sbintrunc (numeral k) (- numeral (Num.Bit0 w)) = | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 525 | sbintrunc (pred_numeral k) (- numeral w) BIT False" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 526 | "sbintrunc (numeral k) (- numeral (Num.Bit1 w)) = | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 527 | sbintrunc (pred_numeral k) (- numeral (w + Num.One)) BIT True" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 528 | "sbintrunc (numeral k) 1 = 1" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 529 | by (simp_all add: sbintrunc_numeral) | 
| 24333 | 530 | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 531 | lemma no_bintr_alt1: "bintrunc n = (\<lambda>w. w mod 2 ^ n :: int)" | 
| 24333 | 532 | by (rule ext) (rule bintrunc_mod2p) | 
| 533 | ||
| 65363 | 534 | lemma range_bintrunc: "range (bintrunc n) = {i. 0 \<le> i \<and> i < 2 ^ n}"
 | 
| 24333 | 535 | apply (unfold no_bintr_alt1) | 
| 536 | apply (auto simp add: image_iff) | |
| 537 | apply (rule exI) | |
| 538 | apply (auto intro: int_mod_lem [THEN iffD1, symmetric]) | |
| 539 | done | |
| 540 | ||
| 65363 | 541 | lemma no_sbintr_alt2: "sbintrunc n = (\<lambda>w. (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n :: int)" | 
| 24333 | 542 | by (rule ext) (simp add : sbintrunc_mod2p) | 
| 543 | ||
| 65363 | 544 | lemma range_sbintrunc: "range (sbintrunc n) = {i. - (2 ^ n) \<le> i \<and> i < 2 ^ n}"
 | 
| 24333 | 545 | apply (unfold no_sbintr_alt2) | 
| 546 | apply (auto simp add: image_iff eq_diff_eq) | |
| 547 | apply (rule exI) | |
| 548 | apply (auto intro: int_mod_lem [THEN iffD1, symmetric]) | |
| 549 | done | |
| 550 | ||
| 65363 | 551 | lemma sb_inc_lem: "a + 2^k < 0 \<Longrightarrow> a + 2^k + 2^(Suc k) \<le> (a + 2^k) mod 2^(Suc k)" | 
| 552 | for a :: int | |
| 25349 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 553 | 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 | 554 | apply (rule TrueI) | 
| 
0d46bea01741
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 555 | done | 
| 24333 | 556 | |
| 65363 | 557 | lemma sb_inc_lem': "a < - (2^k) \<Longrightarrow> a + 2^k + 2^(Suc k) \<le> (a + 2^k) mod 2^(Suc k)" | 
| 558 | for a :: int | |
| 35048 | 559 | by (rule sb_inc_lem) simp | 
| 24333 | 560 | |
| 65363 | 561 | lemma sbintrunc_inc: "x < - (2^n) \<Longrightarrow> x + 2^(Suc n) \<le> sbintrunc n x" | 
| 24333 | 562 | unfolding no_sbintr_alt2 by (drule sb_inc_lem') simp | 
| 563 | ||
| 65363 | 564 | lemma sb_dec_lem: "0 \<le> - (2 ^ k) + a \<Longrightarrow> (a + 2 ^ k) mod (2 * 2 ^ k) \<le> - (2 ^ k) + a" | 
| 565 | for a :: int | |
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53438diff
changeset | 566 | using int_mod_le'[where n = "2 ^ (Suc k)" and b = "a + 2 ^ k"] by simp | 
| 24333 | 567 | |
| 65363 | 568 | lemma sb_dec_lem': "2 ^ k \<le> a \<Longrightarrow> (a + 2 ^ k) mod (2 * 2 ^ k) \<le> - (2 ^ k) + a" | 
| 569 | for a :: int | |
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53438diff
changeset | 570 | by (rule sb_dec_lem) simp | 
| 24333 | 571 | |
| 65363 | 572 | lemma sbintrunc_dec: "x \<ge> (2 ^ n) \<Longrightarrow> x - 2 ^ (Suc n) >= sbintrunc n x" | 
| 24333 | 573 | unfolding no_sbintr_alt2 by (drule sb_dec_lem') simp | 
| 574 | ||
| 45604 | 575 | lemmas m2pths = pos_mod_sign pos_mod_bound [OF zless2p] | 
| 24364 | 576 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 577 | lemma bintr_ge0: "0 \<le> bintrunc n w" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 578 | by (simp add: bintrunc_mod2p) | 
| 24333 | 579 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 580 | lemma bintr_lt2p: "bintrunc n w < 2 ^ n" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 581 | by (simp add: bintrunc_mod2p) | 
| 24333 | 582 | |
| 58410 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
54873diff
changeset | 583 | lemma bintr_Min: "bintrunc n (- 1) = 2 ^ n - 1" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 584 | by (simp add: bintrunc_mod2p m1mod2k) | 
| 24333 | 585 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 586 | lemma sbintr_ge: "- (2 ^ n) \<le> sbintrunc n w" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 587 | by (simp add: sbintrunc_mod2p) | 
| 24333 | 588 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 589 | lemma sbintr_lt: "sbintrunc n w < 2 ^ n" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 590 | by (simp add: sbintrunc_mod2p) | 
| 24333 | 591 | |
| 65363 | 592 | lemma sign_Pls_ge_0: "bin_sign bin = 0 \<longleftrightarrow> bin \<ge> 0" | 
| 593 | for bin :: int | |
| 594 | by (simp add: bin_sign_def) | |
| 24333 | 595 | |
| 65363 | 596 | lemma sign_Min_lt_0: "bin_sign bin = -1 \<longleftrightarrow> bin < 0" | 
| 597 | for bin :: int | |
| 598 | by (simp add: bin_sign_def) | |
| 24333 | 599 | |
| 65363 | 600 | lemma bin_rest_trunc: "bin_rest (bintrunc n bin) = bintrunc (n - 1) (bin_rest bin)" | 
| 45954 
f67d3bb5f09c
use 'induct arbitrary' instead of universal quantifiers
 huffman parents: 
45953diff
changeset | 601 | by (induct n arbitrary: bin) auto | 
| 24333 | 602 | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 603 | lemma bin_rest_power_trunc: | 
| 65363 | 604 | "(bin_rest ^^ k) (bintrunc n bin) = bintrunc (n - k) ((bin_rest ^^ k) bin)" | 
| 24333 | 605 | by (induct k) (auto simp: bin_rest_trunc) | 
| 606 | ||
| 65363 | 607 | lemma bin_rest_trunc_i: "bintrunc n (bin_rest bin) = bin_rest (bintrunc (Suc n) bin)" | 
| 24333 | 608 | by auto | 
| 609 | ||
| 65363 | 610 | lemma bin_rest_strunc: "bin_rest (sbintrunc (Suc n) bin) = sbintrunc n (bin_rest bin)" | 
| 45954 
f67d3bb5f09c
use 'induct arbitrary' instead of universal quantifiers
 huffman parents: 
45953diff
changeset | 611 | by (induct n arbitrary: bin) auto | 
| 24333 | 612 | |
| 65363 | 613 | lemma bintrunc_rest [simp]: "bintrunc n (bin_rest (bintrunc n bin)) = bin_rest (bintrunc n bin)" | 
| 614 | apply (induct n arbitrary: bin) | |
| 615 | apply simp | |
| 24333 | 616 | apply (case_tac bin rule: bin_exhaust) | 
| 617 | apply (auto simp: bintrunc_bintrunc_l) | |
| 618 | done | |
| 619 | ||
| 65363 | 620 | lemma sbintrunc_rest [simp]: "sbintrunc n (bin_rest (sbintrunc n bin)) = bin_rest (sbintrunc n bin)" | 
| 621 | apply (induct n arbitrary: bin) | |
| 622 | apply simp | |
| 24333 | 623 | apply (case_tac bin rule: bin_exhaust) | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54489diff
changeset | 624 | apply (auto simp: bintrunc_bintrunc_l split: bool.splits) | 
| 24333 | 625 | done | 
| 626 | ||
| 65363 | 627 | lemma bintrunc_rest': "bintrunc n \<circ> bin_rest \<circ> bintrunc n = bin_rest \<circ> bintrunc n" | 
| 24333 | 628 | by (rule ext) auto | 
| 629 | ||
| 65363 | 630 | lemma sbintrunc_rest': "sbintrunc n \<circ> bin_rest \<circ> sbintrunc n = bin_rest \<circ> sbintrunc n" | 
| 24333 | 631 | by (rule ext) auto | 
| 632 | ||
| 65363 | 633 | lemma rco_lem: "f \<circ> g \<circ> f = g \<circ> f \<Longrightarrow> f \<circ> (g \<circ> f) ^^ n = g ^^ n \<circ> f" | 
| 24333 | 634 | apply (rule ext) | 
| 635 | apply (induct_tac n) | |
| 636 | apply (simp_all (no_asm)) | |
| 637 | apply (drule fun_cong) | |
| 638 | apply (unfold o_def) | |
| 639 | apply (erule trans) | |
| 640 | apply simp | |
| 641 | done | |
| 642 | ||
| 65363 | 643 | lemmas rco_bintr = bintrunc_rest' | 
| 24333 | 644 | [THEN rco_lem [THEN fun_cong], unfolded o_def] | 
| 65363 | 645 | lemmas rco_sbintr = sbintrunc_rest' | 
| 24333 | 646 | [THEN rco_lem [THEN fun_cong], unfolded o_def] | 
| 647 | ||
| 65363 | 648 | |
| 61799 | 649 | subsection \<open>Splitting and concatenation\<close> | 
| 24364 | 650 | |
| 65363 | 651 | primrec bin_split :: "nat \<Rightarrow> int \<Rightarrow> int \<times> int" | 
| 652 | where | |
| 653 | Z: "bin_split 0 w = (w, 0)" | |
| 654 | | Suc: "bin_split (Suc n) w = | |
| 655 | (let (w1, w2) = bin_split n (bin_rest w) | |
| 656 | in (w1, w2 BIT bin_last w))" | |
| 24364 | 657 | |
| 37667 | 658 | lemma [code]: | 
| 659 | "bin_split (Suc n) w = (let (w1, w2) = bin_split n (bin_rest w) in (w1, w2 BIT bin_last w))" | |
| 660 | "bin_split 0 w = (w, 0)" | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46607diff
changeset | 661 | by simp_all | 
| 37667 | 662 | |
| 65363 | 663 | primrec bin_cat :: "int \<Rightarrow> nat \<Rightarrow> int \<Rightarrow> int" | 
| 664 | where | |
| 665 | Z: "bin_cat w 0 v = w" | |
| 26557 | 666 | | Suc: "bin_cat w (Suc n) v = bin_cat w n (bin_rest v) BIT bin_last v" | 
| 24364 | 667 | |
| 53062 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 haftmann parents: 
47219diff
changeset | 668 | end |