| author | wenzelm |
| Mon, 13 Nov 2006 12:10:49 +0100 | |
| changeset 21318 | edb595802d22 |
| parent 20900 | c1ba49ade6a5 |
| child 21404 | eb85850d3eb7 |
| permissions | -rw-r--r-- |
| 20485 | 1 |
(* Title: HOL/Integ/Numeral.thy |
| 15013 | 2 |
ID: $Id$ |
| 20485 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
4 |
Copyright 1994 University of Cambridge |
|
| 15013 | 5 |
*) |
6 |
||
| 20485 | 7 |
header {* Arithmetic on Binary Integers *}
|
| 15013 | 8 |
|
| 15131 | 9 |
theory Numeral |
|
15620
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
paulson
parents:
15140
diff
changeset
|
10 |
imports IntDef Datatype |
| 16417 | 11 |
uses "../Tools/numeral_syntax.ML" |
| 15131 | 12 |
begin |
| 15013 | 13 |
|
| 20485 | 14 |
text {*
|
15 |
This formalization defines binary arithmetic in terms of the integers |
|
16 |
rather than using a datatype. This avoids multiple representations (leading |
|
17 |
zeroes, etc.) See @{text "ZF/Integ/twos-compl.ML"}, function @{text
|
|
18 |
int_of_binary}, for the numerical interpretation. |
|
| 15013 | 19 |
|
| 20485 | 20 |
The representation expects that @{text "(m mod 2)"} is 0 or 1,
|
21 |
even if m is negative; |
|
22 |
For instance, @{text "-5 div 2 = -3"} and @{text "-5 mod 2 = 1"}; thus
|
|
23 |
@{text "-5 = (-3)*2 + 1"}.
|
|
| 15013 | 24 |
*} |
25 |
||
| 20485 | 26 |
text{*
|
27 |
This datatype avoids the use of type @{typ bool}, which would make
|
|
28 |
all of the rewrite rules higher-order. |
|
29 |
*} |
|
| 15013 | 30 |
|
|
15620
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
paulson
parents:
15140
diff
changeset
|
31 |
datatype bit = B0 | B1 |
|
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
paulson
parents:
15140
diff
changeset
|
32 |
|
| 15013 | 33 |
constdefs |
| 20485 | 34 |
Pls :: int |
35 |
"Pls == 0" |
|
36 |
Min :: int |
|
37 |
"Min == - 1" |
|
38 |
Bit :: "int \<Rightarrow> bit \<Rightarrow> int" (infixl "BIT" 90) |
|
39 |
"k BIT b == (case b of B0 \<Rightarrow> 0 | B1 \<Rightarrow> 1) + k + k" |
|
| 15013 | 40 |
|
41 |
axclass |
|
42 |
number < type -- {* for numeric types: nat, int, real, \dots *}
|
|
43 |
||
44 |
consts |
|
| 20485 | 45 |
number_of :: "int \<Rightarrow> 'a::number" |
| 15013 | 46 |
|
47 |
syntax |
|
| 20485 | 48 |
"_Numeral" :: "num_const \<Rightarrow> 'a" ("_")
|
| 15013 | 49 |
|
50 |
setup NumeralSyntax.setup |
|
51 |
||
| 19380 | 52 |
abbreviation |
| 20485 | 53 |
"Numeral0 \<equiv> number_of Pls" |
54 |
"Numeral1 \<equiv> number_of (Pls BIT B1)" |
|
| 15013 | 55 |
|
| 20485 | 56 |
lemma Let_number_of [simp]: "Let (number_of v) f = f (number_of v)" |
| 15013 | 57 |
-- {* Unfold all @{text let}s involving constants *}
|
| 20485 | 58 |
unfolding Let_def .. |
59 |
||
60 |
lemma Let_0 [simp]: "Let 0 f = f 0" |
|
61 |
unfolding Let_def .. |
|
62 |
||
63 |
lemma Let_1 [simp]: "Let 1 f = f 1" |
|
64 |
unfolding Let_def .. |
|
| 15013 | 65 |
|
| 20485 | 66 |
definition |
67 |
succ :: "int \<Rightarrow> int" |
|
68 |
"succ k = k + 1" |
|
69 |
pred :: "int \<Rightarrow> int" |
|
70 |
"pred k = k - 1" |
|
71 |
||
72 |
lemmas numeral_simps = |
|
73 |
succ_def pred_def Pls_def Min_def Bit_def |
|
| 15013 | 74 |
|
| 20485 | 75 |
text {* Removal of leading zeroes *}
|
76 |
||
| 20699 | 77 |
lemma Pls_0_eq [simp, code func]: |
| 20485 | 78 |
"Pls BIT B0 = Pls" |
79 |
unfolding numeral_simps by simp |
|
80 |
||
| 20699 | 81 |
lemma Min_1_eq [simp, code func]: |
| 20485 | 82 |
"Min BIT B1 = Min" |
83 |
unfolding numeral_simps by simp |
|
| 15013 | 84 |
|
85 |
||
| 20485 | 86 |
subsection {* The Functions @{term succ}, @{term pred} and @{term uminus} *}
|
| 15013 | 87 |
|
| 20485 | 88 |
lemma succ_Pls [simp]: |
89 |
"succ Pls = Pls BIT B1" |
|
90 |
unfolding numeral_simps by simp |
|
| 15013 | 91 |
|
| 20485 | 92 |
lemma succ_Min [simp]: |
93 |
"succ Min = Pls" |
|
94 |
unfolding numeral_simps by simp |
|
| 15013 | 95 |
|
| 20485 | 96 |
lemma succ_1 [simp]: |
97 |
"succ (k BIT B1) = succ k BIT B0" |
|
98 |
unfolding numeral_simps by simp |
|
| 15013 | 99 |
|
| 20485 | 100 |
lemma succ_0 [simp]: |
101 |
"succ (k BIT B0) = k BIT B1" |
|
102 |
unfolding numeral_simps by simp |
|
| 15013 | 103 |
|
| 20485 | 104 |
lemma pred_Pls [simp]: |
105 |
"pred Pls = Min" |
|
106 |
unfolding numeral_simps by simp |
|
| 15013 | 107 |
|
| 20485 | 108 |
lemma pred_Min [simp]: |
109 |
"pred Min = Min BIT B0" |
|
110 |
unfolding numeral_simps by simp |
|
| 15013 | 111 |
|
| 20485 | 112 |
lemma pred_1 [simp]: |
113 |
"pred (k BIT B1) = k BIT B0" |
|
114 |
unfolding numeral_simps by simp |
|
| 15013 | 115 |
|
| 20485 | 116 |
lemma pred_0 [simp]: |
117 |
"pred (k BIT B0) = pred k BIT B1" |
|
118 |
unfolding numeral_simps by simp |
|
| 15013 | 119 |
|
| 20485 | 120 |
lemma minus_Pls [simp]: |
121 |
"- Pls = Pls" |
|
122 |
unfolding numeral_simps by simp |
|
| 15013 | 123 |
|
| 20485 | 124 |
lemma minus_Min [simp]: |
125 |
"- Min = Pls BIT B1" |
|
126 |
unfolding numeral_simps by simp |
|
| 15013 | 127 |
|
| 20485 | 128 |
lemma minus_1 [simp]: |
129 |
"- (k BIT B1) = pred (- k) BIT B1" |
|
130 |
unfolding numeral_simps by simp |
|
| 15013 | 131 |
|
| 20485 | 132 |
lemma minus_0 [simp]: |
133 |
"- (k BIT B0) = (- k) BIT B0" |
|
134 |
unfolding numeral_simps by simp |
|
| 15013 | 135 |
|
136 |
||
| 20485 | 137 |
subsection {*
|
138 |
Binary Addition and Multiplication: @{term "op + \<Colon> int \<Rightarrow> int \<Rightarrow> int"}
|
|
139 |
and @{term "op * \<Colon> int \<Rightarrow> int \<Rightarrow> int"}
|
|
140 |
*} |
|
| 15013 | 141 |
|
| 20485 | 142 |
lemma add_Pls [simp]: |
143 |
"Pls + k = k" |
|
144 |
unfolding numeral_simps by simp |
|
| 15013 | 145 |
|
| 20485 | 146 |
lemma add_Min [simp]: |
147 |
"Min + k = pred k" |
|
148 |
unfolding numeral_simps by simp |
|
| 15013 | 149 |
|
| 20485 | 150 |
lemma add_BIT_11 [simp]: |
151 |
"(k BIT B1) + (l BIT B1) = (k + succ l) BIT B0" |
|
152 |
unfolding numeral_simps by simp |
|
| 15013 | 153 |
|
| 20485 | 154 |
lemma add_BIT_10 [simp]: |
155 |
"(k BIT B1) + (l BIT B0) = (k + l) BIT B1" |
|
156 |
unfolding numeral_simps by simp |
|
| 15013 | 157 |
|
| 20485 | 158 |
lemma add_BIT_0 [simp]: |
159 |
"(k BIT B0) + (l BIT b) = (k + l) BIT b" |
|
160 |
unfolding numeral_simps by simp |
|
| 15013 | 161 |
|
| 20485 | 162 |
lemma add_Pls_right [simp]: |
163 |
"k + Pls = k" |
|
164 |
unfolding numeral_simps by simp |
|
| 15013 | 165 |
|
| 20485 | 166 |
lemma add_Min_right [simp]: |
167 |
"k + Min = pred k" |
|
168 |
unfolding numeral_simps by simp |
|
| 15013 | 169 |
|
| 20485 | 170 |
lemma mult_Pls [simp]: |
171 |
"Pls * w = Pls" |
|
172 |
unfolding numeral_simps by simp |
|
| 15013 | 173 |
|
| 20485 | 174 |
lemma mult_Min [simp]: |
175 |
"Min * k = - k" |
|
176 |
unfolding numeral_simps by simp |
|
| 15013 | 177 |
|
| 20485 | 178 |
lemma mult_num1 [simp]: |
179 |
"(k BIT B1) * l = ((k * l) BIT B0) + l" |
|
180 |
unfolding numeral_simps int_distrib by simp |
|
| 15013 | 181 |
|
| 20485 | 182 |
lemma mult_num0 [simp]: |
183 |
"(k BIT B0) * l = (k * l) BIT B0" |
|
184 |
unfolding numeral_simps int_distrib by simp |
|
| 15013 | 185 |
|
186 |
||
187 |
||
| 20485 | 188 |
subsection {* Converting Numerals to Rings: @{term number_of} *}
|
| 15013 | 189 |
|
190 |
axclass number_ring \<subseteq> number, comm_ring_1 |
|
| 20485 | 191 |
number_of_eq: "number_of k = of_int k" |
| 15013 | 192 |
|
193 |
lemma number_of_succ: |
|
| 20485 | 194 |
"number_of (succ k) = (1 + number_of k ::'a::number_ring)" |
195 |
unfolding number_of_eq numeral_simps by simp |
|
| 15013 | 196 |
|
197 |
lemma number_of_pred: |
|
| 20485 | 198 |
"number_of (pred w) = (- 1 + number_of w ::'a::number_ring)" |
199 |
unfolding number_of_eq numeral_simps by simp |
|
| 15013 | 200 |
|
201 |
lemma number_of_minus: |
|
| 20485 | 202 |
"number_of (uminus w) = (- (number_of w)::'a::number_ring)" |
203 |
unfolding number_of_eq numeral_simps by simp |
|
| 15013 | 204 |
|
205 |
lemma number_of_add: |
|
| 20485 | 206 |
"number_of (v + w) = (number_of v + number_of w::'a::number_ring)" |
207 |
unfolding number_of_eq numeral_simps by simp |
|
| 15013 | 208 |
|
209 |
lemma number_of_mult: |
|
| 20485 | 210 |
"number_of (v * w) = (number_of v * number_of w::'a::number_ring)" |
211 |
unfolding number_of_eq numeral_simps by simp |
|
| 15013 | 212 |
|
| 20485 | 213 |
text {*
|
214 |
The correctness of shifting. |
|
215 |
But it doesn't seem to give a measurable speed-up. |
|
216 |
*} |
|
217 |
||
| 15013 | 218 |
lemma double_number_of_BIT: |
| 20485 | 219 |
"(1 + 1) * number_of w = (number_of (w BIT B0) ::'a::number_ring)" |
220 |
unfolding number_of_eq numeral_simps left_distrib by simp |
|
| 15013 | 221 |
|
| 20485 | 222 |
text {*
|
223 |
Converting numerals 0 and 1 to their abstract versions. |
|
224 |
*} |
|
225 |
||
226 |
lemma numeral_0_eq_0 [simp]: |
|
227 |
"Numeral0 = (0::'a::number_ring)" |
|
228 |
unfolding number_of_eq numeral_simps by simp |
|
| 15013 | 229 |
|
| 20485 | 230 |
lemma numeral_1_eq_1 [simp]: |
231 |
"Numeral1 = (1::'a::number_ring)" |
|
232 |
unfolding number_of_eq numeral_simps by simp |
|
| 15013 | 233 |
|
| 20485 | 234 |
text {*
|
235 |
Special-case simplification for small constants. |
|
236 |
*} |
|
| 15013 | 237 |
|
| 20485 | 238 |
text{*
|
239 |
Unary minus for the abstract constant 1. Cannot be inserted |
|
240 |
as a simprule until later: it is @{text number_of_Min} re-oriented!
|
|
241 |
*} |
|
| 15013 | 242 |
|
| 20485 | 243 |
lemma numeral_m1_eq_minus_1: |
244 |
"(-1::'a::number_ring) = - 1" |
|
245 |
unfolding number_of_eq numeral_simps by simp |
|
| 15013 | 246 |
|
| 20485 | 247 |
lemma mult_minus1 [simp]: |
248 |
"-1 * z = -(z::'a::number_ring)" |
|
249 |
unfolding number_of_eq numeral_simps by simp |
|
250 |
||
251 |
lemma mult_minus1_right [simp]: |
|
252 |
"z * -1 = -(z::'a::number_ring)" |
|
253 |
unfolding number_of_eq numeral_simps by simp |
|
| 15013 | 254 |
|
255 |
(*Negation of a coefficient*) |
|
256 |
lemma minus_number_of_mult [simp]: |
|
| 20485 | 257 |
"- (number_of w) * z = number_of (uminus w) * (z::'a::number_ring)" |
258 |
unfolding number_of_eq by simp |
|
259 |
||
260 |
text {* Subtraction *}
|
|
261 |
||
262 |
lemma diff_number_of_eq: |
|
263 |
"number_of v - number_of w = |
|
264 |
(number_of (v + uminus w)::'a::number_ring)" |
|
265 |
unfolding number_of_eq by simp |
|
| 15013 | 266 |
|
| 20485 | 267 |
lemma number_of_Pls: |
268 |
"number_of Pls = (0::'a::number_ring)" |
|
269 |
unfolding number_of_eq numeral_simps by simp |
|
270 |
||
271 |
lemma number_of_Min: |
|
272 |
"number_of Min = (- 1::'a::number_ring)" |
|
273 |
unfolding number_of_eq numeral_simps by simp |
|
274 |
||
275 |
lemma number_of_BIT: |
|
276 |
"number_of(w BIT x) = (case x of B0 => 0 | B1 => (1::'a::number_ring)) |
|
277 |
+ (number_of w) + (number_of w)" |
|
278 |
unfolding number_of_eq numeral_simps by (simp split: bit.split) |
|
| 15013 | 279 |
|
280 |
||
| 20485 | 281 |
subsection {* Equality of Binary Numbers *}
|
| 15013 | 282 |
|
| 20485 | 283 |
text {* First version by Norbert Voelker *}
|
| 15013 | 284 |
|
285 |
lemma eq_number_of_eq: |
|
286 |
"((number_of x::'a::number_ring) = number_of y) = |
|
| 20485 | 287 |
iszero (number_of (x + uminus y) :: 'a)" |
288 |
unfolding iszero_def number_of_add number_of_minus |
|
289 |
by (simp add: compare_rls) |
|
| 15013 | 290 |
|
| 20485 | 291 |
lemma iszero_number_of_Pls: |
292 |
"iszero ((number_of Pls)::'a::number_ring)" |
|
293 |
unfolding iszero_def numeral_0_eq_0 .. |
|
| 15013 | 294 |
|
| 20485 | 295 |
lemma nonzero_number_of_Min: |
296 |
"~ iszero ((number_of Min)::'a::number_ring)" |
|
297 |
unfolding iszero_def numeral_m1_eq_minus_1 by simp |
|
| 15013 | 298 |
|
299 |
||
| 20485 | 300 |
subsection {* Comparisons, for Ordered Rings *}
|
| 15013 | 301 |
|
| 20485 | 302 |
lemma double_eq_0_iff: |
303 |
"(a + a = 0) = (a = (0::'a::ordered_idom))" |
|
| 15013 | 304 |
proof - |
| 20485 | 305 |
have "a + a = (1 + 1) * a" unfolding left_distrib by simp |
| 15013 | 306 |
with zero_less_two [where 'a = 'a] |
307 |
show ?thesis by force |
|
308 |
qed |
|
309 |
||
310 |
lemma le_imp_0_less: |
|
| 20485 | 311 |
assumes le: "0 \<le> z" |
312 |
shows "(0::int) < 1 + z" |
|
| 15013 | 313 |
proof - |
314 |
have "0 \<le> z" . |
|
315 |
also have "... < z + 1" by (rule less_add_one) |
|
316 |
also have "... = 1 + z" by (simp add: add_ac) |
|
317 |
finally show "0 < 1 + z" . |
|
318 |
qed |
|
319 |
||
| 20485 | 320 |
lemma odd_nonzero: |
321 |
"1 + z + z \<noteq> (0::int)"; |
|
| 15013 | 322 |
proof (cases z rule: int_cases) |
323 |
case (nonneg n) |
|
324 |
have le: "0 \<le> z+z" by (simp add: nonneg add_increasing) |
|
325 |
thus ?thesis using le_imp_0_less [OF le] |
|
326 |
by (auto simp add: add_assoc) |
|
327 |
next |
|
328 |
case (neg n) |
|
329 |
show ?thesis |
|
330 |
proof |
|
331 |
assume eq: "1 + z + z = 0" |
|
332 |
have "0 < 1 + (int n + int n)" |
|
333 |
by (simp add: le_imp_0_less add_increasing) |
|
334 |
also have "... = - (1 + z + z)" |
|
335 |
by (simp add: neg add_assoc [symmetric]) |
|
336 |
also have "... = 0" by (simp add: eq) |
|
337 |
finally have "0<0" .. |
|
338 |
thus False by blast |
|
339 |
qed |
|
340 |
qed |
|
341 |
||
| 20485 | 342 |
text {* The premise involving @{term Ints} prevents @{term "a = 1/2"}. *}
|
| 15013 | 343 |
|
| 20485 | 344 |
lemma Ints_odd_nonzero: |
345 |
assumes in_Ints: "a \<in> Ints" |
|
346 |
shows "1 + a + a \<noteq> (0::'a::ordered_idom)" |
|
347 |
proof - |
|
348 |
from in_Ints have "a \<in> range of_int" unfolding Ints_def [symmetric] . |
|
| 15013 | 349 |
then obtain z where a: "a = of_int z" .. |
350 |
show ?thesis |
|
351 |
proof |
|
352 |
assume eq: "1 + a + a = 0" |
|
353 |
hence "of_int (1 + z + z) = (of_int 0 :: 'a)" by (simp add: a) |
|
354 |
hence "1 + z + z = 0" by (simp only: of_int_eq_iff) |
|
355 |
with odd_nonzero show False by blast |
|
356 |
qed |
|
357 |
qed |
|
358 |
||
| 20485 | 359 |
lemma Ints_number_of: |
360 |
"(number_of w :: 'a::number_ring) \<in> Ints" |
|
361 |
unfolding number_of_eq Ints_def by simp |
|
| 15013 | 362 |
|
363 |
lemma iszero_number_of_BIT: |
|
| 20485 | 364 |
"iszero (number_of (w BIT x)::'a) = |
365 |
(x = B0 \<and> iszero (number_of w::'a::{ordered_idom,number_ring}))"
|
|
366 |
by (simp add: iszero_def number_of_eq numeral_simps double_eq_0_iff |
|
367 |
Ints_odd_nonzero Ints_def split: bit.split) |
|
| 15013 | 368 |
|
369 |
lemma iszero_number_of_0: |
|
| 20485 | 370 |
"iszero (number_of (w BIT B0) :: 'a::{ordered_idom,number_ring}) =
|
371 |
iszero (number_of w :: 'a)" |
|
372 |
by (simp only: iszero_number_of_BIT simp_thms) |
|
| 15013 | 373 |
|
374 |
lemma iszero_number_of_1: |
|
| 20485 | 375 |
"~ iszero (number_of (w BIT B1)::'a::{ordered_idom,number_ring})"
|
376 |
by (simp add: iszero_number_of_BIT) |
|
| 15013 | 377 |
|
378 |
||
| 20485 | 379 |
subsection {* The Less-Than Relation *}
|
| 15013 | 380 |
|
381 |
lemma less_number_of_eq_neg: |
|
| 20485 | 382 |
"((number_of x::'a::{ordered_idom,number_ring}) < number_of y)
|
383 |
= neg (number_of (x + uminus y) :: 'a)" |
|
| 15013 | 384 |
apply (subst less_iff_diff_less_0) |
385 |
apply (simp add: neg_def diff_minus number_of_add number_of_minus) |
|
386 |
done |
|
387 |
||
| 20485 | 388 |
text {*
|
389 |
If @{term Numeral0} is rewritten to 0 then this rule can't be applied:
|
|
390 |
@{term Numeral0} IS @{term "number_of Pls"}
|
|
391 |
*} |
|
392 |
||
| 15013 | 393 |
lemma not_neg_number_of_Pls: |
| 20485 | 394 |
"~ neg (number_of Pls ::'a::{ordered_idom,number_ring})"
|
395 |
by (simp add: neg_def numeral_0_eq_0) |
|
| 15013 | 396 |
|
397 |
lemma neg_number_of_Min: |
|
| 20485 | 398 |
"neg (number_of Min ::'a::{ordered_idom,number_ring})"
|
399 |
by (simp add: neg_def zero_less_one numeral_m1_eq_minus_1) |
|
| 15013 | 400 |
|
| 20485 | 401 |
lemma double_less_0_iff: |
402 |
"(a + a < 0) = (a < (0::'a::ordered_idom))" |
|
| 15013 | 403 |
proof - |
404 |
have "(a + a < 0) = ((1+1)*a < 0)" by (simp add: left_distrib) |
|
405 |
also have "... = (a < 0)" |
|
406 |
by (simp add: mult_less_0_iff zero_less_two |
|
407 |
order_less_not_sym [OF zero_less_two]) |
|
408 |
finally show ?thesis . |
|
409 |
qed |
|
410 |
||
| 20485 | 411 |
lemma odd_less_0: |
412 |
"(1 + z + z < 0) = (z < (0::int))"; |
|
| 15013 | 413 |
proof (cases z rule: int_cases) |
414 |
case (nonneg n) |
|
415 |
thus ?thesis by (simp add: linorder_not_less add_assoc add_increasing |
|
416 |
le_imp_0_less [THEN order_less_imp_le]) |
|
417 |
next |
|
418 |
case (neg n) |
|
419 |
thus ?thesis by (simp del: int_Suc |
|
| 20485 | 420 |
add: int_Suc0_eq_1 [symmetric] zadd_int compare_rls) |
| 15013 | 421 |
qed |
422 |
||
| 20485 | 423 |
text {* The premise involving @{term Ints} prevents @{term "a = 1/2"}. *}
|
424 |
||
| 15013 | 425 |
lemma Ints_odd_less_0: |
| 20485 | 426 |
assumes in_Ints: "a \<in> Ints" |
427 |
shows "(1 + a + a < 0) = (a < (0::'a::ordered_idom))"; |
|
428 |
proof - |
|
429 |
from in_Ints have "a \<in> range of_int" unfolding Ints_def [symmetric] . |
|
| 15013 | 430 |
then obtain z where a: "a = of_int z" .. |
431 |
hence "((1::'a) + a + a < 0) = (of_int (1 + z + z) < (of_int 0 :: 'a))" |
|
432 |
by (simp add: a) |
|
433 |
also have "... = (z < 0)" by (simp only: of_int_less_iff odd_less_0) |
|
434 |
also have "... = (a < 0)" by (simp add: a) |
|
435 |
finally show ?thesis . |
|
436 |
qed |
|
437 |
||
438 |
lemma neg_number_of_BIT: |
|
| 20485 | 439 |
"neg (number_of (w BIT x)::'a) = |
440 |
neg (number_of w :: 'a::{ordered_idom,number_ring})"
|
|
441 |
by (simp add: neg_def number_of_eq numeral_simps double_less_0_iff |
|
442 |
Ints_odd_less_0 Ints_def split: bit.split) |
|
| 15013 | 443 |
|
| 20596 | 444 |
|
| 20485 | 445 |
text {* Less-Than or Equals *}
|
446 |
||
447 |
text {* Reduces @{term "a\<le>b"} to @{term "~ (b<a)"} for ALL numerals. *}
|
|
| 15013 | 448 |
|
449 |
lemmas le_number_of_eq_not_less = |
|
| 20485 | 450 |
linorder_not_less [of "number_of w" "number_of v", symmetric, |
451 |
standard] |
|
| 15013 | 452 |
|
453 |
lemma le_number_of_eq: |
|
454 |
"((number_of x::'a::{ordered_idom,number_ring}) \<le> number_of y)
|
|
| 20485 | 455 |
= (~ (neg (number_of (y + uminus x) :: 'a)))" |
| 15013 | 456 |
by (simp add: le_number_of_eq_not_less less_number_of_eq_neg) |
457 |
||
458 |
||
| 20485 | 459 |
text {* Absolute value (@{term abs}) *}
|
| 15013 | 460 |
|
461 |
lemma abs_number_of: |
|
| 20485 | 462 |
"abs(number_of x::'a::{ordered_idom,number_ring}) =
|
463 |
(if number_of x < (0::'a) then -number_of x else number_of x)" |
|
464 |
by (simp add: abs_if) |
|
| 15013 | 465 |
|
466 |
||
| 20485 | 467 |
text {* Re-orientation of the equation nnn=x *}
|
| 15013 | 468 |
|
| 20485 | 469 |
lemma number_of_reorient: |
470 |
"(number_of w = x) = (x = number_of w)" |
|
471 |
by auto |
|
| 15013 | 472 |
|
473 |
||
| 20485 | 474 |
subsection {* Simplification of arithmetic operations on integer constants. *}
|
| 15013 | 475 |
|
| 20900 | 476 |
lemmas arith_extra_simps [standard] = |
| 20485 | 477 |
number_of_add [symmetric] |
478 |
number_of_minus [symmetric] numeral_m1_eq_minus_1 [symmetric] |
|
479 |
number_of_mult [symmetric] |
|
480 |
diff_number_of_eq abs_number_of |
|
481 |
||
482 |
text {*
|
|
483 |
For making a minimal simpset, one must include these default simprules. |
|
484 |
Also include @{text simp_thms}.
|
|
485 |
*} |
|
| 15013 | 486 |
|
| 20485 | 487 |
lemmas arith_simps = |
488 |
bit.distinct |
|
489 |
Pls_0_eq Min_1_eq |
|
490 |
pred_Pls pred_Min pred_1 pred_0 |
|
491 |
succ_Pls succ_Min succ_1 succ_0 |
|
492 |
add_Pls add_Min add_BIT_0 add_BIT_10 add_BIT_11 |
|
493 |
minus_Pls minus_Min minus_1 minus_0 |
|
494 |
mult_Pls mult_Min mult_num1 mult_num0 |
|
495 |
add_Pls_right add_Min_right |
|
496 |
abs_zero abs_one arith_extra_simps |
|
| 15013 | 497 |
|
| 20485 | 498 |
text {* Simplification of relational operations *}
|
| 15013 | 499 |
|
| 20485 | 500 |
lemmas rel_simps = |
501 |
eq_number_of_eq iszero_number_of_Pls nonzero_number_of_Min |
|
502 |
iszero_number_of_0 iszero_number_of_1 |
|
503 |
less_number_of_eq_neg |
|
504 |
not_neg_number_of_Pls not_neg_0 not_neg_1 not_iszero_1 |
|
505 |
neg_number_of_Min neg_number_of_BIT |
|
506 |
le_number_of_eq |
|
507 |
||
508 |
declare arith_extra_simps [simp] |
|
509 |
declare rel_simps [simp] |
|
| 15013 | 510 |
|
511 |
||
| 20485 | 512 |
subsection {* Simplification of arithmetic when nested to the right. *}
|
| 15013 | 513 |
|
514 |
lemma add_number_of_left [simp]: |
|
| 20485 | 515 |
"number_of v + (number_of w + z) = |
516 |
(number_of(v + w) + z::'a::number_ring)" |
|
517 |
by (simp add: add_assoc [symmetric]) |
|
| 15013 | 518 |
|
519 |
lemma mult_number_of_left [simp]: |
|
| 20485 | 520 |
"number_of v * (number_of w * z) = |
521 |
(number_of(v * w) * z::'a::number_ring)" |
|
522 |
by (simp add: mult_assoc [symmetric]) |
|
| 15013 | 523 |
|
524 |
lemma add_number_of_diff1: |
|
| 20485 | 525 |
"number_of v + (number_of w - c) = |
526 |
number_of(v + w) - (c::'a::number_ring)" |
|
527 |
by (simp add: diff_minus add_number_of_left) |
|
| 15013 | 528 |
|
| 20485 | 529 |
lemma add_number_of_diff2 [simp]: |
530 |
"number_of v + (c - number_of w) = |
|
531 |
number_of (v + uminus w) + (c::'a::number_ring)" |
|
| 15013 | 532 |
apply (subst diff_number_of_eq [symmetric]) |
533 |
apply (simp only: compare_rls) |
|
534 |
done |
|
535 |
||
| 19380 | 536 |
|
| 20500 | 537 |
hide (open) const Pls Min B0 B1 succ pred |
| 19380 | 538 |
|
| 15013 | 539 |
end |