author | haftmann |
Sat, 19 May 2007 11:33:30 +0200 | |
changeset 23024 | 70435ffe077d |
parent 22997 | d4f3b015b50b |
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 |
22801 | 10 |
imports IntDef |
22997 | 11 |
uses ("../Tools/numeral_syntax.ML") |
15131 | 12 |
begin |
15013 | 13 |
|
22801 | 14 |
subsection {* Binary representation *} |
15 |
||
20485 | 16 |
text {* |
17 |
This formalization defines binary arithmetic in terms of the integers |
|
18 |
rather than using a datatype. This avoids multiple representations (leading |
|
19 |
zeroes, etc.) See @{text "ZF/Integ/twos-compl.ML"}, function @{text |
|
20 |
int_of_binary}, for the numerical interpretation. |
|
15013 | 21 |
|
20485 | 22 |
The representation expects that @{text "(m mod 2)"} is 0 or 1, |
23 |
even if m is negative; |
|
24 |
For instance, @{text "-5 div 2 = -3"} and @{text "-5 mod 2 = 1"}; thus |
|
25 |
@{text "-5 = (-3)*2 + 1"}. |
|
15013 | 26 |
*} |
27 |
||
22801 | 28 |
datatype bit = B0 | B1 |
29 |
||
20485 | 30 |
text{* |
22801 | 31 |
Type @{typ bit} avoids the use of type @{typ bool}, which would make |
20485 | 32 |
all of the rewrite rules higher-order. |
33 |
*} |
|
15013 | 34 |
|
21779 | 35 |
definition |
21820
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
haftmann
parents:
21779
diff
changeset
|
36 |
Pls :: int where |
22845 | 37 |
[code func del]:"Pls = 0" |
21779 | 38 |
|
39 |
definition |
|
21820
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
haftmann
parents:
21779
diff
changeset
|
40 |
Min :: int where |
22845 | 41 |
[code func del]:"Min = - 1" |
21779 | 42 |
|
43 |
definition |
|
44 |
Bit :: "int \<Rightarrow> bit \<Rightarrow> int" (infixl "BIT" 90) where |
|
22845 | 45 |
[code func del]: "k BIT b = (case b of B0 \<Rightarrow> 0 | B1 \<Rightarrow> 1) + k + k" |
15013 | 46 |
|
22473 | 47 |
class number = type + -- {* for numeric types: nat, int, real, \dots *} |
21820
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
haftmann
parents:
21779
diff
changeset
|
48 |
fixes number_of :: "int \<Rightarrow> 'a" |
15013 | 49 |
|
50 |
syntax |
|
20485 | 51 |
"_Numeral" :: "num_const \<Rightarrow> 'a" ("_") |
15013 | 52 |
|
22997 | 53 |
use "../Tools/numeral_syntax.ML" |
15013 | 54 |
setup NumeralSyntax.setup |
55 |
||
19380 | 56 |
abbreviation |
20485 | 57 |
"Numeral0 \<equiv> number_of Pls" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20900
diff
changeset
|
58 |
|
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20900
diff
changeset
|
59 |
abbreviation |
20485 | 60 |
"Numeral1 \<equiv> number_of (Pls BIT B1)" |
15013 | 61 |
|
20485 | 62 |
lemma Let_number_of [simp]: "Let (number_of v) f = f (number_of v)" |
15013 | 63 |
-- {* Unfold all @{text let}s involving constants *} |
20485 | 64 |
unfolding Let_def .. |
65 |
||
66 |
lemma Let_0 [simp]: "Let 0 f = f 0" |
|
67 |
unfolding Let_def .. |
|
68 |
||
69 |
lemma Let_1 [simp]: "Let 1 f = f 1" |
|
70 |
unfolding Let_def .. |
|
15013 | 71 |
|
20485 | 72 |
definition |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20900
diff
changeset
|
73 |
succ :: "int \<Rightarrow> int" where |
22845 | 74 |
[code func del]: "succ k = k + 1" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20900
diff
changeset
|
75 |
|
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20900
diff
changeset
|
76 |
definition |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20900
diff
changeset
|
77 |
pred :: "int \<Rightarrow> int" where |
22845 | 78 |
[code func del]: "pred k = k - 1" |
20485 | 79 |
|
22801 | 80 |
lemmas |
81 |
max_number_of [simp] = max_def |
|
82 |
[of "number_of u" "number_of v", standard, simp] |
|
83 |
and |
|
84 |
min_number_of [simp] = min_def |
|
85 |
[of "number_of u" "number_of v", standard, simp] |
|
22187 | 86 |
-- {* unfolding @{text minx} and @{text max} on numerals *} |
87 |
||
20485 | 88 |
lemmas numeral_simps = |
89 |
succ_def pred_def Pls_def Min_def Bit_def |
|
15013 | 90 |
|
20485 | 91 |
text {* Removal of leading zeroes *} |
92 |
||
22801 | 93 |
lemma Pls_0_eq [simp, normal post]: |
20485 | 94 |
"Pls BIT B0 = Pls" |
95 |
unfolding numeral_simps by simp |
|
96 |
||
22801 | 97 |
lemma Min_1_eq [simp, normal post]: |
20485 | 98 |
"Min BIT B1 = Min" |
99 |
unfolding numeral_simps by simp |
|
15013 | 100 |
|
101 |
||
20485 | 102 |
subsection {* The Functions @{term succ}, @{term pred} and @{term uminus} *} |
15013 | 103 |
|
20485 | 104 |
lemma succ_Pls [simp]: |
105 |
"succ Pls = Pls BIT B1" |
|
106 |
unfolding numeral_simps by simp |
|
15013 | 107 |
|
20485 | 108 |
lemma succ_Min [simp]: |
109 |
"succ Min = Pls" |
|
110 |
unfolding numeral_simps by simp |
|
15013 | 111 |
|
20485 | 112 |
lemma succ_1 [simp]: |
113 |
"succ (k BIT B1) = succ k BIT B0" |
|
114 |
unfolding numeral_simps by simp |
|
15013 | 115 |
|
20485 | 116 |
lemma succ_0 [simp]: |
117 |
"succ (k BIT B0) = k BIT B1" |
|
118 |
unfolding numeral_simps by simp |
|
15013 | 119 |
|
20485 | 120 |
lemma pred_Pls [simp]: |
121 |
"pred Pls = Min" |
|
122 |
unfolding numeral_simps by simp |
|
15013 | 123 |
|
20485 | 124 |
lemma pred_Min [simp]: |
125 |
"pred Min = Min BIT B0" |
|
126 |
unfolding numeral_simps by simp |
|
15013 | 127 |
|
20485 | 128 |
lemma pred_1 [simp]: |
129 |
"pred (k BIT B1) = k BIT B0" |
|
130 |
unfolding numeral_simps by simp |
|
15013 | 131 |
|
20485 | 132 |
lemma pred_0 [simp]: |
133 |
"pred (k BIT B0) = pred k BIT B1" |
|
134 |
unfolding numeral_simps by simp |
|
15013 | 135 |
|
20485 | 136 |
lemma minus_Pls [simp]: |
137 |
"- Pls = Pls" |
|
138 |
unfolding numeral_simps by simp |
|
15013 | 139 |
|
20485 | 140 |
lemma minus_Min [simp]: |
141 |
"- Min = Pls BIT B1" |
|
142 |
unfolding numeral_simps by simp |
|
15013 | 143 |
|
20485 | 144 |
lemma minus_1 [simp]: |
145 |
"- (k BIT B1) = pred (- k) BIT B1" |
|
146 |
unfolding numeral_simps by simp |
|
15013 | 147 |
|
20485 | 148 |
lemma minus_0 [simp]: |
149 |
"- (k BIT B0) = (- k) BIT B0" |
|
150 |
unfolding numeral_simps by simp |
|
15013 | 151 |
|
152 |
||
20485 | 153 |
subsection {* |
154 |
Binary Addition and Multiplication: @{term "op + \<Colon> int \<Rightarrow> int \<Rightarrow> int"} |
|
155 |
and @{term "op * \<Colon> int \<Rightarrow> int \<Rightarrow> int"} |
|
156 |
*} |
|
15013 | 157 |
|
20485 | 158 |
lemma add_Pls [simp]: |
159 |
"Pls + k = k" |
|
160 |
unfolding numeral_simps by simp |
|
15013 | 161 |
|
20485 | 162 |
lemma add_Min [simp]: |
163 |
"Min + k = pred k" |
|
164 |
unfolding numeral_simps by simp |
|
15013 | 165 |
|
20485 | 166 |
lemma add_BIT_11 [simp]: |
167 |
"(k BIT B1) + (l BIT B1) = (k + succ l) BIT B0" |
|
168 |
unfolding numeral_simps by simp |
|
15013 | 169 |
|
20485 | 170 |
lemma add_BIT_10 [simp]: |
171 |
"(k BIT B1) + (l BIT B0) = (k + l) BIT B1" |
|
172 |
unfolding numeral_simps by simp |
|
15013 | 173 |
|
20485 | 174 |
lemma add_BIT_0 [simp]: |
175 |
"(k BIT B0) + (l BIT b) = (k + l) BIT b" |
|
176 |
unfolding numeral_simps by simp |
|
15013 | 177 |
|
20485 | 178 |
lemma add_Pls_right [simp]: |
179 |
"k + Pls = k" |
|
180 |
unfolding numeral_simps by simp |
|
15013 | 181 |
|
20485 | 182 |
lemma add_Min_right [simp]: |
183 |
"k + Min = pred k" |
|
184 |
unfolding numeral_simps by simp |
|
15013 | 185 |
|
20485 | 186 |
lemma mult_Pls [simp]: |
187 |
"Pls * w = Pls" |
|
188 |
unfolding numeral_simps by simp |
|
15013 | 189 |
|
20485 | 190 |
lemma mult_Min [simp]: |
191 |
"Min * k = - k" |
|
192 |
unfolding numeral_simps by simp |
|
15013 | 193 |
|
20485 | 194 |
lemma mult_num1 [simp]: |
195 |
"(k BIT B1) * l = ((k * l) BIT B0) + l" |
|
196 |
unfolding numeral_simps int_distrib by simp |
|
15013 | 197 |
|
20485 | 198 |
lemma mult_num0 [simp]: |
199 |
"(k BIT B0) * l = (k * l) BIT B0" |
|
200 |
unfolding numeral_simps int_distrib by simp |
|
15013 | 201 |
|
202 |
||
203 |
||
20485 | 204 |
subsection {* Converting Numerals to Rings: @{term number_of} *} |
15013 | 205 |
|
206 |
axclass number_ring \<subseteq> number, comm_ring_1 |
|
20485 | 207 |
number_of_eq: "number_of k = of_int k" |
15013 | 208 |
|
22801 | 209 |
text {* self-embedding of the intergers *} |
210 |
||
211 |
instance int :: number_ring |
|
212 |
int_number_of_def: "number_of w \<equiv> of_int w" |
|
213 |
by intro_classes (simp only: int_number_of_def) |
|
214 |
||
22845 | 215 |
lemmas [code func del] = int_number_of_def |
22801 | 216 |
|
217 |
lemma number_of_is_id: |
|
218 |
"number_of (k::int) = k" |
|
219 |
unfolding int_number_of_def by simp |
|
220 |
||
15013 | 221 |
lemma number_of_succ: |
20485 | 222 |
"number_of (succ k) = (1 + number_of k ::'a::number_ring)" |
223 |
unfolding number_of_eq numeral_simps by simp |
|
15013 | 224 |
|
225 |
lemma number_of_pred: |
|
20485 | 226 |
"number_of (pred w) = (- 1 + number_of w ::'a::number_ring)" |
227 |
unfolding number_of_eq numeral_simps by simp |
|
15013 | 228 |
|
229 |
lemma number_of_minus: |
|
20485 | 230 |
"number_of (uminus w) = (- (number_of w)::'a::number_ring)" |
231 |
unfolding number_of_eq numeral_simps by simp |
|
15013 | 232 |
|
233 |
lemma number_of_add: |
|
20485 | 234 |
"number_of (v + w) = (number_of v + number_of w::'a::number_ring)" |
235 |
unfolding number_of_eq numeral_simps by simp |
|
15013 | 236 |
|
237 |
lemma number_of_mult: |
|
20485 | 238 |
"number_of (v * w) = (number_of v * number_of w::'a::number_ring)" |
239 |
unfolding number_of_eq numeral_simps by simp |
|
15013 | 240 |
|
20485 | 241 |
text {* |
242 |
The correctness of shifting. |
|
243 |
But it doesn't seem to give a measurable speed-up. |
|
244 |
*} |
|
245 |
||
15013 | 246 |
lemma double_number_of_BIT: |
20485 | 247 |
"(1 + 1) * number_of w = (number_of (w BIT B0) ::'a::number_ring)" |
248 |
unfolding number_of_eq numeral_simps left_distrib by simp |
|
15013 | 249 |
|
20485 | 250 |
text {* |
251 |
Converting numerals 0 and 1 to their abstract versions. |
|
252 |
*} |
|
253 |
||
254 |
lemma numeral_0_eq_0 [simp]: |
|
255 |
"Numeral0 = (0::'a::number_ring)" |
|
256 |
unfolding number_of_eq numeral_simps by simp |
|
15013 | 257 |
|
20485 | 258 |
lemma numeral_1_eq_1 [simp]: |
259 |
"Numeral1 = (1::'a::number_ring)" |
|
260 |
unfolding number_of_eq numeral_simps by simp |
|
15013 | 261 |
|
20485 | 262 |
text {* |
263 |
Special-case simplification for small constants. |
|
264 |
*} |
|
15013 | 265 |
|
20485 | 266 |
text{* |
267 |
Unary minus for the abstract constant 1. Cannot be inserted |
|
268 |
as a simprule until later: it is @{text number_of_Min} re-oriented! |
|
269 |
*} |
|
15013 | 270 |
|
20485 | 271 |
lemma numeral_m1_eq_minus_1: |
272 |
"(-1::'a::number_ring) = - 1" |
|
273 |
unfolding number_of_eq numeral_simps by simp |
|
15013 | 274 |
|
20485 | 275 |
lemma mult_minus1 [simp]: |
276 |
"-1 * z = -(z::'a::number_ring)" |
|
277 |
unfolding number_of_eq numeral_simps by simp |
|
278 |
||
279 |
lemma mult_minus1_right [simp]: |
|
280 |
"z * -1 = -(z::'a::number_ring)" |
|
281 |
unfolding number_of_eq numeral_simps by simp |
|
15013 | 282 |
|
283 |
(*Negation of a coefficient*) |
|
284 |
lemma minus_number_of_mult [simp]: |
|
20485 | 285 |
"- (number_of w) * z = number_of (uminus w) * (z::'a::number_ring)" |
286 |
unfolding number_of_eq by simp |
|
287 |
||
288 |
text {* Subtraction *} |
|
289 |
||
290 |
lemma diff_number_of_eq: |
|
291 |
"number_of v - number_of w = |
|
292 |
(number_of (v + uminus w)::'a::number_ring)" |
|
293 |
unfolding number_of_eq by simp |
|
15013 | 294 |
|
20485 | 295 |
lemma number_of_Pls: |
296 |
"number_of Pls = (0::'a::number_ring)" |
|
297 |
unfolding number_of_eq numeral_simps by simp |
|
298 |
||
299 |
lemma number_of_Min: |
|
300 |
"number_of Min = (- 1::'a::number_ring)" |
|
301 |
unfolding number_of_eq numeral_simps by simp |
|
302 |
||
303 |
lemma number_of_BIT: |
|
304 |
"number_of(w BIT x) = (case x of B0 => 0 | B1 => (1::'a::number_ring)) |
|
305 |
+ (number_of w) + (number_of w)" |
|
306 |
unfolding number_of_eq numeral_simps by (simp split: bit.split) |
|
15013 | 307 |
|
308 |
||
20485 | 309 |
subsection {* Equality of Binary Numbers *} |
15013 | 310 |
|
20485 | 311 |
text {* First version by Norbert Voelker *} |
15013 | 312 |
|
313 |
lemma eq_number_of_eq: |
|
314 |
"((number_of x::'a::number_ring) = number_of y) = |
|
20485 | 315 |
iszero (number_of (x + uminus y) :: 'a)" |
316 |
unfolding iszero_def number_of_add number_of_minus |
|
317 |
by (simp add: compare_rls) |
|
15013 | 318 |
|
20485 | 319 |
lemma iszero_number_of_Pls: |
320 |
"iszero ((number_of Pls)::'a::number_ring)" |
|
321 |
unfolding iszero_def numeral_0_eq_0 .. |
|
15013 | 322 |
|
20485 | 323 |
lemma nonzero_number_of_Min: |
324 |
"~ iszero ((number_of Min)::'a::number_ring)" |
|
325 |
unfolding iszero_def numeral_m1_eq_minus_1 by simp |
|
15013 | 326 |
|
327 |
||
20485 | 328 |
subsection {* Comparisons, for Ordered Rings *} |
15013 | 329 |
|
20485 | 330 |
lemma double_eq_0_iff: |
331 |
"(a + a = 0) = (a = (0::'a::ordered_idom))" |
|
15013 | 332 |
proof - |
20485 | 333 |
have "a + a = (1 + 1) * a" unfolding left_distrib by simp |
15013 | 334 |
with zero_less_two [where 'a = 'a] |
335 |
show ?thesis by force |
|
336 |
qed |
|
337 |
||
338 |
lemma le_imp_0_less: |
|
20485 | 339 |
assumes le: "0 \<le> z" |
340 |
shows "(0::int) < 1 + z" |
|
15013 | 341 |
proof - |
342 |
have "0 \<le> z" . |
|
343 |
also have "... < z + 1" by (rule less_add_one) |
|
344 |
also have "... = 1 + z" by (simp add: add_ac) |
|
345 |
finally show "0 < 1 + z" . |
|
346 |
qed |
|
347 |
||
20485 | 348 |
lemma odd_nonzero: |
349 |
"1 + z + z \<noteq> (0::int)"; |
|
15013 | 350 |
proof (cases z rule: int_cases) |
351 |
case (nonneg n) |
|
352 |
have le: "0 \<le> z+z" by (simp add: nonneg add_increasing) |
|
353 |
thus ?thesis using le_imp_0_less [OF le] |
|
354 |
by (auto simp add: add_assoc) |
|
355 |
next |
|
356 |
case (neg n) |
|
357 |
show ?thesis |
|
358 |
proof |
|
359 |
assume eq: "1 + z + z = 0" |
|
360 |
have "0 < 1 + (int n + int n)" |
|
361 |
by (simp add: le_imp_0_less add_increasing) |
|
362 |
also have "... = - (1 + z + z)" |
|
363 |
by (simp add: neg add_assoc [symmetric]) |
|
364 |
also have "... = 0" by (simp add: eq) |
|
365 |
finally have "0<0" .. |
|
366 |
thus False by blast |
|
367 |
qed |
|
368 |
qed |
|
369 |
||
20485 | 370 |
text {* The premise involving @{term Ints} prevents @{term "a = 1/2"}. *} |
15013 | 371 |
|
22911
2f5e8d70a179
new axclass ring_char_0 for rings with characteristic 0, used for of_int_eq_iff and related lemmas
huffman
parents:
22845
diff
changeset
|
372 |
lemma Ints_double_eq_0_iff: |
2f5e8d70a179
new axclass ring_char_0 for rings with characteristic 0, used for of_int_eq_iff and related lemmas
huffman
parents:
22845
diff
changeset
|
373 |
assumes in_Ints: "a \<in> Ints" |
2f5e8d70a179
new axclass ring_char_0 for rings with characteristic 0, used for of_int_eq_iff and related lemmas
huffman
parents:
22845
diff
changeset
|
374 |
shows "(a + a = 0) = (a = (0::'a::ring_char_0))" |
2f5e8d70a179
new axclass ring_char_0 for rings with characteristic 0, used for of_int_eq_iff and related lemmas
huffman
parents:
22845
diff
changeset
|
375 |
proof - |
2f5e8d70a179
new axclass ring_char_0 for rings with characteristic 0, used for of_int_eq_iff and related lemmas
huffman
parents:
22845
diff
changeset
|
376 |
from in_Ints have "a \<in> range of_int" unfolding Ints_def [symmetric] . |
2f5e8d70a179
new axclass ring_char_0 for rings with characteristic 0, used for of_int_eq_iff and related lemmas
huffman
parents:
22845
diff
changeset
|
377 |
then obtain z where a: "a = of_int z" .. |
2f5e8d70a179
new axclass ring_char_0 for rings with characteristic 0, used for of_int_eq_iff and related lemmas
huffman
parents:
22845
diff
changeset
|
378 |
show ?thesis |
2f5e8d70a179
new axclass ring_char_0 for rings with characteristic 0, used for of_int_eq_iff and related lemmas
huffman
parents:
22845
diff
changeset
|
379 |
proof |
2f5e8d70a179
new axclass ring_char_0 for rings with characteristic 0, used for of_int_eq_iff and related lemmas
huffman
parents:
22845
diff
changeset
|
380 |
assume "a = 0" |
2f5e8d70a179
new axclass ring_char_0 for rings with characteristic 0, used for of_int_eq_iff and related lemmas
huffman
parents:
22845
diff
changeset
|
381 |
thus "a + a = 0" by simp |
2f5e8d70a179
new axclass ring_char_0 for rings with characteristic 0, used for of_int_eq_iff and related lemmas
huffman
parents:
22845
diff
changeset
|
382 |
next |
2f5e8d70a179
new axclass ring_char_0 for rings with characteristic 0, used for of_int_eq_iff and related lemmas
huffman
parents:
22845
diff
changeset
|
383 |
assume eq: "a + a = 0" |
2f5e8d70a179
new axclass ring_char_0 for rings with characteristic 0, used for of_int_eq_iff and related lemmas
huffman
parents:
22845
diff
changeset
|
384 |
hence "of_int (z + z) = (of_int 0 :: 'a)" by (simp add: a) |
2f5e8d70a179
new axclass ring_char_0 for rings with characteristic 0, used for of_int_eq_iff and related lemmas
huffman
parents:
22845
diff
changeset
|
385 |
hence "z + z = 0" by (simp only: of_int_eq_iff) |
2f5e8d70a179
new axclass ring_char_0 for rings with characteristic 0, used for of_int_eq_iff and related lemmas
huffman
parents:
22845
diff
changeset
|
386 |
hence "z = 0" by (simp only: double_eq_0_iff) |
2f5e8d70a179
new axclass ring_char_0 for rings with characteristic 0, used for of_int_eq_iff and related lemmas
huffman
parents:
22845
diff
changeset
|
387 |
thus "a = 0" by (simp add: a) |
2f5e8d70a179
new axclass ring_char_0 for rings with characteristic 0, used for of_int_eq_iff and related lemmas
huffman
parents:
22845
diff
changeset
|
388 |
qed |
2f5e8d70a179
new axclass ring_char_0 for rings with characteristic 0, used for of_int_eq_iff and related lemmas
huffman
parents:
22845
diff
changeset
|
389 |
qed |
2f5e8d70a179
new axclass ring_char_0 for rings with characteristic 0, used for of_int_eq_iff and related lemmas
huffman
parents:
22845
diff
changeset
|
390 |
|
20485 | 391 |
lemma Ints_odd_nonzero: |
392 |
assumes in_Ints: "a \<in> Ints" |
|
22911
2f5e8d70a179
new axclass ring_char_0 for rings with characteristic 0, used for of_int_eq_iff and related lemmas
huffman
parents:
22845
diff
changeset
|
393 |
shows "1 + a + a \<noteq> (0::'a::ring_char_0)" |
20485 | 394 |
proof - |
395 |
from in_Ints have "a \<in> range of_int" unfolding Ints_def [symmetric] . |
|
15013 | 396 |
then obtain z where a: "a = of_int z" .. |
397 |
show ?thesis |
|
398 |
proof |
|
399 |
assume eq: "1 + a + a = 0" |
|
400 |
hence "of_int (1 + z + z) = (of_int 0 :: 'a)" by (simp add: a) |
|
401 |
hence "1 + z + z = 0" by (simp only: of_int_eq_iff) |
|
402 |
with odd_nonzero show False by blast |
|
403 |
qed |
|
404 |
qed |
|
405 |
||
20485 | 406 |
lemma Ints_number_of: |
407 |
"(number_of w :: 'a::number_ring) \<in> Ints" |
|
408 |
unfolding number_of_eq Ints_def by simp |
|
15013 | 409 |
|
410 |
lemma iszero_number_of_BIT: |
|
20485 | 411 |
"iszero (number_of (w BIT x)::'a) = |
22911
2f5e8d70a179
new axclass ring_char_0 for rings with characteristic 0, used for of_int_eq_iff and related lemmas
huffman
parents:
22845
diff
changeset
|
412 |
(x = B0 \<and> iszero (number_of w::'a::{ring_char_0,number_ring}))" |
2f5e8d70a179
new axclass ring_char_0 for rings with characteristic 0, used for of_int_eq_iff and related lemmas
huffman
parents:
22845
diff
changeset
|
413 |
by (simp add: iszero_def number_of_eq numeral_simps Ints_double_eq_0_iff |
20485 | 414 |
Ints_odd_nonzero Ints_def split: bit.split) |
15013 | 415 |
|
416 |
lemma iszero_number_of_0: |
|
22911
2f5e8d70a179
new axclass ring_char_0 for rings with characteristic 0, used for of_int_eq_iff and related lemmas
huffman
parents:
22845
diff
changeset
|
417 |
"iszero (number_of (w BIT B0) :: 'a::{ring_char_0,number_ring}) = |
20485 | 418 |
iszero (number_of w :: 'a)" |
419 |
by (simp only: iszero_number_of_BIT simp_thms) |
|
15013 | 420 |
|
421 |
lemma iszero_number_of_1: |
|
22911
2f5e8d70a179
new axclass ring_char_0 for rings with characteristic 0, used for of_int_eq_iff and related lemmas
huffman
parents:
22845
diff
changeset
|
422 |
"~ iszero (number_of (w BIT B1)::'a::{ring_char_0,number_ring})" |
20485 | 423 |
by (simp add: iszero_number_of_BIT) |
15013 | 424 |
|
425 |
||
20485 | 426 |
subsection {* The Less-Than Relation *} |
15013 | 427 |
|
428 |
lemma less_number_of_eq_neg: |
|
20485 | 429 |
"((number_of x::'a::{ordered_idom,number_ring}) < number_of y) |
430 |
= neg (number_of (x + uminus y) :: 'a)" |
|
15013 | 431 |
apply (subst less_iff_diff_less_0) |
432 |
apply (simp add: neg_def diff_minus number_of_add number_of_minus) |
|
433 |
done |
|
434 |
||
20485 | 435 |
text {* |
436 |
If @{term Numeral0} is rewritten to 0 then this rule can't be applied: |
|
437 |
@{term Numeral0} IS @{term "number_of Pls"} |
|
438 |
*} |
|
439 |
||
15013 | 440 |
lemma not_neg_number_of_Pls: |
20485 | 441 |
"~ neg (number_of Pls ::'a::{ordered_idom,number_ring})" |
442 |
by (simp add: neg_def numeral_0_eq_0) |
|
15013 | 443 |
|
444 |
lemma neg_number_of_Min: |
|
20485 | 445 |
"neg (number_of Min ::'a::{ordered_idom,number_ring})" |
446 |
by (simp add: neg_def zero_less_one numeral_m1_eq_minus_1) |
|
15013 | 447 |
|
20485 | 448 |
lemma double_less_0_iff: |
449 |
"(a + a < 0) = (a < (0::'a::ordered_idom))" |
|
15013 | 450 |
proof - |
451 |
have "(a + a < 0) = ((1+1)*a < 0)" by (simp add: left_distrib) |
|
452 |
also have "... = (a < 0)" |
|
453 |
by (simp add: mult_less_0_iff zero_less_two |
|
454 |
order_less_not_sym [OF zero_less_two]) |
|
455 |
finally show ?thesis . |
|
456 |
qed |
|
457 |
||
20485 | 458 |
lemma odd_less_0: |
459 |
"(1 + z + z < 0) = (z < (0::int))"; |
|
15013 | 460 |
proof (cases z rule: int_cases) |
461 |
case (nonneg n) |
|
462 |
thus ?thesis by (simp add: linorder_not_less add_assoc add_increasing |
|
463 |
le_imp_0_less [THEN order_less_imp_le]) |
|
464 |
next |
|
465 |
case (neg n) |
|
466 |
thus ?thesis by (simp del: int_Suc |
|
20485 | 467 |
add: int_Suc0_eq_1 [symmetric] zadd_int compare_rls) |
15013 | 468 |
qed |
469 |
||
20485 | 470 |
text {* The premise involving @{term Ints} prevents @{term "a = 1/2"}. *} |
471 |
||
15013 | 472 |
lemma Ints_odd_less_0: |
20485 | 473 |
assumes in_Ints: "a \<in> Ints" |
474 |
shows "(1 + a + a < 0) = (a < (0::'a::ordered_idom))"; |
|
475 |
proof - |
|
476 |
from in_Ints have "a \<in> range of_int" unfolding Ints_def [symmetric] . |
|
15013 | 477 |
then obtain z where a: "a = of_int z" .. |
478 |
hence "((1::'a) + a + a < 0) = (of_int (1 + z + z) < (of_int 0 :: 'a))" |
|
479 |
by (simp add: a) |
|
480 |
also have "... = (z < 0)" by (simp only: of_int_less_iff odd_less_0) |
|
481 |
also have "... = (a < 0)" by (simp add: a) |
|
482 |
finally show ?thesis . |
|
483 |
qed |
|
484 |
||
485 |
lemma neg_number_of_BIT: |
|
20485 | 486 |
"neg (number_of (w BIT x)::'a) = |
487 |
neg (number_of w :: 'a::{ordered_idom,number_ring})" |
|
488 |
by (simp add: neg_def number_of_eq numeral_simps double_less_0_iff |
|
489 |
Ints_odd_less_0 Ints_def split: bit.split) |
|
15013 | 490 |
|
20596 | 491 |
|
20485 | 492 |
text {* Less-Than or Equals *} |
493 |
||
494 |
text {* Reduces @{term "a\<le>b"} to @{term "~ (b<a)"} for ALL numerals. *} |
|
15013 | 495 |
|
496 |
lemmas le_number_of_eq_not_less = |
|
20485 | 497 |
linorder_not_less [of "number_of w" "number_of v", symmetric, |
498 |
standard] |
|
15013 | 499 |
|
500 |
lemma le_number_of_eq: |
|
501 |
"((number_of x::'a::{ordered_idom,number_ring}) \<le> number_of y) |
|
20485 | 502 |
= (~ (neg (number_of (y + uminus x) :: 'a)))" |
15013 | 503 |
by (simp add: le_number_of_eq_not_less less_number_of_eq_neg) |
504 |
||
505 |
||
20485 | 506 |
text {* Absolute value (@{term abs}) *} |
15013 | 507 |
|
508 |
lemma abs_number_of: |
|
20485 | 509 |
"abs(number_of x::'a::{ordered_idom,number_ring}) = |
510 |
(if number_of x < (0::'a) then -number_of x else number_of x)" |
|
511 |
by (simp add: abs_if) |
|
15013 | 512 |
|
513 |
||
20485 | 514 |
text {* Re-orientation of the equation nnn=x *} |
15013 | 515 |
|
20485 | 516 |
lemma number_of_reorient: |
517 |
"(number_of w = x) = (x = number_of w)" |
|
518 |
by auto |
|
15013 | 519 |
|
520 |
||
20485 | 521 |
subsection {* Simplification of arithmetic operations on integer constants. *} |
15013 | 522 |
|
22801 | 523 |
lemmas arith_extra_simps [standard, simp] = |
20485 | 524 |
number_of_add [symmetric] |
525 |
number_of_minus [symmetric] numeral_m1_eq_minus_1 [symmetric] |
|
526 |
number_of_mult [symmetric] |
|
527 |
diff_number_of_eq abs_number_of |
|
528 |
||
529 |
text {* |
|
530 |
For making a minimal simpset, one must include these default simprules. |
|
531 |
Also include @{text simp_thms}. |
|
532 |
*} |
|
15013 | 533 |
|
20485 | 534 |
lemmas arith_simps = |
535 |
bit.distinct |
|
536 |
Pls_0_eq Min_1_eq |
|
537 |
pred_Pls pred_Min pred_1 pred_0 |
|
538 |
succ_Pls succ_Min succ_1 succ_0 |
|
539 |
add_Pls add_Min add_BIT_0 add_BIT_10 add_BIT_11 |
|
540 |
minus_Pls minus_Min minus_1 minus_0 |
|
541 |
mult_Pls mult_Min mult_num1 mult_num0 |
|
542 |
add_Pls_right add_Min_right |
|
543 |
abs_zero abs_one arith_extra_simps |
|
15013 | 544 |
|
20485 | 545 |
text {* Simplification of relational operations *} |
15013 | 546 |
|
22801 | 547 |
lemmas rel_simps [simp] = |
20485 | 548 |
eq_number_of_eq iszero_number_of_Pls nonzero_number_of_Min |
549 |
iszero_number_of_0 iszero_number_of_1 |
|
550 |
less_number_of_eq_neg |
|
551 |
not_neg_number_of_Pls not_neg_0 not_neg_1 not_iszero_1 |
|
552 |
neg_number_of_Min neg_number_of_BIT |
|
553 |
le_number_of_eq |
|
554 |
||
15013 | 555 |
|
20485 | 556 |
subsection {* Simplification of arithmetic when nested to the right. *} |
15013 | 557 |
|
558 |
lemma add_number_of_left [simp]: |
|
20485 | 559 |
"number_of v + (number_of w + z) = |
560 |
(number_of(v + w) + z::'a::number_ring)" |
|
561 |
by (simp add: add_assoc [symmetric]) |
|
15013 | 562 |
|
563 |
lemma mult_number_of_left [simp]: |
|
20485 | 564 |
"number_of v * (number_of w * z) = |
565 |
(number_of(v * w) * z::'a::number_ring)" |
|
566 |
by (simp add: mult_assoc [symmetric]) |
|
15013 | 567 |
|
568 |
lemma add_number_of_diff1: |
|
20485 | 569 |
"number_of v + (number_of w - c) = |
570 |
number_of(v + w) - (c::'a::number_ring)" |
|
571 |
by (simp add: diff_minus add_number_of_left) |
|
15013 | 572 |
|
20485 | 573 |
lemma add_number_of_diff2 [simp]: |
574 |
"number_of v + (c - number_of w) = |
|
575 |
number_of (v + uminus w) + (c::'a::number_ring)" |
|
15013 | 576 |
apply (subst diff_number_of_eq [symmetric]) |
577 |
apply (simp only: compare_rls) |
|
578 |
done |
|
579 |
||
19380 | 580 |
|
22801 | 581 |
subsection {* Configuration of the code generator *} |
582 |
||
583 |
instance int :: eq .. |
|
584 |
||
585 |
code_datatype Pls Min Bit "number_of \<Colon> int \<Rightarrow> int" |
|
586 |
||
587 |
definition |
|
588 |
int_aux :: "int \<Rightarrow> nat \<Rightarrow> int" where |
|
589 |
"int_aux i n = (i + int n)" |
|
590 |
||
591 |
lemma [code]: |
|
592 |
"int_aux i 0 = i" |
|
593 |
"int_aux i (Suc n) = int_aux (i + 1) n" -- {* tail recursive *} |
|
594 |
by (simp add: int_aux_def)+ |
|
595 |
||
596 |
lemma [code]: |
|
597 |
"int n = int_aux 0 n" |
|
598 |
by (simp add: int_aux_def) |
|
599 |
||
600 |
definition |
|
601 |
nat_aux :: "nat \<Rightarrow> int \<Rightarrow> nat" where |
|
602 |
"nat_aux n i = (n + nat i)" |
|
603 |
||
604 |
lemma [code]: "nat_aux n i = (if i <= 0 then n else nat_aux (Suc n) (i - 1))" |
|
605 |
-- {* tail recursive *} |
|
606 |
by (auto simp add: nat_aux_def nat_eq_iff linorder_not_le order_less_imp_le |
|
607 |
dest: zless_imp_add1_zle) |
|
608 |
||
609 |
lemma [code]: "nat i = nat_aux 0 i" |
|
610 |
by (simp add: nat_aux_def) |
|
611 |
||
612 |
lemma zero_is_num_zero [code func, code inline, symmetric, normal post]: |
|
613 |
"(0\<Colon>int) = number_of Numeral.Pls" |
|
614 |
by simp |
|
615 |
||
616 |
lemma one_is_num_one [code func, code inline, symmetric, normal post]: |
|
617 |
"(1\<Colon>int) = number_of (Numeral.Pls BIT bit.B1)" |
|
618 |
by simp |
|
619 |
||
620 |
code_modulename SML |
|
621 |
IntDef Integer |
|
622 |
||
623 |
code_modulename OCaml |
|
624 |
IntDef Integer |
|
625 |
||
626 |
code_modulename Haskell |
|
627 |
IntDef Integer |
|
628 |
||
629 |
code_modulename SML |
|
630 |
Numeral Integer |
|
631 |
||
632 |
code_modulename OCaml |
|
633 |
Numeral Integer |
|
634 |
||
635 |
code_modulename Haskell |
|
636 |
Numeral Integer |
|
637 |
||
638 |
(*FIXME: the IntInf.fromInt below hides a dependence on fixed-precision ints!*) |
|
639 |
||
640 |
types_code |
|
641 |
"int" ("int") |
|
642 |
attach (term_of) {* |
|
643 |
val term_of_int = HOLogic.mk_number HOLogic.intT o IntInf.fromInt; |
|
644 |
*} |
|
645 |
attach (test) {* |
|
646 |
fun gen_int i = one_of [~1, 1] * random_range 0 i; |
|
647 |
*} |
|
648 |
||
649 |
setup {* |
|
650 |
let |
|
651 |
||
22997 | 652 |
fun number_of_codegen thy defs gr dep module b (Const (@{const_name Numeral.number_of}, Type ("fun", [_, T])) $ t) = |
22801 | 653 |
if T = HOLogic.intT then |
654 |
(SOME (fst (Codegen.invoke_tycodegen thy defs dep module false (gr, T)), |
|
655 |
(Pretty.str o IntInf.toString o HOLogic.dest_numeral) t) handle TERM _ => NONE) |
|
656 |
else if T = HOLogic.natT then |
|
657 |
SOME (Codegen.invoke_codegen thy defs dep module b (gr, |
|
658 |
Const ("IntDef.nat", HOLogic.intT --> HOLogic.natT) $ |
|
22997 | 659 |
(Const (@{const_name Numeral.number_of}, HOLogic.intT --> HOLogic.intT) $ t))) |
22801 | 660 |
else NONE |
661 |
| number_of_codegen _ _ _ _ _ _ _ = NONE; |
|
662 |
||
663 |
in |
|
664 |
||
665 |
Codegen.add_codegen "number_of_codegen" number_of_codegen |
|
666 |
||
667 |
end |
|
668 |
*} |
|
669 |
||
670 |
consts_code |
|
22921
475ff421a6a3
consts in consts_code Isar commands are now referred to by usual term syntax
haftmann
parents:
22911
diff
changeset
|
671 |
"0 :: int" ("0") |
475ff421a6a3
consts in consts_code Isar commands are now referred to by usual term syntax
haftmann
parents:
22911
diff
changeset
|
672 |
"1 :: int" ("1") |
475ff421a6a3
consts in consts_code Isar commands are now referred to by usual term syntax
haftmann
parents:
22911
diff
changeset
|
673 |
"uminus :: int => int" ("~") |
475ff421a6a3
consts in consts_code Isar commands are now referred to by usual term syntax
haftmann
parents:
22911
diff
changeset
|
674 |
"op + :: int => int => int" ("(_ +/ _)") |
475ff421a6a3
consts in consts_code Isar commands are now referred to by usual term syntax
haftmann
parents:
22911
diff
changeset
|
675 |
"op * :: int => int => int" ("(_ */ _)") |
475ff421a6a3
consts in consts_code Isar commands are now referred to by usual term syntax
haftmann
parents:
22911
diff
changeset
|
676 |
"op \<le> :: int => int => bool" ("(_ <=/ _)") |
475ff421a6a3
consts in consts_code Isar commands are now referred to by usual term syntax
haftmann
parents:
22911
diff
changeset
|
677 |
"op < :: int => int => bool" ("(_ </ _)") |
22801 | 678 |
|
679 |
quickcheck_params [default_type = int] |
|
680 |
||
22921
475ff421a6a3
consts in consts_code Isar commands are now referred to by usual term syntax
haftmann
parents:
22911
diff
changeset
|
681 |
(*setup continues in theory Presburger*) |
22801 | 682 |
|
20500 | 683 |
hide (open) const Pls Min B0 B1 succ pred |
19380 | 684 |
|
15013 | 685 |
end |