src/Pure/Syntax/mixfix.ML
changeset 384 a8d204d8071d
child 472 bbaa2a02bd82
equal deleted inserted replaced
383:fcea89074e4c 384:a8d204d8071d
       
     1 (*  Title:      Pure/Syntax/mixfix.ML
       
     2     ID:         $Id$
       
     3     Author:     Markus Wenzel, TU Muenchen
       
     4 
       
     5 Mixfix declarations, infixes, binders. Part of the Pure syntax.
       
     6 
       
     7 TODO:
       
     8   note: needs mk_binder_tr/tr'
       
     9   Mixfix.mixfix -> mixfix
       
    10   remove ....Mixfix
       
    11 *)
       
    12 
       
    13 signature MIXFIX0 =
       
    14 sig
       
    15   datatype mixfix =
       
    16     NoSyn |
       
    17     Mixfix of string * int list * int |
       
    18     Delimfix of string |
       
    19     Infixl of int |
       
    20     Infixr of int |
       
    21     Binder of string * int
       
    22   val max_pri: int
       
    23 end;
       
    24 
       
    25 signature MIXFIX1 =
       
    26 sig
       
    27   include MIXFIX0
       
    28   val type_name: string -> mixfix -> string
       
    29   val const_name: string -> mixfix -> string
       
    30   val pure_types: (string * int * mixfix) list
       
    31   val pure_syntax: (string * string * mixfix) list
       
    32 end;
       
    33 
       
    34 signature MIXFIX =
       
    35 sig
       
    36   include MIXFIX1
       
    37   structure SynExt: SYN_EXT
       
    38   local open SynExt in
       
    39     val syn_ext_types: string list -> (string * int * mixfix) list -> syn_ext
       
    40     val syn_ext_consts: string list -> (string * typ * mixfix) list -> syn_ext
       
    41   end
       
    42 end;
       
    43 
       
    44 functor MixfixFun(structure Lexicon: LEXICON and SynExt: SYN_EXT
       
    45   and SExtension: SEXTENSION): MIXFIX =
       
    46 struct
       
    47 
       
    48 structure SynExt = SynExt;
       
    49 open Lexicon SynExt SynExt.Ast SExtension;
       
    50 
       
    51 
       
    52 (** mixfix declarations **)
       
    53 
       
    54 datatype mixfix =
       
    55   NoSyn |
       
    56   Mixfix of string * int list * int |
       
    57   Delimfix of string |
       
    58   Infixl of int |
       
    59   Infixr of int |
       
    60   Binder of string * int;
       
    61 
       
    62 
       
    63 (* type / const names *)
       
    64 
       
    65 fun strip ("'" :: c :: cs) = c :: strip cs
       
    66   | strip ["'"] = []
       
    67   | strip (c :: cs) = c :: strip cs
       
    68   | strip [] = [];
       
    69 
       
    70 val strip_esc = implode o strip o explode;
       
    71 
       
    72 
       
    73 fun type_name t (Infixl _) = strip_esc t
       
    74   | type_name t (Infixr _) = strip_esc t
       
    75   | type_name t _ = t;
       
    76 
       
    77 fun infix_name c = "op " ^ strip_esc c;
       
    78 
       
    79 fun const_name c (Infixl _) = infix_name c
       
    80   | const_name c (Infixr _) = infix_name c
       
    81   | const_name c _ = c;
       
    82 
       
    83 
       
    84 (* syn_ext_types *)
       
    85 
       
    86 fun syn_ext_types all_roots type_decls =
       
    87   let
       
    88     fun name_of (t, _, mx) = type_name t mx;
       
    89 
       
    90     fun mk_infix t p1 p2 p3 =
       
    91       Mfix ("(_ " ^ t ^ "/ _)", [typeT, typeT] ---> typeT,
       
    92         strip_esc t, [p1, p2], p3);
       
    93 
       
    94     fun mfix_of (_, _, NoSyn) = None
       
    95       | mfix_of (t, 2, Infixl p) = Some (mk_infix t p (p + 1) p)
       
    96       | mfix_of (t, 2, Infixr p) = Some (mk_infix t (p + 1) p p)
       
    97       | mfix_of decl = error ("Bad mixfix declaration for type " ^
       
    98           quote (name_of decl));
       
    99 
       
   100     val mfix = mapfilter mfix_of type_decls;
       
   101     val xconsts = map name_of type_decls;
       
   102   in
       
   103     syn_ext all_roots [] mfix xconsts ([], [], [], []) ([], [])
       
   104   end;
       
   105 
       
   106 
       
   107 (* syn_ext_consts *)
       
   108 
       
   109 fun syn_ext_consts all_roots const_decls =
       
   110   let
       
   111     fun name_of (c, _, mx) = const_name c mx;
       
   112 
       
   113     fun mk_infix sy ty c p1 p2 p3 =
       
   114       [Mfix ("(_ " ^ sy ^ "/ _)", ty, c, [p1, p2], p3),
       
   115        Mfix ("op " ^ sy, ty, c, [], max_pri)];
       
   116 
       
   117     fun binder_typ _ (Type ("fun", [Type ("fun", [_, ty2]), ty3])) =
       
   118           [Type ("idts", []), ty2] ---> ty3
       
   119       | binder_typ c _ = error ("Bad type of binder " ^ quote c);
       
   120 
       
   121     fun mfix_of (_, _, NoSyn) = []
       
   122       | mfix_of (c, ty, Mixfix (sy, ps, p)) = [Mfix (sy, ty, c, ps, p)]
       
   123       | mfix_of (c, ty, Delimfix sy) = [Mfix (sy, ty, c, [], max_pri)]
       
   124       | mfix_of (sy, ty, Infixl p) = mk_infix sy ty (infix_name sy) p (p + 1) p
       
   125       | mfix_of (sy, ty, Infixr p) = mk_infix sy ty (infix_name sy) (p + 1) p p
       
   126       | mfix_of (c, ty, Binder (sy, p)) =
       
   127           [Mfix ("(3" ^ sy ^ "_./ _)", binder_typ c ty, sy, [], p)];
       
   128 
       
   129     fun binder (c, _, Binder (sy, _)) = Some (sy, c)
       
   130       | binder _ = None;
       
   131 
       
   132     val mfix = flat (map mfix_of const_decls);
       
   133     val xconsts = map name_of const_decls;
       
   134     val binders = mapfilter binder const_decls;
       
   135     val binder_trs = map mk_binder_tr binders;
       
   136     val binder_trs' = map (mk_binder_tr' o swap) binders;
       
   137   in
       
   138     syn_ext all_roots [] mfix xconsts ([], binder_trs, binder_trs', []) ([], [])
       
   139   end;
       
   140 
       
   141 
       
   142 
       
   143 (** Pure syntax **)
       
   144 
       
   145 val pure_types =
       
   146   map (fn t => (t, 0, NoSyn))
       
   147     (terminals @ [logic, "type", "types", "sort", "classes", args, "idt",
       
   148       "idts", "aprop", "asms"]);
       
   149 
       
   150 val pure_syntax =
       
   151  [("_lambda",   "[idts, 'a] => ('b => 'a)",       Mixfix ("(3%_./ _)", [], 0)),
       
   152   ("",          "'a => " ^ args,                  Delimfix "_"),
       
   153   ("_args",     "['a, " ^ args ^ "] => " ^ args,  Delimfix "_,/ _"),
       
   154   ("",          "id => idt",                      Delimfix "_"),
       
   155   ("_idtyp",    "[id, type] => idt",              Mixfix ("_::_", [], 0)),
       
   156   ("",          "idt => idt",                     Delimfix "'(_')"),
       
   157   ("",          "idt => idts",                    Delimfix "_"),
       
   158   ("_idts",     "[idt, idts] => idts",            Mixfix ("_/ _", [1, 0], 0)),
       
   159   ("",          "id => aprop",                    Delimfix "_"),
       
   160   ("",          "var => aprop",                   Delimfix "_"),
       
   161   (applC,       "[('b => 'a), " ^ args ^ "] => aprop",
       
   162     Mixfix ("(1_/(1'(_')))", [max_pri, 0], max_pri)),
       
   163   ("_aprop",    "aprop => prop",                  Delimfix "PROP _"),
       
   164   ("",          "prop => asms",                   Delimfix "_"),
       
   165   ("_asms",     "[prop, asms] => asms",           Delimfix "_;/ _"),
       
   166   ("_bigimpl",  "[asms, prop] => prop",           Mixfix ("((3[| _ |]) ==>/ _)", [0, 1], 1)),
       
   167   ("_insort",   "[type, 'a] => prop",             Delimfix "(2'(| _ :/ _ |'))"),
       
   168   ("_K",        "'a",                             NoSyn),
       
   169   ("_explode",  "'a",                             NoSyn),
       
   170   ("_implode",  "'a",                             NoSyn)];
       
   171 
       
   172 
       
   173 end;
       
   174