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