author | wenzelm |
Sat, 03 Nov 2001 18:42:38 +0100 | |
changeset 12038 | 343a9888e875 |
parent 11868 | 56db9f3a6b3e |
child 12098 | 784fe681ba26 |
permissions | -rw-r--r-- |
6905 | 1 |
(* Title: HOL/Numeral.thy |
2 |
ID: $Id$ |
|
3 |
Author: Larry Paulson and Markus Wenzel |
|
4 |
||
9410
612ee826a409
removed selector syntax -- improper tuples are broken beyond repair :-(
wenzelm
parents:
9358
diff
changeset
|
5 |
Generic numerals represented as twos-complement bit strings. |
6905 | 6 |
*) |
7 |
||
8 |
theory Numeral = Datatype |
|
9035 | 9 |
files "Tools/numeral_syntax.ML": |
6905 | 10 |
|
11 |
datatype |
|
12 |
bin = Pls |
|
13 |
| Min |
|
9358
973672495414
added syntax for proper / improper selector functions;
wenzelm
parents:
9035
diff
changeset
|
14 |
| Bit bin bool (infixl "BIT" 90) |
6905 | 15 |
|
16 |
axclass |
|
9035 | 17 |
number < "term" (*for numeric types: nat, int, real, ...*) |
6905 | 18 |
|
19 |
consts |
|
9035 | 20 |
number_of :: "bin => 'a::number" |
6905 | 21 |
|
11488
4ff900551340
constify numeral tokens in order to allow translations;
wenzelm
parents:
9410
diff
changeset
|
22 |
nonterminals |
4ff900551340
constify numeral tokens in order to allow translations;
wenzelm
parents:
9410
diff
changeset
|
23 |
numeral |
4ff900551340
constify numeral tokens in order to allow translations;
wenzelm
parents:
9410
diff
changeset
|
24 |
|
6905 | 25 |
syntax |
11699 | 26 |
"_constify" :: "num => numeral" ("_") |
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11699
diff
changeset
|
27 |
"_Numeral" :: "numeral => 'a" ("_") |
11699 | 28 |
Numeral0 :: 'a |
29 |
Numeral1 :: 'a |
|
30 |
||
31 |
translations |
|
32 |
"Numeral0" == "number_of Pls" |
|
33 |
"Numeral1" == "number_of (Pls BIT True)" |
|
11488
4ff900551340
constify numeral tokens in order to allow translations;
wenzelm
parents:
9410
diff
changeset
|
34 |
|
6905 | 35 |
|
9035 | 36 |
setup NumeralSyntax.setup |
6905 | 37 |
|
11699 | 38 |
|
39 |
(*Unfold all "let"s involving constants*) |
|
40 |
lemma Let_number_of [simp]: "Let (number_of v) f == f (number_of v)" |
|
41 |
by (simp add: Let_def) |
|
42 |
||
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
43 |
lemma Let_0 [simp]: "Let 0 f == f 0" |
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
44 |
by (simp add: Let_def) |
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
45 |
|
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
46 |
lemma Let_1 [simp]: "Let 1 f == f 1" |
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
47 |
by (simp add: Let_def) |
11699 | 48 |
|
9035 | 49 |
end |