src/HOL/Tools/numeral_syntax.ML
author wenzelm
Wed, 13 Jul 2005 16:07:35 +0200
changeset 16813 67140ae50e77
parent 16365 838c65dad23a
child 18708 4b3dadb4fe33
permissions -rw-r--r--
removed ad-hoc atp_hook, cal_atp; removed depth_of; tuned;

(*  Title:      HOL/Tools/numeral_syntax.ML
    ID:         $Id$
    Author:     Markus Wenzel, TU Muenchen

Concrete syntax for generic numerals.  Assumes consts and syntax of
theory HOL/Numeral.
*)

signature NUMERAL_SYNTAX =
sig
  val setup: (theory -> theory) list
end;

structure NumeralSyntax: NUMERAL_SYNTAX =
struct


(* bit strings *)   (*we try to handle superfluous leading digits nicely*)

fun prefix_len _ [] = 0
  | prefix_len pred (x :: xs) =
      if pred x then 1 + prefix_len pred xs else 0;

fun dest_bin_str tm =
  let
    val rev_digs = HOLogic.bin_of tm handle TERM _ => raise Match
    val (sign, zs) =
      (case rev rev_digs of
        ~1 :: bs => ("-", prefix_len (equal 1) bs)
      | bs => ("", prefix_len (equal 0) bs));
    val i = HOLogic.int_of rev_digs;
    val num = IntInf.toString (IntInf.abs i);
  in
    if i = IntInf.fromInt 0 orelse i = IntInf.fromInt 1 then raise Match
    else sign ^ implode (replicate zs "0") ^ num
  end;

(* translation of integer numeral tokens to and from bitstrings *)

fun numeral_tr (*"_Numeral"*) [t as Const (str, _)] =
      (Syntax.const "Numeral.number_of" $ (HOLogic.mk_bin (valOf (IntInf.fromString str))))
  | numeral_tr (*"_Numeral"*) ts = raise TERM ("numeral_tr", ts);

fun numeral_tr' show_sorts (*"number_of"*) (Type ("fun", [_, T])) (t :: ts) =
      let val t' = Syntax.const "_Numeral" $ Syntax.free (dest_bin_str t) in
        if not (! show_types) andalso can Term.dest_Type T then t'
        else Syntax.const Syntax.constrainC $ t' $ Syntax.term_of_typ show_sorts T
      end
  | numeral_tr' _ (*"number_of"*) T (t :: ts) =
      if T = dummyT then Syntax.const "_Numeral" $ Syntax.free (dest_bin_str t)
      else raise Match
  | numeral_tr' _ (*"number_of"*) _ _ = raise Match;


(* theory setup *)

val setup =
 [Theory.hide_consts_i false ["Numeral.Pls", "Numeral.Min"],
  Theory.hide_consts_i false ["Numeral.bit.B0", "Numeral.bit.B1"],
   Theory.add_trfuns ([], [("_Numeral", numeral_tr)], [], []),
  Theory.add_trfunsT [("number_of", numeral_tr'), ("Numeral.number_of", numeral_tr')]];

end;