| author | berghofe | 
| Tue, 11 Dec 2001 14:54:18 +0100 | |
| changeset 12461 | 23686cad32d6 | 
| parent 12338 | de0f4a63baa5 | 
| child 12738 | 9d80e3746eb0 | 
| 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  | 
|
| 
12338
 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 
wenzelm 
parents: 
12098 
diff
changeset
 | 
17  | 
  number < type  -- {* for numeric types: nat, int, real, \dots *}
 | 
| 6905 | 18  | 
|
19  | 
consts  | 
|
| 9035 | 20  | 
number_of :: "bin => 'a::number"  | 
| 6905 | 21  | 
|
22  | 
syntax  | 
|
| 12098 | 23  | 
  "_Numeral" :: "num_const => 'a"    ("_")
 | 
| 11699 | 24  | 
Numeral0 :: 'a  | 
25  | 
Numeral1 :: 'a  | 
|
26  | 
||
27  | 
translations  | 
|
28  | 
"Numeral0" == "number_of Pls"  | 
|
29  | 
"Numeral1" == "number_of (Pls BIT True)"  | 
|
| 
11488
 
4ff900551340
constify numeral tokens in order to allow translations;
 
wenzelm 
parents: 
9410 
diff
changeset
 | 
30  | 
|
| 6905 | 31  | 
|
| 9035 | 32  | 
setup NumeralSyntax.setup  | 
| 6905 | 33  | 
|
| 11699 | 34  | 
|
35  | 
(*Unfold all "let"s involving constants*)  | 
|
36  | 
lemma Let_number_of [simp]: "Let (number_of v) f == f (number_of v)"  | 
|
37  | 
by (simp add: Let_def)  | 
|
38  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
39  | 
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
 | 
40  | 
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
 | 
41  | 
|
| 
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
42  | 
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
 | 
43  | 
by (simp add: Let_def)  | 
| 11699 | 44  | 
|
| 9035 | 45  | 
end  |