src/ZF/Tools/numeral_syntax.ML
author wenzelm
Mon, 03 Dec 2018 22:21:35 +0100
changeset 69399 9e488aeb9870
parent 58421 37cbbd8eb460
child 69593 3dda49e08b9d
permissions -rw-r--r--
updated to jedit_build-20181203 (according to d70767e508d7);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
     1
(*  Title:      ZF/Tools/numeral_syntax.ML
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
     3
42246
8f798ba04393 misc tuning and simplification;
wenzelm
parents: 41310
diff changeset
     4
Concrete syntax for generic numerals.
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
     5
*)
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
     6
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
     7
signature NUMERAL_SYNTAX =
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
     8
sig
35112
ff6f60e6ab85 numeral syntax: clarify parse trees vs. actual terms;
wenzelm
parents: 32960
diff changeset
     9
  val make_binary: int -> int list
ff6f60e6ab85 numeral syntax: clarify parse trees vs. actual terms;
wenzelm
parents: 32960
diff changeset
    10
  val dest_binary: int list -> int
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    11
end;
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    12
35123
e286d5df187a modernized structures;
wenzelm
parents: 35112
diff changeset
    13
structure Numeral_Syntax: NUMERAL_SYNTAX =
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    14
struct
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    15
35112
ff6f60e6ab85 numeral syntax: clarify parse trees vs. actual terms;
wenzelm
parents: 32960
diff changeset
    16
(* bits *)
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    17
41310
65631ca437c9 proper identifiers for consts and types;
wenzelm
parents: 40314
diff changeset
    18
fun mk_bit 0 = Syntax.const @{const_syntax zero}
65631ca437c9 proper identifiers for consts and types;
wenzelm
parents: 40314
diff changeset
    19
  | mk_bit 1 = Syntax.const @{const_syntax succ} $ Syntax.const @{const_syntax zero}
40314
b5ec88d9ac03 replaced ancient sys_error by raise Fail, assuming that the latter is not handled specifically by the environment;
wenzelm
parents: 35123
diff changeset
    20
  | mk_bit _ = raise Fail "mk_bit";
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    21
41310
65631ca437c9 proper identifiers for consts and types;
wenzelm
parents: 40314
diff changeset
    22
fun dest_bit (Const (@{const_syntax zero}, _)) = 0
58421
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 52143
diff changeset
    23
  | dest_bit (Const (@{const_syntax one}, _)) = 1
41310
65631ca437c9 proper identifiers for consts and types;
wenzelm
parents: 40314
diff changeset
    24
  | dest_bit (Const (@{const_syntax succ}, _) $ Const (@{const_syntax zero}, _)) = 1
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    25
  | dest_bit _ = raise Match;
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    26
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    27
35112
ff6f60e6ab85 numeral syntax: clarify parse trees vs. actual terms;
wenzelm
parents: 32960
diff changeset
    28
(* bit strings *)
ff6f60e6ab85 numeral syntax: clarify parse trees vs. actual terms;
wenzelm
parents: 32960
diff changeset
    29
ff6f60e6ab85 numeral syntax: clarify parse trees vs. actual terms;
wenzelm
parents: 32960
diff changeset
    30
fun make_binary 0 = []
ff6f60e6ab85 numeral syntax: clarify parse trees vs. actual terms;
wenzelm
parents: 32960
diff changeset
    31
  | make_binary ~1 = [~1]
ff6f60e6ab85 numeral syntax: clarify parse trees vs. actual terms;
wenzelm
parents: 32960
diff changeset
    32
  | make_binary n = (n mod 2) :: make_binary (n div 2);
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    33
35112
ff6f60e6ab85 numeral syntax: clarify parse trees vs. actual terms;
wenzelm
parents: 32960
diff changeset
    34
fun dest_binary [] = 0
ff6f60e6ab85 numeral syntax: clarify parse trees vs. actual terms;
wenzelm
parents: 32960
diff changeset
    35
  | dest_binary (b :: bs) = b + 2 * dest_binary bs;
ff6f60e6ab85 numeral syntax: clarify parse trees vs. actual terms;
wenzelm
parents: 32960
diff changeset
    36
ff6f60e6ab85 numeral syntax: clarify parse trees vs. actual terms;
wenzelm
parents: 32960
diff changeset
    37
ff6f60e6ab85 numeral syntax: clarify parse trees vs. actual terms;
wenzelm
parents: 32960
diff changeset
    38
(*try to handle superfluous leading digits nicely*)
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    39
fun prefix_len _ [] = 0
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    40
  | prefix_len pred (x :: xs) =
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    41
      if pred x then 1 + prefix_len pred xs else 0;
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    42
24630
351a308ab58d simplified type int (eliminated IntInf.int, integer);
wenzelm
parents: 23146
diff changeset
    43
fun mk_bin i =
35112
ff6f60e6ab85 numeral syntax: clarify parse trees vs. actual terms;
wenzelm
parents: 32960
diff changeset
    44
  let
ff6f60e6ab85 numeral syntax: clarify parse trees vs. actual terms;
wenzelm
parents: 32960
diff changeset
    45
    fun term_of [] = Syntax.const @{const_syntax Pls}
ff6f60e6ab85 numeral syntax: clarify parse trees vs. actual terms;
wenzelm
parents: 32960
diff changeset
    46
      | term_of [~1] = Syntax.const @{const_syntax Min}
ff6f60e6ab85 numeral syntax: clarify parse trees vs. actual terms;
wenzelm
parents: 32960
diff changeset
    47
      | term_of (b :: bs) = Syntax.const @{const_syntax Bit} $ term_of bs $ mk_bit b;
ff6f60e6ab85 numeral syntax: clarify parse trees vs. actual terms;
wenzelm
parents: 32960
diff changeset
    48
  in term_of (make_binary i) end;
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    49
35112
ff6f60e6ab85 numeral syntax: clarify parse trees vs. actual terms;
wenzelm
parents: 32960
diff changeset
    50
fun bin_of (Const (@{const_syntax Pls}, _)) = []
ff6f60e6ab85 numeral syntax: clarify parse trees vs. actual terms;
wenzelm
parents: 32960
diff changeset
    51
  | bin_of (Const (@{const_syntax Min}, _)) = [~1]
ff6f60e6ab85 numeral syntax: clarify parse trees vs. actual terms;
wenzelm
parents: 32960
diff changeset
    52
  | bin_of (Const (@{const_syntax Bit}, _) $ bs $ b) = dest_bit b :: bin_of bs
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    53
  | bin_of _ = raise Match;
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    54
35112
ff6f60e6ab85 numeral syntax: clarify parse trees vs. actual terms;
wenzelm
parents: 32960
diff changeset
    55
(*Leading 0s and (for negative numbers) -1s cause complications, though they 
15965
f422f8283491 Use of IntInf.int instead of int in most numeric simprocs; avoids
paulson
parents: 10917
diff changeset
    56
  should never arise in normal use. The formalization used in HOL prevents 
f422f8283491 Use of IntInf.int instead of int in most numeric simprocs; avoids
paulson
parents: 10917
diff changeset
    57
  them altogether.*)
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    58
fun show_int t =
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    59
  let
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    60
    val rev_digs = bin_of t;
58421
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 52143
diff changeset
    61
    val (c, zs) =
35112
ff6f60e6ab85 numeral syntax: clarify parse trees vs. actual terms;
wenzelm
parents: 32960
diff changeset
    62
      (case rev rev_digs of
58421
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 52143
diff changeset
    63
         ~1 :: bs => (@{syntax_const "_Neg_Int"}, prefix_len (equal 1) bs)
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 52143
diff changeset
    64
      | bs => (@{syntax_const "_Int"},  prefix_len (equal 0) bs));
35112
ff6f60e6ab85 numeral syntax: clarify parse trees vs. actual terms;
wenzelm
parents: 32960
diff changeset
    65
    val num = string_of_int (abs (dest_binary rev_digs));
58421
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 52143
diff changeset
    66
  in (c, implode (replicate zs "0") ^ num) end;
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    67
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    68
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    69
(* translation of integer constant tokens to and from binary *)
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    70
58421
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 52143
diff changeset
    71
fun int_tr [Free (s, _)] =
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 52143
diff changeset
    72
      Syntax.const @{const_syntax integ_of} $ mk_bin (#value (Lexicon.read_num s))
42246
8f798ba04393 misc tuning and simplification;
wenzelm
parents: 41310
diff changeset
    73
  | int_tr ts = raise TERM ("int_tr", ts);
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    74
58421
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 52143
diff changeset
    75
fun neg_int_tr [Free (s, _)] =
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 52143
diff changeset
    76
      Syntax.const @{const_syntax integ_of} $ mk_bin (~ (#value (Lexicon.read_num s)))
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 52143
diff changeset
    77
  | neg_int_tr ts = raise TERM ("neg_int_tr", ts);
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    78
58421
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 52143
diff changeset
    79
fun integ_of_tr' [t] =
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 52143
diff changeset
    80
      let val (c, s) = show_int t
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 52143
diff changeset
    81
      in Syntax.const c $ Syntax.free s end
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 52143
diff changeset
    82
  | integ_of_tr' _ = raise Match;
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 52143
diff changeset
    83
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 52143
diff changeset
    84
val _ = Theory.setup
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 52143
diff changeset
    85
 (Sign.parse_translation
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 52143
diff changeset
    86
   [(@{syntax_const "_Int"}, K int_tr),
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 52143
diff changeset
    87
    (@{syntax_const "_Neg_Int"}, K neg_int_tr)] #>
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 52143
diff changeset
    88
  Sign.print_translation
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 52143
diff changeset
    89
   [(@{const_syntax integ_of}, K integ_of_tr')]);
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    90
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents:
diff changeset
    91
end;