author | wenzelm |
Sun, 17 Sep 2000 22:25:18 +0200 | |
changeset 10012 | 4961c73b5f60 |
parent 9410 | 612ee826a409 |
child 11488 | 4ff900551340 |
permissions | -rw-r--r-- |
(* Title: HOL/Numeral.thy ID: $Id$ Author: Larry Paulson and Markus Wenzel Generic numerals represented as twos-complement bit strings. *) theory Numeral = Datatype files "Tools/numeral_syntax.ML": datatype bin = Pls | Min | Bit bin bool (infixl "BIT" 90) axclass number < "term" (*for numeric types: nat, int, real, ...*) consts number_of :: "bin => 'a::number" syntax "_Numeral" :: "xnum => 'a" ("_") setup NumeralSyntax.setup end