src/HOL/Tools/numeral.ML
author haftmann
Tue, 15 Jan 2008 16:19:23 +0100
changeset 25919 8b1c0d434824
parent 24630 351a308ab58d
child 25932 db0fd0ecdcd4
permissions -rw-r--r--
joined theories IntDef, Numeral, IntArith to theory Int

(*  Title:      HOL/Tools/Int.ML
    ID:         $Id$
    Author:     Makarius

Logical operations on numerals (see also HOL/hologic.ML).
*)

signature NUMERAL =
sig
  val mk_cnumeral: int -> cterm
  val mk_cnumber: ctyp -> int -> cterm
end;

structure Numeral: NUMERAL =
struct

(* numeral *)

fun mk_cbit 0 = @{cterm "Int.bit.B0"}
  | mk_cbit 1 = @{cterm "Int.bit.B1"}
  | mk_cbit _ = raise CTERM ("mk_cbit", []);

fun mk_cnumeral 0 = @{cterm "Int.Pls"}
  | mk_cnumeral ~1 = @{cterm "Int.Min"}
  | mk_cnumeral i =
      let val (q, r) = Integer.div_mod i 2 in
        Thm.capply (Thm.capply @{cterm "Int.Bit"} (mk_cnumeral q)) (mk_cbit r)
      end;


(* number *)

local

val zero = @{cpat "0"};
val zeroT = Thm.ctyp_of_term zero;

val one = @{cpat "1"};
val oneT = Thm.ctyp_of_term one;

val number_of = @{cpat "number_of"};
val numberT = Thm.ctyp_of @{theory} (Term.range_type (Thm.typ_of (Thm.ctyp_of_term number_of)));

fun instT T V = Thm.instantiate_cterm ([(V, T)], []);

in

fun mk_cnumber T 0 = instT T zeroT zero
  | mk_cnumber T 1 = instT T oneT one
  | mk_cnumber T i = Thm.capply (instT T numberT number_of) (mk_cnumeral i);

end;

end;