| author | wenzelm |
| Sat, 03 Dec 2011 13:11:50 +0100 | |
| changeset 45744 | 0ad063afa3d6 |
| parent 45604 | 29cf40fe8daf |
| child 45845 | 4158f35a5c6f |
| permissions | -rw-r--r-- |
| 24333 | 1 |
(* |
2 |
Author: Jeremy Dawson and Gerwin Klein, NICTA |
|
3 |
||
|
44939
5930d35c976d
removed unused legacy lemma names, some comment cleanup.
kleing
parents:
37667
diff
changeset
|
4 |
Definitions and basic theorems for bit-wise logical operations |
| 24333 | 5 |
for integers expressed using Pls, Min, BIT, |
|
44939
5930d35c976d
removed unused legacy lemma names, some comment cleanup.
kleing
parents:
37667
diff
changeset
|
6 |
and converting them to and from lists of bools. |
| 24333 | 7 |
*) |
8 |
||
| 24350 | 9 |
header {* Bitwise Operations on Binary Integers *}
|
10 |
||
| 37658 | 11 |
theory Bit_Int |
12 |
imports Bit_Representation Bit_Operations |
|
| 24333 | 13 |
begin |
14 |
||
| 37667 | 15 |
subsection {* Recursion combinators for bitstrings *}
|
16 |
||
17 |
function bin_rec :: "'a \<Rightarrow> 'a \<Rightarrow> (int \<Rightarrow> bit \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> int \<Rightarrow> 'a" where |
|
18 |
"bin_rec f1 f2 f3 bin = (if bin = 0 then f1 |
|
19 |
else if bin = - 1 then f2 |
|
20 |
else f3 (bin_rest bin) (bin_last bin) (bin_rec f1 f2 f3 (bin_rest bin)))" |
|
21 |
by pat_completeness auto |
|
22 |
||
23 |
termination by (relation "measure (nat o abs o snd o snd o snd)") |
|
24 |
(simp_all add: bin_last_def bin_rest_def) |
|
25 |
||
26 |
declare bin_rec.simps [simp del] |
|
27 |
||
28 |
lemma bin_rec_PM: |
|
29 |
"f = bin_rec f1 f2 f3 ==> f Int.Pls = f1 & f Int.Min = f2" |
|
30 |
by (unfold Pls_def Min_def) (simp add: bin_rec.simps) |
|
31 |
||
32 |
lemma bin_rec_Pls: "bin_rec f1 f2 f3 Int.Pls = f1" |
|
33 |
by (unfold Pls_def Min_def) (simp add: bin_rec.simps) |
|
34 |
||
35 |
lemma bin_rec_Min: "bin_rec f1 f2 f3 Int.Min = f2" |
|
36 |
by (unfold Pls_def Min_def) (simp add: bin_rec.simps) |
|
37 |
||
38 |
lemma bin_rec_Bit0: |
|
39 |
"f3 Int.Pls (0::bit) f1 = f1 \<Longrightarrow> |
|
40 |
bin_rec f1 f2 f3 (Int.Bit0 w) = f3 w (0::bit) (bin_rec f1 f2 f3 w)" |
|
41 |
by (unfold Pls_def Min_def Bit0_def Bit1_def) (simp add: bin_rec.simps bin_last_def bin_rest_def) |
|
42 |
||
43 |
lemma bin_rec_Bit1: |
|
44 |
"f3 Int.Min (1::bit) f2 = f2 \<Longrightarrow> |
|
45 |
bin_rec f1 f2 f3 (Int.Bit1 w) = f3 w (1::bit) (bin_rec f1 f2 f3 w)" |
|
46 |
apply (cases "w = Int.Min") |
|
47 |
apply (simp add: bin_rec_Min) |
|
48 |
apply (cases "w = Int.Pls") |
|
49 |
apply (simp add: bin_rec_Pls number_of_is_id Pls_def [symmetric] bin_rec.simps) |
|
50 |
apply (subst bin_rec.simps) |
|
51 |
apply auto unfolding Pls_def Min_def Bit0_def Bit1_def number_of_is_id apply auto |
|
52 |
done |
|
53 |
||
54 |
lemma bin_rec_Bit: |
|
55 |
"f = bin_rec f1 f2 f3 ==> f3 Int.Pls (0::bit) f1 = f1 ==> |
|
56 |
f3 Int.Min (1::bit) f2 = f2 ==> f (w BIT b) = f3 w b (f w)" |
|
57 |
by (cases b, simp add: bin_rec_Bit0, simp add: bin_rec_Bit1) |
|
58 |
||
59 |
lemmas bin_rec_simps = refl [THEN bin_rec_Bit] bin_rec_Pls bin_rec_Min |
|
60 |
bin_rec_Bit0 bin_rec_Bit1 |
|
61 |
||
62 |
||
| 24364 | 63 |
subsection {* Logical operations *}
|
| 24353 | 64 |
|
65 |
text "bit-wise logical operations on the int type" |
|
66 |
||
| 25762 | 67 |
instantiation int :: bit |
68 |
begin |
|
69 |
||
70 |
definition |
|
| 37667 | 71 |
int_not_def: "bitNOT = bin_rec (- 1) 0 |
| 24353 | 72 |
(\<lambda>w b s. s BIT (NOT b))" |
| 25762 | 73 |
|
74 |
definition |
|
| 37667 | 75 |
int_and_def: "bitAND = bin_rec (\<lambda>x. 0) (\<lambda>y. y) |
| 24353 | 76 |
(\<lambda>w b s y. s (bin_rest y) BIT (b AND bin_last y))" |
| 25762 | 77 |
|
78 |
definition |
|
| 37667 | 79 |
int_or_def: "bitOR = bin_rec (\<lambda>x. x) (\<lambda>y. - 1) |
| 24353 | 80 |
(\<lambda>w b s y. s (bin_rest y) BIT (b OR bin_last y))" |
| 25762 | 81 |
|
82 |
definition |
|
| 37667 | 83 |
int_xor_def: "bitXOR = bin_rec (\<lambda>x. x) bitNOT |
| 24353 | 84 |
(\<lambda>w b s y. s (bin_rest y) BIT (b XOR bin_last y))" |
| 25762 | 85 |
|
86 |
instance .. |
|
87 |
||
88 |
end |
|
| 24353 | 89 |
|
|
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
90 |
subsubsection {* Basic simplification rules *}
|
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
91 |
|
| 24333 | 92 |
lemma int_not_simps [simp]: |
|
25919
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents:
25762
diff
changeset
|
93 |
"NOT Int.Pls = Int.Min" |
|
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents:
25762
diff
changeset
|
94 |
"NOT Int.Min = Int.Pls" |
|
26086
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
95 |
"NOT (Int.Bit0 w) = Int.Bit1 (NOT w)" |
|
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
96 |
"NOT (Int.Bit1 w) = Int.Bit0 (NOT w)" |
| 26558 | 97 |
"NOT (w BIT b) = (NOT w) BIT (NOT b)" |
| 37667 | 98 |
unfolding int_not_def Pls_def [symmetric] Min_def [symmetric] by (simp_all add: bin_rec_simps) |
| 24333 | 99 |
|
| 37667 | 100 |
lemma int_xor_Pls [simp]: |
|
25919
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents:
25762
diff
changeset
|
101 |
"Int.Pls XOR x = x" |
| 37667 | 102 |
unfolding int_xor_def Pls_def [symmetric] Min_def [symmetric] by (simp add: bin_rec_PM) |
| 24333 | 103 |
|
| 37667 | 104 |
lemma int_xor_Min [simp]: |
|
25919
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents:
25762
diff
changeset
|
105 |
"Int.Min XOR x = NOT x" |
| 37667 | 106 |
unfolding int_xor_def Pls_def [symmetric] Min_def [symmetric] by (simp add: bin_rec_PM) |
| 24333 | 107 |
|
108 |
lemma int_xor_Bits [simp]: |
|
| 24353 | 109 |
"(x BIT b) XOR (y BIT c) = (x XOR y) BIT (b XOR c)" |
| 37667 | 110 |
apply (unfold int_xor_def Pls_def [symmetric] Min_def [symmetric]) |
| 24333 | 111 |
apply (rule bin_rec_simps (1) [THEN fun_cong, THEN trans]) |
112 |
apply (rule ext, simp) |
|
113 |
prefer 2 |
|
114 |
apply simp |
|
115 |
apply (rule ext) |
|
116 |
apply (simp add: int_not_simps [symmetric]) |
|
117 |
done |
|
118 |
||
| 37667 | 119 |
lemma int_xor_Bits2 [simp]: |
|
26086
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
120 |
"(Int.Bit0 x) XOR (Int.Bit0 y) = Int.Bit0 (x XOR y)" |
|
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
121 |
"(Int.Bit0 x) XOR (Int.Bit1 y) = Int.Bit1 (x XOR y)" |
|
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
122 |
"(Int.Bit1 x) XOR (Int.Bit0 y) = Int.Bit1 (x XOR y)" |
|
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
123 |
"(Int.Bit1 x) XOR (Int.Bit1 y) = Int.Bit0 (x XOR y)" |
|
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
124 |
unfolding BIT_simps [symmetric] int_xor_Bits by simp_all |
|
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
125 |
|
| 37667 | 126 |
lemma int_or_Pls [simp]: |
|
25919
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents:
25762
diff
changeset
|
127 |
"Int.Pls OR x = x" |
| 24465 | 128 |
by (unfold int_or_def) (simp add: bin_rec_PM) |
| 24333 | 129 |
|
| 37667 | 130 |
lemma int_or_Min [simp]: |
|
25919
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents:
25762
diff
changeset
|
131 |
"Int.Min OR x = Int.Min" |
| 37667 | 132 |
by (unfold int_or_def Pls_def [symmetric] Min_def [symmetric]) (simp add: bin_rec_PM) |
| 24333 | 133 |
|
134 |
lemma int_or_Bits [simp]: |
|
| 24353 | 135 |
"(x BIT b) OR (y BIT c) = (x OR y) BIT (b OR c)" |
| 37667 | 136 |
unfolding int_or_def Pls_def [symmetric] Min_def [symmetric] by (simp add: bin_rec_simps) |
| 24333 | 137 |
|
| 37667 | 138 |
lemma int_or_Bits2 [simp]: |
|
26086
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
139 |
"(Int.Bit0 x) OR (Int.Bit0 y) = Int.Bit0 (x OR y)" |
|
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
140 |
"(Int.Bit0 x) OR (Int.Bit1 y) = Int.Bit1 (x OR y)" |
|
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
141 |
"(Int.Bit1 x) OR (Int.Bit0 y) = Int.Bit1 (x OR y)" |
|
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
142 |
"(Int.Bit1 x) OR (Int.Bit1 y) = Int.Bit1 (x OR y)" |
|
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
143 |
unfolding BIT_simps [symmetric] int_or_Bits by simp_all |
|
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
144 |
|
| 37667 | 145 |
lemma int_and_Pls [simp]: |
|
25919
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents:
25762
diff
changeset
|
146 |
"Int.Pls AND x = Int.Pls" |
| 24465 | 147 |
unfolding int_and_def by (simp add: bin_rec_PM) |
| 24333 | 148 |
|
| 37667 | 149 |
lemma int_and_Min [simp]: |
|
25919
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents:
25762
diff
changeset
|
150 |
"Int.Min AND x = x" |
| 24465 | 151 |
unfolding int_and_def by (simp add: bin_rec_PM) |
| 24333 | 152 |
|
153 |
lemma int_and_Bits [simp]: |
|
| 24353 | 154 |
"(x BIT b) AND (y BIT c) = (x AND y) BIT (b AND c)" |
| 37667 | 155 |
unfolding int_and_def Pls_def [symmetric] Min_def [symmetric] by (simp add: bin_rec_simps) |
| 24333 | 156 |
|
| 37667 | 157 |
lemma int_and_Bits2 [simp]: |
|
26086
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
158 |
"(Int.Bit0 x) AND (Int.Bit0 y) = Int.Bit0 (x AND y)" |
|
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
159 |
"(Int.Bit0 x) AND (Int.Bit1 y) = Int.Bit0 (x AND y)" |
|
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
160 |
"(Int.Bit1 x) AND (Int.Bit0 y) = Int.Bit0 (x AND y)" |
|
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
161 |
"(Int.Bit1 x) AND (Int.Bit1 y) = Int.Bit1 (x AND y)" |
|
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
162 |
unfolding BIT_simps [symmetric] int_and_Bits by simp_all |
|
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
163 |
|
|
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
164 |
subsubsection {* Binary destructors *}
|
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
165 |
|
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
166 |
lemma bin_rest_NOT [simp]: "bin_rest (NOT x) = NOT (bin_rest x)" |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
167 |
by (cases x rule: bin_exhaust, simp) |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
168 |
|
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
169 |
lemma bin_last_NOT [simp]: "bin_last (NOT x) = NOT (bin_last x)" |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
170 |
by (cases x rule: bin_exhaust, simp) |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
171 |
|
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
172 |
lemma bin_rest_AND [simp]: "bin_rest (x AND y) = bin_rest x AND bin_rest y" |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
173 |
by (cases x rule: bin_exhaust, cases y rule: bin_exhaust, simp) |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
174 |
|
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
175 |
lemma bin_last_AND [simp]: "bin_last (x AND y) = bin_last x AND bin_last y" |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
176 |
by (cases x rule: bin_exhaust, cases y rule: bin_exhaust, simp) |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
177 |
|
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
178 |
lemma bin_rest_OR [simp]: "bin_rest (x OR y) = bin_rest x OR bin_rest y" |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
179 |
by (cases x rule: bin_exhaust, cases y rule: bin_exhaust, simp) |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
180 |
|
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
181 |
lemma bin_last_OR [simp]: "bin_last (x OR y) = bin_last x OR bin_last y" |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
182 |
by (cases x rule: bin_exhaust, cases y rule: bin_exhaust, simp) |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
183 |
|
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
184 |
lemma bin_rest_XOR [simp]: "bin_rest (x XOR y) = bin_rest x XOR bin_rest y" |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
185 |
by (cases x rule: bin_exhaust, cases y rule: bin_exhaust, simp) |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
186 |
|
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
187 |
lemma bin_last_XOR [simp]: "bin_last (x XOR y) = bin_last x XOR bin_last y" |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
188 |
by (cases x rule: bin_exhaust, cases y rule: bin_exhaust, simp) |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
189 |
|
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
190 |
lemma bit_NOT_eq_1_iff [simp]: "NOT (b::bit) = 1 \<longleftrightarrow> b = 0" |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
191 |
by (induct b, simp_all) |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
192 |
|
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
193 |
lemma bit_AND_eq_1_iff [simp]: "(a::bit) AND b = 1 \<longleftrightarrow> a = 1 \<and> b = 1" |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
194 |
by (induct a, simp_all) |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
195 |
|
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
196 |
lemma bin_nth_ops: |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
197 |
"!!x y. bin_nth (x AND y) n = (bin_nth x n & bin_nth y n)" |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
198 |
"!!x y. bin_nth (x OR y) n = (bin_nth x n | bin_nth y n)" |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
199 |
"!!x y. bin_nth (x XOR y) n = (bin_nth x n ~= bin_nth y n)" |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
200 |
"!!x. bin_nth (NOT x) n = (~ bin_nth x n)" |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
201 |
by (induct n) auto |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
202 |
|
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
203 |
subsubsection {* Derived properties *}
|
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
204 |
|
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
205 |
lemma int_xor_extra_simps [simp]: |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
206 |
"w XOR Int.Pls = w" |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
207 |
"w XOR Int.Min = NOT w" |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
208 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
209 |
|
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
210 |
lemma int_or_extra_simps [simp]: |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
211 |
"w OR Int.Pls = w" |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
212 |
"w OR Int.Min = Int.Min" |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
213 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
| 24333 | 214 |
|
| 37667 | 215 |
lemma int_and_extra_simps [simp]: |
|
26086
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
216 |
"w AND Int.Pls = Int.Pls" |
|
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
217 |
"w AND Int.Min = w" |
|
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
218 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
| 24333 | 219 |
|
220 |
(* commutativity of the above *) |
|
221 |
lemma bin_ops_comm: |
|
222 |
shows |
|
| 24353 | 223 |
int_and_comm: "!!y::int. x AND y = y AND x" and |
224 |
int_or_comm: "!!y::int. x OR y = y OR x" and |
|
225 |
int_xor_comm: "!!y::int. x XOR y = y XOR x" |
|
|
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
226 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
| 24333 | 227 |
|
228 |
lemma bin_ops_same [simp]: |
|
| 24353 | 229 |
"(x::int) AND x = x" |
230 |
"(x::int) OR x = x" |
|
|
25919
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents:
25762
diff
changeset
|
231 |
"(x::int) XOR x = Int.Pls" |
|
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
232 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
| 24333 | 233 |
|
| 24353 | 234 |
lemma int_not_not [simp]: "NOT (NOT (x::int)) = x" |
|
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
235 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
| 24333 | 236 |
|
237 |
lemmas bin_log_esimps = |
|
238 |
int_and_extra_simps int_or_extra_simps int_xor_extra_simps |
|
239 |
int_and_Pls int_and_Min int_or_Pls int_or_Min int_xor_Pls int_xor_Min |
|
240 |
||
241 |
(* basic properties of logical (bit-wise) operations *) |
|
242 |
||
243 |
lemma bbw_ao_absorb: |
|
| 24353 | 244 |
"!!y::int. x AND (y OR x) = x & x OR (y AND x) = x" |
|
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
245 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
| 24333 | 246 |
|
247 |
lemma bbw_ao_absorbs_other: |
|
| 24353 | 248 |
"x AND (x OR y) = x \<and> (y AND x) OR x = (x::int)" |
249 |
"(y OR x) AND x = x \<and> x OR (x AND y) = (x::int)" |
|
250 |
"(x OR y) AND x = x \<and> (x AND y) OR x = (x::int)" |
|
|
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
251 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
| 24353 | 252 |
|
| 24333 | 253 |
lemmas bbw_ao_absorbs [simp] = bbw_ao_absorb bbw_ao_absorbs_other |
254 |
||
255 |
lemma int_xor_not: |
|
| 24353 | 256 |
"!!y::int. (NOT x) XOR y = NOT (x XOR y) & |
257 |
x XOR (NOT y) = NOT (x XOR y)" |
|
|
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
258 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
| 24333 | 259 |
|
260 |
lemma int_and_assoc: |
|
| 24353 | 261 |
"(x AND y) AND (z::int) = x AND (y AND z)" |
|
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
262 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
| 24333 | 263 |
|
264 |
lemma int_or_assoc: |
|
| 24353 | 265 |
"(x OR y) OR (z::int) = x OR (y OR z)" |
|
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
266 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
| 24333 | 267 |
|
268 |
lemma int_xor_assoc: |
|
| 24353 | 269 |
"(x XOR y) XOR (z::int) = x XOR (y XOR z)" |
|
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
270 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
| 24333 | 271 |
|
272 |
lemmas bbw_assocs = int_and_assoc int_or_assoc int_xor_assoc |
|
273 |
||
|
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
274 |
(* BH: Why are these declared as simp rules??? *) |
| 24333 | 275 |
lemma bbw_lcs [simp]: |
| 24353 | 276 |
"(y::int) AND (x AND z) = x AND (y AND z)" |
277 |
"(y::int) OR (x OR z) = x OR (y OR z)" |
|
278 |
"(y::int) XOR (x XOR z) = x XOR (y XOR z)" |
|
|
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
279 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
| 24333 | 280 |
|
281 |
lemma bbw_not_dist: |
|
| 24353 | 282 |
"!!y::int. NOT (x OR y) = (NOT x) AND (NOT y)" |
283 |
"!!y::int. NOT (x AND y) = (NOT x) OR (NOT y)" |
|
|
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
284 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
| 24333 | 285 |
|
286 |
lemma bbw_oa_dist: |
|
| 24353 | 287 |
"!!y z::int. (x AND y) OR z = |
288 |
(x OR z) AND (y OR z)" |
|
|
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
289 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
| 24333 | 290 |
|
291 |
lemma bbw_ao_dist: |
|
| 24353 | 292 |
"!!y z::int. (x OR y) AND z = |
293 |
(x AND z) OR (y AND z)" |
|
|
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
294 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
| 24333 | 295 |
|
|
24367
3e29eafabe16
AC rules for bitwise logical operators no longer declared simp
huffman
parents:
24366
diff
changeset
|
296 |
(* |
|
3e29eafabe16
AC rules for bitwise logical operators no longer declared simp
huffman
parents:
24366
diff
changeset
|
297 |
Why were these declared simp??? |
| 24333 | 298 |
declare bin_ops_comm [simp] bbw_assocs [simp] |
|
24367
3e29eafabe16
AC rules for bitwise logical operators no longer declared simp
huffman
parents:
24366
diff
changeset
|
299 |
*) |
| 24333 | 300 |
|
|
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
301 |
subsubsection {* Interactions with arithmetic *}
|
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
302 |
|
| 24333 | 303 |
lemma plus_and_or [rule_format]: |
| 24353 | 304 |
"ALL y::int. (x AND y) + (x OR y) = x + y" |
| 24333 | 305 |
apply (induct x rule: bin_induct) |
306 |
apply clarsimp |
|
307 |
apply clarsimp |
|
308 |
apply clarsimp |
|
309 |
apply (case_tac y rule: bin_exhaust) |
|
310 |
apply clarsimp |
|
311 |
apply (unfold Bit_def) |
|
312 |
apply clarsimp |
|
313 |
apply (erule_tac x = "x" in allE) |
|
| 37667 | 314 |
apply (simp add: bitval_def split: bit.split) |
| 24333 | 315 |
done |
316 |
||
317 |
lemma le_int_or: |
|
| 37667 | 318 |
"bin_sign (y::int) = Int.Pls ==> x <= x OR y" |
319 |
apply (induct y arbitrary: x rule: bin_induct) |
|
| 24333 | 320 |
apply clarsimp |
321 |
apply clarsimp |
|
322 |
apply (case_tac x rule: bin_exhaust) |
|
323 |
apply (case_tac b) |
|
324 |
apply (case_tac [!] bit) |
|
| 26514 | 325 |
apply (auto simp: less_eq_int_code) |
| 24333 | 326 |
done |
327 |
||
328 |
lemmas int_and_le = |
|
| 45475 | 329 |
xtr3 [OF bbw_ao_absorbs (2) [THEN conjunct2, symmetric] le_int_or] |
| 24333 | 330 |
|
| 24364 | 331 |
(* interaction between bit-wise and arithmetic *) |
332 |
(* good example of bin_induction *) |
|
|
25919
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents:
25762
diff
changeset
|
333 |
lemma bin_add_not: "x + NOT x = Int.Min" |
| 24364 | 334 |
apply (induct x rule: bin_induct) |
335 |
apply clarsimp |
|
336 |
apply clarsimp |
|
337 |
apply (case_tac bit, auto) |
|
338 |
done |
|
339 |
||
|
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
340 |
subsubsection {* Truncating results of bit-wise operations *}
|
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
341 |
|
| 24364 | 342 |
lemma bin_trunc_ao: |
343 |
"!!x y. (bintrunc n x) AND (bintrunc n y) = bintrunc n (x AND y)" |
|
344 |
"!!x y. (bintrunc n x) OR (bintrunc n y) = bintrunc n (x OR y)" |
|
|
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
345 |
by (auto simp add: bin_eq_iff bin_nth_ops nth_bintr) |
| 24364 | 346 |
|
347 |
lemma bin_trunc_xor: |
|
348 |
"!!x y. bintrunc n (bintrunc n x XOR bintrunc n y) = |
|
349 |
bintrunc n (x XOR y)" |
|
|
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
350 |
by (auto simp add: bin_eq_iff bin_nth_ops nth_bintr) |
| 24364 | 351 |
|
352 |
lemma bin_trunc_not: |
|
353 |
"!!x. bintrunc n (NOT (bintrunc n x)) = bintrunc n (NOT x)" |
|
|
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
354 |
by (auto simp add: bin_eq_iff bin_nth_ops nth_bintr) |
| 24364 | 355 |
|
356 |
(* want theorems of the form of bin_trunc_xor *) |
|
357 |
lemma bintr_bintr_i: |
|
358 |
"x = bintrunc n y ==> bintrunc n x = bintrunc n y" |
|
359 |
by auto |
|
360 |
||
361 |
lemmas bin_trunc_and = bin_trunc_ao(1) [THEN bintr_bintr_i] |
|
362 |
lemmas bin_trunc_or = bin_trunc_ao(2) [THEN bintr_bintr_i] |
|
363 |
||
364 |
subsection {* Setting and clearing bits *}
|
|
365 |
||
| 26558 | 366 |
primrec |
| 24364 | 367 |
bin_sc :: "nat => bit => int => int" |
| 26558 | 368 |
where |
369 |
Z: "bin_sc 0 b w = bin_rest w BIT b" |
|
370 |
| Suc: "bin_sc (Suc n) b w = bin_sc n b (bin_rest w) BIT bin_last w" |
|
| 24364 | 371 |
|
| 24333 | 372 |
(** nth bit, set/clear **) |
373 |
||
374 |
lemma bin_nth_sc [simp]: |
|
| 37654 | 375 |
"!!w. bin_nth (bin_sc n b w) n = (b = 1)" |
| 24333 | 376 |
by (induct n) auto |
377 |
||
378 |
lemma bin_sc_sc_same [simp]: |
|
379 |
"!!w. bin_sc n c (bin_sc n b w) = bin_sc n c w" |
|
380 |
by (induct n) auto |
|
381 |
||
382 |
lemma bin_sc_sc_diff: |
|
383 |
"!!w m. m ~= n ==> |
|
384 |
bin_sc m c (bin_sc n b w) = bin_sc n b (bin_sc m c w)" |
|
385 |
apply (induct n) |
|
386 |
apply (case_tac [!] m) |
|
387 |
apply auto |
|
388 |
done |
|
389 |
||
390 |
lemma bin_nth_sc_gen: |
|
| 37654 | 391 |
"!!w m. bin_nth (bin_sc n b w) m = (if m = n then b = 1 else bin_nth w m)" |
| 24333 | 392 |
by (induct n) (case_tac [!] m, auto) |
393 |
||
394 |
lemma bin_sc_nth [simp]: |
|
| 37654 | 395 |
"!!w. (bin_sc n (If (bin_nth w n) 1 0) w) = w" |
| 24465 | 396 |
by (induct n) auto |
| 24333 | 397 |
|
398 |
lemma bin_sign_sc [simp]: |
|
399 |
"!!w. bin_sign (bin_sc n b w) = bin_sign w" |
|
400 |
by (induct n) auto |
|
401 |
||
402 |
lemma bin_sc_bintr [simp]: |
|
403 |
"!!w m. bintrunc m (bin_sc n x (bintrunc m (w))) = bintrunc m (bin_sc n x w)" |
|
404 |
apply (induct n) |
|
405 |
apply (case_tac [!] w rule: bin_exhaust) |
|
406 |
apply (case_tac [!] m, auto) |
|
407 |
done |
|
408 |
||
409 |
lemma bin_clr_le: |
|
| 37654 | 410 |
"!!w. bin_sc n 0 w <= w" |
| 24333 | 411 |
apply (induct n) |
412 |
apply (case_tac [!] w rule: bin_exhaust) |
|
|
26086
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
413 |
apply (auto simp del: BIT_simps) |
| 24333 | 414 |
apply (unfold Bit_def) |
| 37667 | 415 |
apply (simp_all add: bitval_def split: bit.split) |
| 24333 | 416 |
done |
417 |
||
418 |
lemma bin_set_ge: |
|
| 37654 | 419 |
"!!w. bin_sc n 1 w >= w" |
| 24333 | 420 |
apply (induct n) |
421 |
apply (case_tac [!] w rule: bin_exhaust) |
|
|
26086
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
422 |
apply (auto simp del: BIT_simps) |
| 24333 | 423 |
apply (unfold Bit_def) |
| 37667 | 424 |
apply (simp_all add: bitval_def split: bit.split) |
| 24333 | 425 |
done |
426 |
||
427 |
lemma bintr_bin_clr_le: |
|
| 37654 | 428 |
"!!w m. bintrunc n (bin_sc m 0 w) <= bintrunc n w" |
| 24333 | 429 |
apply (induct n) |
430 |
apply simp |
|
431 |
apply (case_tac w rule: bin_exhaust) |
|
432 |
apply (case_tac m) |
|
|
26086
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
433 |
apply (auto simp del: BIT_simps) |
| 24333 | 434 |
apply (unfold Bit_def) |
| 37667 | 435 |
apply (simp_all add: bitval_def split: bit.split) |
| 24333 | 436 |
done |
437 |
||
438 |
lemma bintr_bin_set_ge: |
|
| 37654 | 439 |
"!!w m. bintrunc n (bin_sc m 1 w) >= bintrunc n w" |
| 24333 | 440 |
apply (induct n) |
441 |
apply simp |
|
442 |
apply (case_tac w rule: bin_exhaust) |
|
443 |
apply (case_tac m) |
|
|
26086
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
444 |
apply (auto simp del: BIT_simps) |
| 24333 | 445 |
apply (unfold Bit_def) |
| 37667 | 446 |
apply (simp_all add: bitval_def split: bit.split) |
| 24333 | 447 |
done |
448 |
||
| 37654 | 449 |
lemma bin_sc_FP [simp]: "bin_sc n 0 Int.Pls = Int.Pls" |
| 24333 | 450 |
by (induct n) auto |
451 |
||
| 37654 | 452 |
lemma bin_sc_TM [simp]: "bin_sc n 1 Int.Min = Int.Min" |
| 24333 | 453 |
by (induct n) auto |
454 |
||
455 |
lemmas bin_sc_simps = bin_sc.Z bin_sc.Suc bin_sc_TM bin_sc_FP |
|
456 |
||
457 |
lemma bin_sc_minus: |
|
458 |
"0 < n ==> bin_sc (Suc (n - 1)) b w = bin_sc n b w" |
|
459 |
by auto |
|
460 |
||
461 |
lemmas bin_sc_Suc_minus = |
|
| 45604 | 462 |
trans [OF bin_sc_minus [symmetric] bin_sc.Suc] |
| 24333 | 463 |
|
464 |
lemmas bin_sc_Suc_pred [simp] = |
|
| 45604 | 465 |
bin_sc_Suc_minus [of "number_of bin", simplified nobm1] for bin |
| 24333 | 466 |
|
| 24465 | 467 |
|
| 24364 | 468 |
subsection {* Splitting and concatenation *}
|
| 24333 | 469 |
|
| 26558 | 470 |
definition bin_rcat :: "nat \<Rightarrow> int list \<Rightarrow> int" where |
471 |
"bin_rcat n = foldl (%u v. bin_cat u n v) Int.Pls" |
|
472 |
||
| 37667 | 473 |
lemma [code]: |
474 |
"bin_rcat n = foldl (\<lambda>u v. bin_cat u n v) 0" |
|
475 |
by (simp add: bin_rcat_def Pls_def) |
|
476 |
||
| 28042 | 477 |
fun bin_rsplit_aux :: "nat \<Rightarrow> nat \<Rightarrow> int \<Rightarrow> int list \<Rightarrow> int list" where |
| 26558 | 478 |
"bin_rsplit_aux n m c bs = |
| 24364 | 479 |
(if m = 0 | n = 0 then bs else |
480 |
let (a, b) = bin_split n c |
|
| 26558 | 481 |
in bin_rsplit_aux n (m - n) a (b # bs))" |
| 24364 | 482 |
|
| 26558 | 483 |
definition bin_rsplit :: "nat \<Rightarrow> nat \<times> int \<Rightarrow> int list" where |
484 |
"bin_rsplit n w = bin_rsplit_aux n (fst w) (snd w) []" |
|
485 |
||
| 28042 | 486 |
fun bin_rsplitl_aux :: "nat \<Rightarrow> nat \<Rightarrow> int \<Rightarrow> int list \<Rightarrow> int list" where |
| 26558 | 487 |
"bin_rsplitl_aux n m c bs = |
| 24364 | 488 |
(if m = 0 | n = 0 then bs else |
489 |
let (a, b) = bin_split (min m n) c |
|
| 26558 | 490 |
in bin_rsplitl_aux n (m - n) a (b # bs))" |
| 24364 | 491 |
|
| 26558 | 492 |
definition bin_rsplitl :: "nat \<Rightarrow> nat \<times> int \<Rightarrow> int list" where |
493 |
"bin_rsplitl n w = bin_rsplitl_aux n (fst w) (snd w) []" |
|
494 |
||
| 24364 | 495 |
declare bin_rsplit_aux.simps [simp del] |
496 |
declare bin_rsplitl_aux.simps [simp del] |
|
497 |
||
498 |
lemma bin_sign_cat: |
|
499 |
"!!y. bin_sign (bin_cat x n y) = bin_sign x" |
|
500 |
by (induct n) auto |
|
501 |
||
502 |
lemma bin_cat_Suc_Bit: |
|
503 |
"bin_cat w (Suc n) (v BIT b) = bin_cat w n v BIT b" |
|
504 |
by auto |
|
505 |
||
506 |
lemma bin_nth_cat: |
|
507 |
"!!n y. bin_nth (bin_cat x k y) n = |
|
508 |
(if n < k then bin_nth y n else bin_nth x (n - k))" |
|
509 |
apply (induct k) |
|
510 |
apply clarsimp |
|
511 |
apply (case_tac n, auto) |
|
| 24333 | 512 |
done |
513 |
||
| 24364 | 514 |
lemma bin_nth_split: |
515 |
"!!b c. bin_split n c = (a, b) ==> |
|
516 |
(ALL k. bin_nth a k = bin_nth c (n + k)) & |
|
517 |
(ALL k. bin_nth b k = (k < n & bin_nth c k))" |
|
| 24333 | 518 |
apply (induct n) |
| 24364 | 519 |
apply clarsimp |
520 |
apply (clarsimp simp: Let_def split: ls_splits) |
|
521 |
apply (case_tac k) |
|
522 |
apply auto |
|
523 |
done |
|
524 |
||
525 |
lemma bin_cat_assoc: |
|
526 |
"!!z. bin_cat (bin_cat x m y) n z = bin_cat x (m + n) (bin_cat y n z)" |
|
527 |
by (induct n) auto |
|
528 |
||
529 |
lemma bin_cat_assoc_sym: "!!z m. |
|
530 |
bin_cat x m (bin_cat y n z) = bin_cat (bin_cat x (m - n) y) (min m n) z" |
|
531 |
apply (induct n, clarsimp) |
|
532 |
apply (case_tac m, auto) |
|
| 24333 | 533 |
done |
534 |
||
| 24364 | 535 |
lemma bin_cat_Pls [simp]: |
|
25919
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents:
25762
diff
changeset
|
536 |
"!!w. bin_cat Int.Pls n w = bintrunc n w" |
| 24364 | 537 |
by (induct n) auto |
538 |
||
539 |
lemma bintr_cat1: |
|
540 |
"!!b. bintrunc (k + n) (bin_cat a n b) = bin_cat (bintrunc k a) n b" |
|
541 |
by (induct n) auto |
|
542 |
||
543 |
lemma bintr_cat: "bintrunc m (bin_cat a n b) = |
|
544 |
bin_cat (bintrunc (m - n) a) n (bintrunc (min m n) b)" |
|
545 |
by (rule bin_eqI) (auto simp: bin_nth_cat nth_bintr) |
|
546 |
||
547 |
lemma bintr_cat_same [simp]: |
|
548 |
"bintrunc n (bin_cat a n b) = bintrunc n b" |
|
549 |
by (auto simp add : bintr_cat) |
|
550 |
||
551 |
lemma cat_bintr [simp]: |
|
552 |
"!!b. bin_cat a n (bintrunc n b) = bin_cat a n b" |
|
553 |
by (induct n) auto |
|
554 |
||
555 |
lemma split_bintrunc: |
|
556 |
"!!b c. bin_split n c = (a, b) ==> b = bintrunc n c" |
|
557 |
by (induct n) (auto simp: Let_def split: ls_splits) |
|
558 |
||
559 |
lemma bin_cat_split: |
|
560 |
"!!v w. bin_split n w = (u, v) ==> w = bin_cat u n v" |
|
561 |
by (induct n) (auto simp: Let_def split: ls_splits) |
|
562 |
||
563 |
lemma bin_split_cat: |
|
564 |
"!!w. bin_split n (bin_cat v n w) = (v, bintrunc n w)" |
|
565 |
by (induct n) auto |
|
566 |
||
567 |
lemma bin_split_Pls [simp]: |
|
|
25919
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents:
25762
diff
changeset
|
568 |
"bin_split n Int.Pls = (Int.Pls, Int.Pls)" |
| 24364 | 569 |
by (induct n) (auto simp: Let_def split: ls_splits) |
570 |
||
571 |
lemma bin_split_Min [simp]: |
|
|
25919
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents:
25762
diff
changeset
|
572 |
"bin_split n Int.Min = (Int.Min, bintrunc n Int.Min)" |
| 24364 | 573 |
by (induct n) (auto simp: Let_def split: ls_splits) |
574 |
||
575 |
lemma bin_split_trunc: |
|
576 |
"!!m b c. bin_split (min m n) c = (a, b) ==> |
|
577 |
bin_split n (bintrunc m c) = (bintrunc (m - n) a, b)" |
|
578 |
apply (induct n, clarsimp) |
|
579 |
apply (simp add: bin_rest_trunc Let_def split: ls_splits) |
|
580 |
apply (case_tac m) |
|
581 |
apply (auto simp: Let_def split: ls_splits) |
|
| 24333 | 582 |
done |
583 |
||
| 24364 | 584 |
lemma bin_split_trunc1: |
585 |
"!!m b c. bin_split n c = (a, b) ==> |
|
586 |
bin_split n (bintrunc m c) = (bintrunc (m - n) a, bintrunc m b)" |
|
587 |
apply (induct n, clarsimp) |
|
588 |
apply (simp add: bin_rest_trunc Let_def split: ls_splits) |
|
589 |
apply (case_tac m) |
|
590 |
apply (auto simp: Let_def split: ls_splits) |
|
591 |
done |
|
| 24333 | 592 |
|
| 24364 | 593 |
lemma bin_cat_num: |
594 |
"!!b. bin_cat a n b = a * 2 ^ n + bintrunc n b" |
|
595 |
apply (induct n, clarsimp) |
|
596 |
apply (simp add: Bit_def cong: number_of_False_cong) |
|
597 |
done |
|
598 |
||
599 |
lemma bin_split_num: |
|
600 |
"!!b. bin_split n b = (b div 2 ^ n, b mod 2 ^ n)" |
|
601 |
apply (induct n, clarsimp) |
|
|
45529
0e1037d4e049
remove redundant lemmas bin_last_mod and bin_rest_div, use bin_last_def and bin_rest_def instead
huffman
parents:
45475
diff
changeset
|
602 |
apply (simp add: bin_rest_def zdiv_zmult2_eq) |
| 24364 | 603 |
apply (case_tac b rule: bin_exhaust) |
604 |
apply simp |
|
| 37667 | 605 |
apply (simp add: Bit_def mod_mult_mult1 p1mod22k bitval_def |
| 24364 | 606 |
split: bit.split |
607 |
cong: number_of_False_cong) |
|
608 |
done |
|
609 |
||
610 |
subsection {* Miscellaneous lemmas *}
|
|
| 24333 | 611 |
|
612 |
lemma nth_2p_bin: |
|
613 |
"!!m. bin_nth (2 ^ n) m = (m = n)" |
|
614 |
apply (induct n) |
|
615 |
apply clarsimp |
|
616 |
apply safe |
|
617 |
apply (case_tac m) |
|
618 |
apply (auto simp: trans [OF numeral_1_eq_1 [symmetric] number_of_eq]) |
|
619 |
apply (case_tac m) |
|
620 |
apply (auto simp: Bit_B0_2t [symmetric]) |
|
621 |
done |
|
622 |
||
623 |
(* for use when simplifying with bin_nth_Bit *) |
|
624 |
||
625 |
lemma ex_eq_or: |
|
626 |
"(EX m. n = Suc m & (m = k | P m)) = (n = Suc k | (EX m. n = Suc m & P m))" |
|
627 |
by auto |
|
628 |
||
629 |
end |
|
630 |