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