src/HOL/Numeral.thy
author berghofe
Mon, 29 Jan 2001 13:28:15 +0100
changeset 10989 87f8a7644f91
parent 9410 612ee826a409
child 11488 4ff900551340
permissions -rw-r--r--
New function complete_split_rule for complete splitting of partially splitted rules (as generated by inductive definition package).

(*  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