src/Pure/Syntax/mixfix.ML
author clasohm
Mon, 03 Jul 1995 15:39:53 +0200
changeset 1178 b28c6ecc3e6d
parent 1067 00ed040f66e1
child 1181 c4e90fb7f8fa
permissions -rw-r--r--
added cargs for curried function application
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
384
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Syntax/mixfix.ML
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
551
4c139c37dbaf minor cleanings;
wenzelm
parents: 472
diff changeset
     3
    Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen
384
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
     4
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
     5
Mixfix declarations, infixes, binders. Part of the Pure syntax.
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
     6
*)
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
     7
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
     8
signature MIXFIX0 =
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
     9
sig
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    10
  datatype mixfix =
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    11
    NoSyn |
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    12
    Mixfix of string * int list * int |
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    13
    Delimfix of string |
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    14
    Infixl of int |
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    15
    Infixr of int |
865
b38c67991122 added optional precedence for body of binder;
clasohm
parents: 842
diff changeset
    16
    Binder of string * int * int
384
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    17
  val max_pri: int
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    18
end;
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    19
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    20
signature MIXFIX1 =
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    21
sig
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    22
  include MIXFIX0
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    23
  val type_name: string -> mixfix -> string
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    24
  val const_name: string -> mixfix -> string
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    25
  val pure_types: (string * int * mixfix) list
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    26
  val pure_syntax: (string * string * mixfix) list
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    27
end;
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    28
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    29
signature MIXFIX =
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    30
sig
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    31
  include MIXFIX1
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    32
  structure SynExt: SYN_EXT
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    33
  local open SynExt in
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    34
    val syn_ext_types: string list -> (string * int * mixfix) list -> syn_ext
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    35
    val syn_ext_consts: string list -> (string * typ * mixfix) list -> syn_ext
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    36
  end
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    37
end;
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    38
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    39
functor MixfixFun(structure Lexicon: LEXICON and SynExt: SYN_EXT
551
4c139c37dbaf minor cleanings;
wenzelm
parents: 472
diff changeset
    40
  and SynTrans: SYN_TRANS): MIXFIX =
384
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    41
struct
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    42
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    43
structure SynExt = SynExt;
551
4c139c37dbaf minor cleanings;
wenzelm
parents: 472
diff changeset
    44
open Lexicon SynExt SynExt.Ast SynTrans;
384
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    45
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    46
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    47
(** mixfix declarations **)
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    48
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    49
datatype mixfix =
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    50
  NoSyn |
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    51
  Mixfix of string * int list * int |
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    52
  Delimfix of string |
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    53
  Infixl of int |
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    54
  Infixr of int |
887
6a054d83acb2 *** empty log message ***
wenzelm
parents: 865
diff changeset
    55
  Binder of string * int * int;
384
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    56
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    57
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    58
(* type / const names *)
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    59
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    60
fun strip ("'" :: c :: cs) = c :: strip cs
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    61
  | strip ["'"] = []
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    62
  | strip (c :: cs) = c :: strip cs
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    63
  | strip [] = [];
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    64
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    65
val strip_esc = implode o strip o explode;
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    66
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    67
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    68
fun type_name t (Infixl _) = strip_esc t
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    69
  | type_name t (Infixr _) = strip_esc t
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    70
  | type_name t _ = t;
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    71
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    72
fun infix_name c = "op " ^ strip_esc c;
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    73
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    74
fun const_name c (Infixl _) = infix_name c
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    75
  | const_name c (Infixr _) = infix_name c
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    76
  | const_name c _ = c;
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    77
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    78
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    79
(* syn_ext_types *)
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    80
764
b60e77395d1a changed Pure's grammar and the way types are converted to nonterminals
clasohm
parents: 639
diff changeset
    81
fun syn_ext_types logtypes type_decls =
384
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    82
  let
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    83
    fun name_of (t, _, mx) = type_name t mx;
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    84
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    85
    fun mk_infix t p1 p2 p3 =
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    86
      Mfix ("(_ " ^ t ^ "/ _)", [typeT, typeT] ---> typeT,
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    87
        strip_esc t, [p1, p2], p3);
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    88
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    89
    fun mfix_of (_, _, NoSyn) = None
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    90
      | mfix_of (t, 2, Infixl p) = Some (mk_infix t p (p + 1) p)
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    91
      | mfix_of (t, 2, Infixr p) = Some (mk_infix t (p + 1) p p)
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    92
      | mfix_of decl = error ("Bad mixfix declaration for type " ^
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    93
          quote (name_of decl));
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    94
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    95
    val mfix = mapfilter mfix_of type_decls;
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    96
    val xconsts = map name_of type_decls;
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    97
  in
764
b60e77395d1a changed Pure's grammar and the way types are converted to nonterminals
clasohm
parents: 639
diff changeset
    98
    syn_ext logtypes mfix xconsts ([], [], [], []) ([], [])
384
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
    99
  end;
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   100
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   101
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   102
(* syn_ext_consts *)
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   103
764
b60e77395d1a changed Pure's grammar and the way types are converted to nonterminals
clasohm
parents: 639
diff changeset
   104
fun syn_ext_consts logtypes const_decls =
384
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   105
  let
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   106
    fun name_of (c, _, mx) = const_name c mx;
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   107
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   108
    fun mk_infix sy ty c p1 p2 p3 =
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   109
      [Mfix ("(_ " ^ sy ^ "/ _)", ty, c, [p1, p2], p3),
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   110
       Mfix ("op " ^ sy, ty, c, [], max_pri)];
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   111
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   112
    fun binder_typ _ (Type ("fun", [Type ("fun", [_, ty2]), ty3])) =
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   113
          [Type ("idts", []), ty2] ---> ty3
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   114
      | binder_typ c _ = error ("Bad type of binder " ^ quote c);
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   115
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   116
    fun mfix_of (_, _, NoSyn) = []
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   117
      | mfix_of (c, ty, Mixfix (sy, ps, p)) = [Mfix (sy, ty, c, ps, p)]
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   118
      | mfix_of (c, ty, Delimfix sy) = [Mfix (sy, ty, c, [], max_pri)]
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   119
      | mfix_of (sy, ty, Infixl p) = mk_infix sy ty (infix_name sy) p (p + 1) p
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   120
      | mfix_of (sy, ty, Infixr p) = mk_infix sy ty (infix_name sy) (p + 1) p p
887
6a054d83acb2 *** empty log message ***
wenzelm
parents: 865
diff changeset
   121
      | mfix_of (c, ty, Binder (sy, p, q)) =
6a054d83acb2 *** empty log message ***
wenzelm
parents: 865
diff changeset
   122
          [Mfix ("(3" ^ sy ^ "_./ _)", binder_typ c ty, sy, [0, p], q)];
384
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   123
865
b38c67991122 added optional precedence for body of binder;
clasohm
parents: 842
diff changeset
   124
    fun binder (c, _, Binder (sy, _, _)) = Some (sy, c)
384
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   125
      | binder _ = None;
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   126
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   127
    val mfix = flat (map mfix_of const_decls);
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   128
    val xconsts = map name_of const_decls;
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   129
    val binders = mapfilter binder const_decls;
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   130
    val binder_trs = map mk_binder_tr binders;
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   131
    val binder_trs' = map (mk_binder_tr' o swap) binders;
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   132
  in
764
b60e77395d1a changed Pure's grammar and the way types are converted to nonterminals
clasohm
parents: 639
diff changeset
   133
    syn_ext logtypes mfix xconsts ([], binder_trs, binder_trs', []) ([], [])
384
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   134
  end;
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   135
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   136
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   137
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   138
(** Pure syntax **)
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   139
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   140
val pure_types =
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   141
  map (fn t => (t, 0, NoSyn))
1178
b28c6ecc3e6d added cargs for curried function application
clasohm
parents: 1067
diff changeset
   142
    (terminals @ [logic, "type", "types", "sort", "classes", args, cargs,
1067
00ed040f66e1 I have modified the grammar for idts (sequences of identifiers with optional
nipkow
parents: 922
diff changeset
   143
      "pttrn", "idt", "idts", "aprop", "asms", any, sprop]);
384
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   144
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   145
val pure_syntax =
781
9ab8873bf9b3 added any, sprop to pure_types;
wenzelm
parents: 764
diff changeset
   146
 [("_lambda",   "[idts, 'a] => logic",            Mixfix ("(3%_./ _)", [], 0)),
921
6bee3815c0bf added declaration of syntactic const "_abs";
wenzelm
parents: 887
diff changeset
   147
  ("_abs",      "'a",                             NoSyn),
384
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   148
  ("",          "'a => " ^ args,                  Delimfix "_"),
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   149
  ("_args",     "['a, " ^ args ^ "] => " ^ args,  Delimfix "_,/ _"),
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   150
  ("",          "id => idt",                      Delimfix "_"),
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   151
  ("_idtyp",    "[id, type] => idt",              Mixfix ("_::_", [], 0)),
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   152
  ("",          "idt => idt",                     Delimfix "'(_')"),
1067
00ed040f66e1 I have modified the grammar for idts (sequences of identifiers with optional
nipkow
parents: 922
diff changeset
   153
  ("",          "idt => pttrn",                   Delimfix "_"),
00ed040f66e1 I have modified the grammar for idts (sequences of identifiers with optional
nipkow
parents: 922
diff changeset
   154
  ("",          "pttrn => idts",                  Delimfix "_"),
00ed040f66e1 I have modified the grammar for idts (sequences of identifiers with optional
nipkow
parents: 922
diff changeset
   155
  ("_idts",     "[pttrn, idts] => idts",          Mixfix ("_/ _", [1, 0], 0)),
384
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   156
  ("",          "id => aprop",                    Delimfix "_"),
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   157
  ("",          "var => aprop",                   Delimfix "_"),
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   158
  ("_aprop",    "aprop => prop",                  Delimfix "PROP _"),
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   159
  ("",          "prop => asms",                   Delimfix "_"),
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   160
  ("_asms",     "[prop, asms] => asms",           Delimfix "_;/ _"),
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   161
  ("_bigimpl",  "[asms, prop] => prop",           Mixfix ("((3[| _ |]) ==>/ _)", [0, 1], 1)),
842
8d45c937a485 slightly changed OFCLASS syntax;
wenzelm
parents: 781
diff changeset
   162
  ("_ofclass",  "[type, logic] => prop",          Delimfix "(1OFCLASS/(1'(_,/ _')))"),
384
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   163
  ("_K",        "'a",                             NoSyn),
624
33b9b5da3e6f made major changes to grammar;
clasohm
parents: 551
diff changeset
   164
  ("",          "id => logic",                    Delimfix "_"),
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 921
diff changeset
   165
  ("",          "var => logic",                   Delimfix "_")]
384
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   166
a8d204d8071d support for new style mixfix annotations;
wenzelm
parents:
diff changeset
   167
end;