author | kuncar |
Wed, 23 Apr 2014 17:57:56 +0200 | |
changeset 56677 | 660ffb526069 |
parent 54874 | c55c5dacd6a1 |
child 58410 | 6d46ad54a2ab |
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 |
||
54854
3324a0078636
prefer "Bits" as theory name for abstract bit operations, similar to "Orderings", "Lattices", "Groups" etc.
haftmann
parents:
54848
diff
changeset
|
11 |
theory Bits_Int |
3324a0078636
prefer "Bits" as theory name for abstract bit operations, similar to "Orderings", "Lattices", "Groups" etc.
haftmann
parents:
54848
diff
changeset
|
12 |
imports Bits Bit_Representation |
24333 | 13 |
begin |
14 |
||
24364 | 15 |
subsection {* Logical operations *} |
24353 | 16 |
|
17 |
text "bit-wise logical operations on the int type" |
|
18 |
||
25762 | 19 |
instantiation int :: bit |
20 |
begin |
|
21 |
||
46019 | 22 |
definition int_not_def: |
23 |
"bitNOT = (\<lambda>x::int. - x - 1)" |
|
24 |
||
25 |
function bitAND_int where |
|
26 |
"bitAND_int x y = |
|
27 |
(if x = 0 then 0 else if x = -1 then y else |
|
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
28 |
(bin_rest x AND bin_rest y) BIT (bin_last x \<and> bin_last y))" |
46019 | 29 |
by pat_completeness simp |
25762 | 30 |
|
46019 | 31 |
termination |
32 |
by (relation "measure (nat o abs o fst)", simp_all add: bin_rest_def) |
|
33 |
||
34 |
declare bitAND_int.simps [simp del] |
|
25762 | 35 |
|
46019 | 36 |
definition int_or_def: |
37 |
"bitOR = (\<lambda>x y::int. NOT (NOT x AND NOT y))" |
|
25762 | 38 |
|
46019 | 39 |
definition int_xor_def: |
40 |
"bitXOR = (\<lambda>x y::int. (x AND NOT y) OR (NOT x AND y))" |
|
25762 | 41 |
|
42 |
instance .. |
|
43 |
||
44 |
end |
|
24353 | 45 |
|
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
46 |
subsubsection {* Basic simplification rules *} |
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
47 |
|
46016 | 48 |
lemma int_not_BIT [simp]: |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
49 |
"NOT (w BIT b) = (NOT w) BIT (\<not> b)" |
46016 | 50 |
unfolding int_not_def Bit_def by (cases b, simp_all) |
51 |
||
24333 | 52 |
lemma int_not_simps [simp]: |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
53 |
"NOT (0::int) = -1" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
54 |
"NOT (1::int) = -2" |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54427
diff
changeset
|
55 |
"NOT (- 1::int) = 0" |
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54427
diff
changeset
|
56 |
"NOT (numeral w::int) = - numeral (w + Num.One)" |
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54427
diff
changeset
|
57 |
"NOT (- numeral (Num.Bit0 w)::int) = numeral (Num.BitM w)" |
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54427
diff
changeset
|
58 |
"NOT (- numeral (Num.Bit1 w)::int) = numeral (Num.Bit0 w)" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
59 |
unfolding int_not_def by simp_all |
24333 | 60 |
|
46017 | 61 |
lemma int_not_not [simp]: "NOT (NOT (x::int)) = x" |
62 |
unfolding int_not_def by simp |
|
63 |
||
46019 | 64 |
lemma int_and_0 [simp]: "(0::int) AND x = 0" |
65 |
by (simp add: bitAND_int.simps) |
|
66 |
||
67 |
lemma int_and_m1 [simp]: "(-1::int) AND x = x" |
|
68 |
by (simp add: bitAND_int.simps) |
|
69 |
||
46017 | 70 |
lemma int_and_Bits [simp]: |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
71 |
"(x BIT b) AND (y BIT c) = (x AND y) BIT (b \<and> c)" |
46019 | 72 |
by (subst bitAND_int.simps, simp add: Bit_eq_0_iff Bit_eq_m1_iff) |
46017 | 73 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
74 |
lemma int_or_zero [simp]: "(0::int) OR x = x" |
46017 | 75 |
unfolding int_or_def by simp |
46018 | 76 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
77 |
lemma int_or_minus1 [simp]: "(-1::int) OR x = -1" |
46017 | 78 |
unfolding int_or_def by simp |
79 |
||
24333 | 80 |
lemma int_or_Bits [simp]: |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
81 |
"(x BIT b) OR (y BIT c) = (x OR y) BIT (b \<or> c)" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
82 |
unfolding int_or_def by simp |
24333 | 83 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
84 |
lemma int_xor_zero [simp]: "(0::int) XOR x = x" |
46018 | 85 |
unfolding int_xor_def by simp |
86 |
||
87 |
lemma int_xor_Bits [simp]: |
|
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
88 |
"(x BIT b) XOR (y BIT c) = (x XOR y) BIT ((b \<or> c) \<and> \<not> (b \<and> c))" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
89 |
unfolding int_xor_def by auto |
46018 | 90 |
|
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
91 |
subsubsection {* Binary destructors *} |
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
92 |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
93 |
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
|
94 |
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
|
95 |
|
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
96 |
lemma bin_last_NOT [simp]: "bin_last (NOT x) \<longleftrightarrow> \<not> bin_last x" |
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
97 |
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
|
98 |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
99 |
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
|
100 |
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
|
101 |
|
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
102 |
lemma bin_last_AND [simp]: "bin_last (x AND y) \<longleftrightarrow> bin_last x \<and> bin_last y" |
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
103 |
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
|
104 |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
105 |
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
|
106 |
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
|
107 |
|
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
108 |
lemma bin_last_OR [simp]: "bin_last (x OR y) \<longleftrightarrow> bin_last x \<or> bin_last y" |
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
109 |
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
|
110 |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
111 |
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
|
112 |
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
|
113 |
|
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
114 |
lemma bin_last_XOR [simp]: "bin_last (x XOR y) \<longleftrightarrow> (bin_last x \<or> bin_last y) \<and> \<not> (bin_last x \<and> bin_last y)" |
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
115 |
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
|
116 |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
117 |
lemma bin_nth_ops: |
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
118 |
"!!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
|
119 |
"!!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
|
120 |
"!!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
|
121 |
"!!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
|
122 |
by (induct n) auto |
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
123 |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
124 |
subsubsection {* Derived properties *} |
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
125 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
126 |
lemma int_xor_minus1 [simp]: "(-1::int) XOR x = NOT x" |
46018 | 127 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
128 |
||
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
129 |
lemma int_xor_extra_simps [simp]: |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
130 |
"w XOR (0::int) = w" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
131 |
"w XOR (-1::int) = NOT w" |
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
132 |
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
|
133 |
|
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
134 |
lemma int_or_extra_simps [simp]: |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
135 |
"w OR (0::int) = w" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
136 |
"w OR (-1::int) = -1" |
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
137 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
24333 | 138 |
|
37667 | 139 |
lemma int_and_extra_simps [simp]: |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
140 |
"w AND (0::int) = 0" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
141 |
"w AND (-1::int) = w" |
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
142 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
24333 | 143 |
|
144 |
(* commutativity of the above *) |
|
145 |
lemma bin_ops_comm: |
|
146 |
shows |
|
24353 | 147 |
int_and_comm: "!!y::int. x AND y = y AND x" and |
148 |
int_or_comm: "!!y::int. x OR y = y OR x" and |
|
149 |
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
|
150 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
24333 | 151 |
|
152 |
lemma bin_ops_same [simp]: |
|
24353 | 153 |
"(x::int) AND x = x" |
154 |
"(x::int) OR x = x" |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
155 |
"(x::int) XOR x = 0" |
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
156 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
24333 | 157 |
|
158 |
lemmas bin_log_esimps = |
|
159 |
int_and_extra_simps int_or_extra_simps int_xor_extra_simps |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
160 |
int_and_0 int_and_m1 int_or_zero int_or_minus1 int_xor_zero int_xor_minus1 |
24333 | 161 |
|
162 |
(* basic properties of logical (bit-wise) operations *) |
|
163 |
||
164 |
lemma bbw_ao_absorb: |
|
24353 | 165 |
"!!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
|
166 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
24333 | 167 |
|
168 |
lemma bbw_ao_absorbs_other: |
|
24353 | 169 |
"x AND (x OR y) = x \<and> (y AND x) OR x = (x::int)" |
170 |
"(y OR x) AND x = x \<and> x OR (x AND y) = (x::int)" |
|
171 |
"(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
|
172 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
24353 | 173 |
|
24333 | 174 |
lemmas bbw_ao_absorbs [simp] = bbw_ao_absorb bbw_ao_absorbs_other |
175 |
||
176 |
lemma int_xor_not: |
|
24353 | 177 |
"!!y::int. (NOT x) XOR y = NOT (x XOR y) & |
178 |
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
|
179 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
24333 | 180 |
|
181 |
lemma int_and_assoc: |
|
24353 | 182 |
"(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
|
183 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
24333 | 184 |
|
185 |
lemma int_or_assoc: |
|
24353 | 186 |
"(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
|
187 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
24333 | 188 |
|
189 |
lemma int_xor_assoc: |
|
24353 | 190 |
"(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
|
191 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
24333 | 192 |
|
193 |
lemmas bbw_assocs = int_and_assoc int_or_assoc int_xor_assoc |
|
194 |
||
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
195 |
(* BH: Why are these declared as simp rules??? *) |
24333 | 196 |
lemma bbw_lcs [simp]: |
24353 | 197 |
"(y::int) AND (x AND z) = x AND (y AND z)" |
198 |
"(y::int) OR (x OR z) = x OR (y OR z)" |
|
199 |
"(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
|
200 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
24333 | 201 |
|
202 |
lemma bbw_not_dist: |
|
24353 | 203 |
"!!y::int. NOT (x OR y) = (NOT x) AND (NOT y)" |
204 |
"!!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
|
205 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
24333 | 206 |
|
207 |
lemma bbw_oa_dist: |
|
24353 | 208 |
"!!y z::int. (x AND y) OR z = |
209 |
(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
|
210 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
24333 | 211 |
|
212 |
lemma bbw_ao_dist: |
|
24353 | 213 |
"!!y z::int. (x OR y) AND z = |
214 |
(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
|
215 |
by (auto simp add: bin_eq_iff bin_nth_ops) |
24333 | 216 |
|
24367
3e29eafabe16
AC rules for bitwise logical operators no longer declared simp
huffman
parents:
24366
diff
changeset
|
217 |
(* |
3e29eafabe16
AC rules for bitwise logical operators no longer declared simp
huffman
parents:
24366
diff
changeset
|
218 |
Why were these declared simp??? |
24333 | 219 |
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
|
220 |
*) |
24333 | 221 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
222 |
subsubsection {* Simplification with numerals *} |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
223 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
224 |
text {* Cases for @{text "0"} and @{text "-1"} are already covered by |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
225 |
other simp rules. *} |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
226 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
227 |
lemma bin_rl_eqI: "\<lbrakk>bin_rest x = bin_rest y; bin_last x = bin_last y\<rbrakk> \<Longrightarrow> x = y" |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
228 |
by (metis (mono_tags) BIT_eq_iff bin_ex_rl bin_last_BIT bin_rest_BIT) |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
229 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
230 |
lemma bin_rest_neg_numeral_BitM [simp]: |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54427
diff
changeset
|
231 |
"bin_rest (- numeral (Num.BitM w)) = - numeral w" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
232 |
by (simp only: BIT_bin_simps [symmetric] bin_rest_BIT) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
233 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
234 |
lemma bin_last_neg_numeral_BitM [simp]: |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
235 |
"bin_last (- numeral (Num.BitM w))" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
236 |
by (simp only: BIT_bin_simps [symmetric] bin_last_BIT) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
237 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
238 |
text {* FIXME: The rule sets below are very large (24 rules for each |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
239 |
operator). Is there a simpler way to do this? *} |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
240 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
241 |
lemma int_and_numerals [simp]: |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
242 |
"numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (numeral x AND numeral y) BIT False" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
243 |
"numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (numeral x AND numeral y) BIT False" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
244 |
"numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (numeral x AND numeral y) BIT False" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
245 |
"numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = (numeral x AND numeral y) BIT True" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
246 |
"numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (numeral x AND - numeral y) BIT False" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
247 |
"numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (numeral x AND - numeral (y + Num.One)) BIT False" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
248 |
"numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (numeral x AND - numeral y) BIT False" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
249 |
"numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = (numeral x AND - numeral (y + Num.One)) BIT True" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
250 |
"- numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (- numeral x AND numeral y) BIT False" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
251 |
"- numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (- numeral x AND numeral y) BIT False" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
252 |
"- numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (- numeral (x + Num.One) AND numeral y) BIT False" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
253 |
"- numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = (- numeral (x + Num.One) AND numeral y) BIT True" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
254 |
"- numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (- numeral x AND - numeral y) BIT False" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
255 |
"- numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (- numeral x AND - numeral (y + Num.One)) BIT False" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
256 |
"- numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (- numeral (x + Num.One) AND - numeral y) BIT False" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
257 |
"- numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = (- numeral (x + Num.One) AND - numeral (y + Num.One)) BIT True" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
258 |
"(1::int) AND numeral (Num.Bit0 y) = 0" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
259 |
"(1::int) AND numeral (Num.Bit1 y) = 1" |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54427
diff
changeset
|
260 |
"(1::int) AND - numeral (Num.Bit0 y) = 0" |
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54427
diff
changeset
|
261 |
"(1::int) AND - numeral (Num.Bit1 y) = 1" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
262 |
"numeral (Num.Bit0 x) AND (1::int) = 0" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
263 |
"numeral (Num.Bit1 x) AND (1::int) = 1" |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54427
diff
changeset
|
264 |
"- numeral (Num.Bit0 x) AND (1::int) = 0" |
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54427
diff
changeset
|
265 |
"- numeral (Num.Bit1 x) AND (1::int) = 1" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
266 |
by (rule bin_rl_eqI, simp, simp)+ |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
267 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
268 |
lemma int_or_numerals [simp]: |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
269 |
"numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (numeral x OR numeral y) BIT False" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
270 |
"numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = (numeral x OR numeral y) BIT True" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
271 |
"numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = (numeral x OR numeral y) BIT True" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
272 |
"numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = (numeral x OR numeral y) BIT True" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
273 |
"numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (numeral x OR - numeral y) BIT False" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
274 |
"numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = (numeral x OR - numeral (y + Num.One)) BIT True" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
275 |
"numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = (numeral x OR - numeral y) BIT True" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
276 |
"numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = (numeral x OR - numeral (y + Num.One)) BIT True" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
277 |
"- numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (- numeral x OR numeral y) BIT False" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
278 |
"- numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = (- numeral x OR numeral y) BIT True" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
279 |
"- numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = (- numeral (x + Num.One) OR numeral y) BIT True" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
280 |
"- numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = (- numeral (x + Num.One) OR numeral y) BIT True" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
281 |
"- numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (- numeral x OR - numeral y) BIT False" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
282 |
"- numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = (- numeral x OR - numeral (y + Num.One)) BIT True" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
283 |
"- numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = (- numeral (x + Num.One) OR - numeral y) BIT True" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
284 |
"- numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = (- numeral (x + Num.One) OR - numeral (y + Num.One)) BIT True" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
285 |
"(1::int) OR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
286 |
"(1::int) OR numeral (Num.Bit1 y) = numeral (Num.Bit1 y)" |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54427
diff
changeset
|
287 |
"(1::int) OR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)" |
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54427
diff
changeset
|
288 |
"(1::int) OR - numeral (Num.Bit1 y) = - numeral (Num.Bit1 y)" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
289 |
"numeral (Num.Bit0 x) OR (1::int) = numeral (Num.Bit1 x)" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
290 |
"numeral (Num.Bit1 x) OR (1::int) = numeral (Num.Bit1 x)" |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54427
diff
changeset
|
291 |
"- numeral (Num.Bit0 x) OR (1::int) = - numeral (Num.BitM x)" |
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54427
diff
changeset
|
292 |
"- numeral (Num.Bit1 x) OR (1::int) = - numeral (Num.Bit1 x)" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
293 |
by (rule bin_rl_eqI, simp, simp)+ |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
294 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
295 |
lemma int_xor_numerals [simp]: |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
296 |
"numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (numeral x XOR numeral y) BIT False" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
297 |
"numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = (numeral x XOR numeral y) BIT True" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
298 |
"numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = (numeral x XOR numeral y) BIT True" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
299 |
"numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (numeral x XOR numeral y) BIT False" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
300 |
"numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (numeral x XOR - numeral y) BIT False" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
301 |
"numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = (numeral x XOR - numeral (y + Num.One)) BIT True" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
302 |
"numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = (numeral x XOR - numeral y) BIT True" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
303 |
"numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (numeral x XOR - numeral (y + Num.One)) BIT False" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
304 |
"- numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (- numeral x XOR numeral y) BIT False" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
305 |
"- numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = (- numeral x XOR numeral y) BIT True" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
306 |
"- numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = (- numeral (x + Num.One) XOR numeral y) BIT True" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
307 |
"- numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (- numeral (x + Num.One) XOR numeral y) BIT False" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
308 |
"- numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (- numeral x XOR - numeral y) BIT False" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
309 |
"- numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = (- numeral x XOR - numeral (y + Num.One)) BIT True" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
310 |
"- numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = (- numeral (x + Num.One) XOR - numeral y) BIT True" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
311 |
"- numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (- numeral (x + Num.One) XOR - numeral (y + Num.One)) BIT False" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
312 |
"(1::int) XOR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
313 |
"(1::int) XOR numeral (Num.Bit1 y) = numeral (Num.Bit0 y)" |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54427
diff
changeset
|
314 |
"(1::int) XOR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)" |
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54427
diff
changeset
|
315 |
"(1::int) XOR - numeral (Num.Bit1 y) = - numeral (Num.Bit0 (y + Num.One))" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
316 |
"numeral (Num.Bit0 x) XOR (1::int) = numeral (Num.Bit1 x)" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
317 |
"numeral (Num.Bit1 x) XOR (1::int) = numeral (Num.Bit0 x)" |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54427
diff
changeset
|
318 |
"- numeral (Num.Bit0 x) XOR (1::int) = - numeral (Num.BitM x)" |
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54427
diff
changeset
|
319 |
"- numeral (Num.Bit1 x) XOR (1::int) = - numeral (Num.Bit0 (x + Num.One))" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
320 |
by (rule bin_rl_eqI, simp, simp)+ |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
321 |
|
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
322 |
subsubsection {* Interactions with arithmetic *} |
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
323 |
|
24333 | 324 |
lemma plus_and_or [rule_format]: |
24353 | 325 |
"ALL y::int. (x AND y) + (x OR y) = x + y" |
24333 | 326 |
apply (induct x rule: bin_induct) |
327 |
apply clarsimp |
|
328 |
apply clarsimp |
|
329 |
apply clarsimp |
|
330 |
apply (case_tac y rule: bin_exhaust) |
|
331 |
apply clarsimp |
|
332 |
apply (unfold Bit_def) |
|
333 |
apply clarsimp |
|
334 |
apply (erule_tac x = "x" in allE) |
|
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
335 |
apply simp |
24333 | 336 |
done |
337 |
||
338 |
lemma le_int_or: |
|
46604
9f9e85264e4d
make uses of bin_sign respect int/bin distinction
huffman
parents:
46023
diff
changeset
|
339 |
"bin_sign (y::int) = 0 ==> x <= x OR y" |
37667 | 340 |
apply (induct y arbitrary: x rule: bin_induct) |
24333 | 341 |
apply clarsimp |
342 |
apply clarsimp |
|
343 |
apply (case_tac x rule: bin_exhaust) |
|
344 |
apply (case_tac b) |
|
345 |
apply (case_tac [!] bit) |
|
46604
9f9e85264e4d
make uses of bin_sign respect int/bin distinction
huffman
parents:
46023
diff
changeset
|
346 |
apply (auto simp: le_Bits) |
24333 | 347 |
done |
348 |
||
349 |
lemmas int_and_le = |
|
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
|
350 |
xtrans(3) [OF bbw_ao_absorbs (2) [THEN conjunct2, symmetric] le_int_or] |
24333 | 351 |
|
24364 | 352 |
(* interaction between bit-wise and arithmetic *) |
353 |
(* good example of bin_induction *) |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
354 |
lemma bin_add_not: "x + NOT x = (-1::int)" |
24364 | 355 |
apply (induct x rule: bin_induct) |
356 |
apply clarsimp |
|
357 |
apply clarsimp |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
358 |
apply (case_tac bit, auto) |
24364 | 359 |
done |
360 |
||
45543
827bf668c822
HOL-Word: add simp rules for bin_rest, bin_last; shorten several proofs
huffman
parents:
45529
diff
changeset
|
361 |
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
|
362 |
|
24364 | 363 |
lemma bin_trunc_ao: |
364 |
"!!x y. (bintrunc n x) AND (bintrunc n y) = bintrunc n (x AND y)" |
|
365 |
"!!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
|
366 |
by (auto simp add: bin_eq_iff bin_nth_ops nth_bintr) |
24364 | 367 |
|
368 |
lemma bin_trunc_xor: |
|
369 |
"!!x y. bintrunc n (bintrunc n x XOR bintrunc n y) = |
|
370 |
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
|
371 |
by (auto simp add: bin_eq_iff bin_nth_ops nth_bintr) |
24364 | 372 |
|
373 |
lemma bin_trunc_not: |
|
374 |
"!!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
|
375 |
by (auto simp add: bin_eq_iff bin_nth_ops nth_bintr) |
24364 | 376 |
|
377 |
(* want theorems of the form of bin_trunc_xor *) |
|
378 |
lemma bintr_bintr_i: |
|
379 |
"x = bintrunc n y ==> bintrunc n x = bintrunc n y" |
|
380 |
by auto |
|
381 |
||
382 |
lemmas bin_trunc_and = bin_trunc_ao(1) [THEN bintr_bintr_i] |
|
383 |
lemmas bin_trunc_or = bin_trunc_ao(2) [THEN bintr_bintr_i] |
|
384 |
||
385 |
subsection {* Setting and clearing bits *} |
|
386 |
||
54874 | 387 |
(** nth bit, set/clear **) |
388 |
||
26558 | 389 |
primrec |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
390 |
bin_sc :: "nat => bool => int => int" |
26558 | 391 |
where |
392 |
Z: "bin_sc 0 b w = bin_rest w BIT b" |
|
393 |
| Suc: "bin_sc (Suc n) b w = bin_sc n b (bin_rest w) BIT bin_last w" |
|
24364 | 394 |
|
24333 | 395 |
lemma bin_nth_sc [simp]: |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
396 |
"bin_nth (bin_sc n b w) n \<longleftrightarrow> b" |
45955 | 397 |
by (induct n arbitrary: w) auto |
24333 | 398 |
|
399 |
lemma bin_sc_sc_same [simp]: |
|
45955 | 400 |
"bin_sc n c (bin_sc n b w) = bin_sc n c w" |
401 |
by (induct n arbitrary: w) auto |
|
24333 | 402 |
|
403 |
lemma bin_sc_sc_diff: |
|
45955 | 404 |
"m ~= n ==> |
24333 | 405 |
bin_sc m c (bin_sc n b w) = bin_sc n b (bin_sc m c w)" |
45955 | 406 |
apply (induct n arbitrary: w m) |
24333 | 407 |
apply (case_tac [!] m) |
408 |
apply auto |
|
409 |
done |
|
410 |
||
411 |
lemma bin_nth_sc_gen: |
|
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
412 |
"bin_nth (bin_sc n b w) m = (if m = n then b else bin_nth w m)" |
45955 | 413 |
by (induct n arbitrary: w m) (case_tac [!] m, auto) |
24333 | 414 |
|
415 |
lemma bin_sc_nth [simp]: |
|
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
416 |
"(bin_sc n (bin_nth w n) w) = w" |
45955 | 417 |
by (induct n arbitrary: w) auto |
24333 | 418 |
|
419 |
lemma bin_sign_sc [simp]: |
|
45955 | 420 |
"bin_sign (bin_sc n b w) = bin_sign w" |
421 |
by (induct n arbitrary: w) auto |
|
24333 | 422 |
|
423 |
lemma bin_sc_bintr [simp]: |
|
45955 | 424 |
"bintrunc m (bin_sc n x (bintrunc m (w))) = bintrunc m (bin_sc n x w)" |
425 |
apply (induct n arbitrary: w m) |
|
24333 | 426 |
apply (case_tac [!] w rule: bin_exhaust) |
427 |
apply (case_tac [!] m, auto) |
|
428 |
done |
|
429 |
||
430 |
lemma bin_clr_le: |
|
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
431 |
"bin_sc n False w <= w" |
45955 | 432 |
apply (induct n arbitrary: w) |
24333 | 433 |
apply (case_tac [!] w rule: bin_exhaust) |
46605 | 434 |
apply (auto simp: le_Bits) |
24333 | 435 |
done |
436 |
||
437 |
lemma bin_set_ge: |
|
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
438 |
"bin_sc n True w >= w" |
45955 | 439 |
apply (induct n arbitrary: w) |
24333 | 440 |
apply (case_tac [!] w rule: bin_exhaust) |
46605 | 441 |
apply (auto simp: le_Bits) |
24333 | 442 |
done |
443 |
||
444 |
lemma bintr_bin_clr_le: |
|
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
445 |
"bintrunc n (bin_sc m False w) <= bintrunc n w" |
45955 | 446 |
apply (induct n arbitrary: w m) |
24333 | 447 |
apply simp |
448 |
apply (case_tac w rule: bin_exhaust) |
|
449 |
apply (case_tac m) |
|
46605 | 450 |
apply (auto simp: le_Bits) |
24333 | 451 |
done |
452 |
||
453 |
lemma bintr_bin_set_ge: |
|
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
454 |
"bintrunc n (bin_sc m True w) >= bintrunc n w" |
45955 | 455 |
apply (induct n arbitrary: w m) |
24333 | 456 |
apply simp |
457 |
apply (case_tac w rule: bin_exhaust) |
|
458 |
apply (case_tac m) |
|
46605 | 459 |
apply (auto simp: le_Bits) |
24333 | 460 |
done |
461 |
||
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
462 |
lemma bin_sc_FP [simp]: "bin_sc n False 0 = 0" |
46608
37e383cc7831
make uses of constant bin_sc respect int/bin distinction
huffman
parents:
46605
diff
changeset
|
463 |
by (induct n) auto |
24333 | 464 |
|
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
465 |
lemma bin_sc_TM [simp]: "bin_sc n True -1 = -1" |
46608
37e383cc7831
make uses of constant bin_sc respect int/bin distinction
huffman
parents:
46605
diff
changeset
|
466 |
by (induct n) auto |
24333 | 467 |
|
468 |
lemmas bin_sc_simps = bin_sc.Z bin_sc.Suc bin_sc_TM bin_sc_FP |
|
469 |
||
470 |
lemma bin_sc_minus: |
|
471 |
"0 < n ==> bin_sc (Suc (n - 1)) b w = bin_sc n b w" |
|
472 |
by auto |
|
473 |
||
474 |
lemmas bin_sc_Suc_minus = |
|
45604 | 475 |
trans [OF bin_sc_minus [symmetric] bin_sc.Suc] |
24333 | 476 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
477 |
lemma bin_sc_numeral [simp]: |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46610
diff
changeset
|
478 |
"bin_sc (numeral k) b w = |
47219
172c031ad743
restate various simp rules for word operations using pred_numeral
huffman
parents:
47108
diff
changeset
|
479 |
bin_sc (pred_numeral k) b (bin_rest w) BIT bin_last w" |
172c031ad743
restate various simp rules for word operations using pred_numeral
huffman
parents:
47108
diff
changeset
|
480 |
by (simp add: numeral_eq_Suc) |
24333 | 481 |
|
24465 | 482 |
|
24364 | 483 |
subsection {* Splitting and concatenation *} |
24333 | 484 |
|
54848 | 485 |
definition bin_rcat :: "nat \<Rightarrow> int list \<Rightarrow> int" |
486 |
where |
|
37667 | 487 |
"bin_rcat n = foldl (\<lambda>u v. bin_cat u n v) 0" |
488 |
||
54848 | 489 |
fun bin_rsplit_aux :: "nat \<Rightarrow> nat \<Rightarrow> int \<Rightarrow> int list \<Rightarrow> int list" |
490 |
where |
|
26558 | 491 |
"bin_rsplit_aux n m c bs = |
24364 | 492 |
(if m = 0 | n = 0 then bs else |
493 |
let (a, b) = bin_split n c |
|
26558 | 494 |
in bin_rsplit_aux n (m - n) a (b # bs))" |
24364 | 495 |
|
54848 | 496 |
definition bin_rsplit :: "nat \<Rightarrow> nat \<times> int \<Rightarrow> int list" |
497 |
where |
|
26558 | 498 |
"bin_rsplit n w = bin_rsplit_aux n (fst w) (snd w) []" |
499 |
||
54848 | 500 |
fun bin_rsplitl_aux :: "nat \<Rightarrow> nat \<Rightarrow> int \<Rightarrow> int list \<Rightarrow> int list" |
501 |
where |
|
26558 | 502 |
"bin_rsplitl_aux n m c bs = |
24364 | 503 |
(if m = 0 | n = 0 then bs else |
504 |
let (a, b) = bin_split (min m n) c |
|
26558 | 505 |
in bin_rsplitl_aux n (m - n) a (b # bs))" |
24364 | 506 |
|
54848 | 507 |
definition bin_rsplitl :: "nat \<Rightarrow> nat \<times> int \<Rightarrow> int list" |
508 |
where |
|
26558 | 509 |
"bin_rsplitl n w = bin_rsplitl_aux n (fst w) (snd w) []" |
510 |
||
24364 | 511 |
declare bin_rsplit_aux.simps [simp del] |
512 |
declare bin_rsplitl_aux.simps [simp del] |
|
513 |
||
514 |
lemma bin_sign_cat: |
|
45955 | 515 |
"bin_sign (bin_cat x n y) = bin_sign x" |
516 |
by (induct n arbitrary: y) auto |
|
24364 | 517 |
|
518 |
lemma bin_cat_Suc_Bit: |
|
519 |
"bin_cat w (Suc n) (v BIT b) = bin_cat w n v BIT b" |
|
520 |
by auto |
|
521 |
||
522 |
lemma bin_nth_cat: |
|
45955 | 523 |
"bin_nth (bin_cat x k y) n = |
24364 | 524 |
(if n < k then bin_nth y n else bin_nth x (n - k))" |
45955 | 525 |
apply (induct k arbitrary: n y) |
24364 | 526 |
apply clarsimp |
527 |
apply (case_tac n, auto) |
|
24333 | 528 |
done |
529 |
||
24364 | 530 |
lemma bin_nth_split: |
45955 | 531 |
"bin_split n c = (a, b) ==> |
24364 | 532 |
(ALL k. bin_nth a k = bin_nth c (n + k)) & |
533 |
(ALL k. bin_nth b k = (k < n & bin_nth c k))" |
|
45955 | 534 |
apply (induct n arbitrary: b c) |
24364 | 535 |
apply clarsimp |
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
|
536 |
apply (clarsimp simp: Let_def split: prod.split_asm) |
24364 | 537 |
apply (case_tac k) |
538 |
apply auto |
|
539 |
done |
|
540 |
||
541 |
lemma bin_cat_assoc: |
|
45955 | 542 |
"bin_cat (bin_cat x m y) n z = bin_cat x (m + n) (bin_cat y n z)" |
543 |
by (induct n arbitrary: z) auto |
|
24364 | 544 |
|
45955 | 545 |
lemma bin_cat_assoc_sym: |
546 |
"bin_cat x m (bin_cat y n z) = bin_cat (bin_cat x (m - n) y) (min m n) z" |
|
547 |
apply (induct n arbitrary: z m, clarsimp) |
|
24364 | 548 |
apply (case_tac m, auto) |
24333 | 549 |
done |
550 |
||
45956 | 551 |
lemma bin_cat_zero [simp]: "bin_cat 0 n w = bintrunc n w" |
46001
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
huffman
parents:
45956
diff
changeset
|
552 |
by (induct n arbitrary: w) auto |
45956 | 553 |
|
24364 | 554 |
lemma bintr_cat1: |
45955 | 555 |
"bintrunc (k + n) (bin_cat a n b) = bin_cat (bintrunc k a) n b" |
556 |
by (induct n arbitrary: b) auto |
|
24364 | 557 |
|
558 |
lemma bintr_cat: "bintrunc m (bin_cat a n b) = |
|
559 |
bin_cat (bintrunc (m - n) a) n (bintrunc (min m n) b)" |
|
560 |
by (rule bin_eqI) (auto simp: bin_nth_cat nth_bintr) |
|
561 |
||
562 |
lemma bintr_cat_same [simp]: |
|
563 |
"bintrunc n (bin_cat a n b) = bintrunc n b" |
|
564 |
by (auto simp add : bintr_cat) |
|
565 |
||
566 |
lemma cat_bintr [simp]: |
|
45955 | 567 |
"bin_cat a n (bintrunc n b) = bin_cat a n b" |
568 |
by (induct n arbitrary: b) auto |
|
24364 | 569 |
|
570 |
lemma split_bintrunc: |
|
45955 | 571 |
"bin_split n c = (a, b) ==> b = bintrunc n c" |
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
|
572 |
by (induct n arbitrary: b c) (auto simp: Let_def split: prod.split_asm) |
24364 | 573 |
|
574 |
lemma bin_cat_split: |
|
45955 | 575 |
"bin_split n w = (u, v) ==> w = bin_cat u n v" |
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
|
576 |
by (induct n arbitrary: v w) (auto simp: Let_def split: prod.split_asm) |
24364 | 577 |
|
578 |
lemma bin_split_cat: |
|
45955 | 579 |
"bin_split n (bin_cat v n w) = (v, bintrunc n w)" |
580 |
by (induct n arbitrary: w) auto |
|
24364 | 581 |
|
45956 | 582 |
lemma bin_split_zero [simp]: "bin_split n 0 = (0, 0)" |
46001
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
huffman
parents:
45956
diff
changeset
|
583 |
by (induct n) auto |
45956 | 584 |
|
46610
0c3a5e28f425
make uses of bin_split respect int/bin distinction
huffman
parents:
46609
diff
changeset
|
585 |
lemma bin_split_minus1 [simp]: |
0c3a5e28f425
make uses of bin_split respect int/bin distinction
huffman
parents:
46609
diff
changeset
|
586 |
"bin_split n -1 = (-1, bintrunc n -1)" |
0c3a5e28f425
make uses of bin_split respect int/bin distinction
huffman
parents:
46609
diff
changeset
|
587 |
by (induct n) auto |
24364 | 588 |
|
589 |
lemma bin_split_trunc: |
|
45955 | 590 |
"bin_split (min m n) c = (a, b) ==> |
24364 | 591 |
bin_split n (bintrunc m c) = (bintrunc (m - n) a, b)" |
45955 | 592 |
apply (induct n arbitrary: m b c, clarsimp) |
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
|
593 |
apply (simp add: bin_rest_trunc Let_def split: prod.split_asm) |
24364 | 594 |
apply (case_tac m) |
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
|
595 |
apply (auto simp: Let_def split: prod.split_asm) |
24333 | 596 |
done |
597 |
||
24364 | 598 |
lemma bin_split_trunc1: |
45955 | 599 |
"bin_split n c = (a, b) ==> |
24364 | 600 |
bin_split n (bintrunc m c) = (bintrunc (m - n) a, bintrunc m b)" |
45955 | 601 |
apply (induct n arbitrary: m b c, clarsimp) |
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
|
602 |
apply (simp add: bin_rest_trunc Let_def split: prod.split_asm) |
24364 | 603 |
apply (case_tac m) |
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
|
604 |
apply (auto simp: Let_def split: prod.split_asm) |
24364 | 605 |
done |
24333 | 606 |
|
24364 | 607 |
lemma bin_cat_num: |
45955 | 608 |
"bin_cat a n b = a * 2 ^ n + bintrunc n b" |
609 |
apply (induct n arbitrary: b, clarsimp) |
|
46001
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
huffman
parents:
45956
diff
changeset
|
610 |
apply (simp add: Bit_def) |
24364 | 611 |
done |
612 |
||
613 |
lemma bin_split_num: |
|
45955 | 614 |
"bin_split n b = (b div 2 ^ n, b mod 2 ^ n)" |
46610
0c3a5e28f425
make uses of bin_split respect int/bin distinction
huffman
parents:
46609
diff
changeset
|
615 |
apply (induct n arbitrary: b, simp) |
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
|
616 |
apply (simp add: bin_rest_def zdiv_zmult2_eq) |
24364 | 617 |
apply (case_tac b rule: bin_exhaust) |
618 |
apply simp |
|
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
619 |
apply (simp add: Bit_def mod_mult_mult1 p1mod22k) |
45955 | 620 |
done |
24364 | 621 |
|
622 |
subsection {* Miscellaneous lemmas *} |
|
24333 | 623 |
|
624 |
lemma nth_2p_bin: |
|
45955 | 625 |
"bin_nth (2 ^ n) m = (m = n)" |
626 |
apply (induct n arbitrary: m) |
|
24333 | 627 |
apply clarsimp |
628 |
apply safe |
|
629 |
apply (case_tac m) |
|
630 |
apply (auto simp: Bit_B0_2t [symmetric]) |
|
631 |
done |
|
632 |
||
633 |
(* for use when simplifying with bin_nth_Bit *) |
|
634 |
||
635 |
lemma ex_eq_or: |
|
636 |
"(EX m. n = Suc m & (m = k | P m)) = (n = Suc k | (EX m. n = Suc m & P m))" |
|
637 |
by auto |
|
638 |
||
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
639 |
lemma power_BIT: "2 ^ (Suc n) - 1 = (2 ^ n - 1) BIT True" |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
640 |
unfolding Bit_B1 |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
641 |
by (induct n) simp_all |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
642 |
|
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
643 |
lemma mod_BIT: |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
644 |
"bin BIT bit mod 2 ^ Suc n = (bin mod 2 ^ n) BIT bit" |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
645 |
proof - |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
646 |
have "bin mod 2 ^ n < 2 ^ n" by simp |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
647 |
then have "bin mod 2 ^ n \<le> 2 ^ n - 1" by simp |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
648 |
then have "2 * (bin mod 2 ^ n) \<le> 2 * (2 ^ n - 1)" |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
649 |
by (rule mult_left_mono) simp |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
650 |
then have "2 * (bin mod 2 ^ n) + 1 < 2 * 2 ^ n" by simp |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
651 |
then show ?thesis |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
652 |
by (auto simp add: Bit_def mod_mult_mult1 mod_add_left_eq [of "2 * bin"] |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
653 |
mod_pos_pos_trivial) |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
654 |
qed |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
655 |
|
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
656 |
lemma AND_mod: |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
657 |
fixes x :: int |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
658 |
shows "x AND 2 ^ n - 1 = x mod 2 ^ n" |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
659 |
proof (induct x arbitrary: n rule: bin_induct) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
660 |
case 1 |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
661 |
then show ?case |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
662 |
by simp |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
663 |
next |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
664 |
case 2 |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
665 |
then show ?case |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
666 |
by (simp, simp add: m1mod2k) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
667 |
next |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
668 |
case (3 bin bit) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
669 |
show ?case |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
670 |
proof (cases n) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
671 |
case 0 |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54489
diff
changeset
|
672 |
then show ?thesis by simp |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
673 |
next |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
674 |
case (Suc m) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
675 |
with 3 show ?thesis |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
676 |
by (simp only: power_BIT mod_BIT int_and_Bits) simp |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
677 |
qed |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
678 |
qed |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
54224
diff
changeset
|
679 |
|
24333 | 680 |
end |
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
|
681 |