src/Pure/Syntax/printer.ML
author wenzelm
Mon, 31 May 2010 21:06:57 +0200
changeset 37216 3165bc303f66
parent 35545 e6c03f397eb8
child 38120 a902f158b4fc
permissions -rw-r--r--
modernized some structure names, keeping a few legacy aliases;
     1 (*  Title:      Pure/Syntax/printer.ML
     2     Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen
     3 
     4 Pretty printing of asts, terms, types and print (ast) translation.
     5 *)
     6 
     7 signature PRINTER0 =
     8 sig
     9   val show_brackets: bool Unsynchronized.ref
    10   val show_sorts: bool Unsynchronized.ref
    11   val show_types: bool Unsynchronized.ref
    12   val show_no_free_types: bool Unsynchronized.ref
    13   val show_all_types: bool Unsynchronized.ref
    14   val show_structs: bool Unsynchronized.ref
    15   val pp_show_brackets: Pretty.pp -> Pretty.pp
    16 end;
    17 
    18 signature PRINTER =
    19 sig
    20   include PRINTER0
    21   val sort_to_ast: Proof.context ->
    22     (string -> (Proof.context -> bool -> typ -> term list -> term) list) -> sort -> Ast.ast
    23   val typ_to_ast: Proof.context ->
    24     (string -> (Proof.context -> bool -> typ -> term list -> term) list) -> typ -> Ast.ast
    25   val term_to_ast: {structs: string list, fixes: string list} -> string list -> Proof.context ->
    26     (string -> (Proof.context -> bool -> typ -> term list -> term) list) -> term -> Ast.ast
    27   type prtabs
    28   val empty_prtabs: prtabs
    29   val update_prtabs: string -> Syn_Ext.xprod list -> prtabs -> prtabs
    30   val remove_prtabs: string -> Syn_Ext.xprod list -> prtabs -> prtabs
    31   val merge_prtabs: prtabs -> prtabs -> prtabs
    32   val pretty_term_ast: {extern_class: string -> xstring, extern_type: string -> xstring,
    33       extern_const: string -> xstring} -> Proof.context -> bool -> prtabs ->
    34     (string -> (Proof.context -> Ast.ast list -> Ast.ast) list) ->
    35     (string -> (Proof.context -> string -> Pretty.T) option) -> Ast.ast -> Pretty.T list
    36   val pretty_typ_ast: {extern_class: string -> xstring, extern_type: string -> xstring} ->
    37     Proof.context -> bool -> prtabs ->
    38     (string -> (Proof.context -> Ast.ast list -> Ast.ast) list) ->
    39     (string -> (Proof.context -> string -> Pretty.T) option) -> Ast.ast -> Pretty.T list
    40 end;
    41 
    42 structure Printer: PRINTER =
    43 struct
    44 
    45 
    46 (** options for printing **)
    47 
    48 val show_types = Unsynchronized.ref false;
    49 val show_sorts = Unsynchronized.ref false;
    50 val show_brackets = Unsynchronized.ref false;
    51 val show_no_free_types = Unsynchronized.ref false;
    52 val show_all_types = Unsynchronized.ref false;
    53 val show_structs = Unsynchronized.ref false;
    54 
    55 fun pp_show_brackets pp = Pretty.pp (setmp_CRITICAL show_brackets true (Pretty.term pp),
    56   Pretty.typ pp, Pretty.sort pp, Pretty.classrel pp, Pretty.arity pp);
    57 
    58 
    59 
    60 (** misc utils **)
    61 
    62 (* apply print (ast) translation function *)
    63 
    64 fun apply_trans ctxt fns args =
    65   let
    66     fun app_first [] = raise Match
    67       | app_first (f :: fs) = f ctxt args handle Match => app_first fs;
    68   in app_first fns end;
    69 
    70 fun apply_typed x y fs = map (fn f => fn ctxt => f ctxt x y) fs;
    71 
    72 
    73 (* simple_ast_of *)
    74 
    75 fun simple_ast_of (Const (c, _)) = Ast.Constant c
    76   | simple_ast_of (Free (x, _)) = Ast.Variable x
    77   | simple_ast_of (Var (xi, _)) = Ast.Variable (Term.string_of_vname xi)
    78   | simple_ast_of (t as _ $ _) =
    79       let val (f, args) = strip_comb t in
    80         Ast.mk_appl (simple_ast_of f) (map simple_ast_of args)
    81       end
    82   | simple_ast_of (Bound i) = Ast.Variable ("B." ^ string_of_int i)
    83   | simple_ast_of (Abs _) = sys_error "simple_ast_of: Abs";
    84 
    85 
    86 
    87 (** sort_to_ast, typ_to_ast **)
    88 
    89 fun ast_of_termT ctxt trf tm =
    90   let
    91     fun ast_of (t as Const ("_tfree", _) $ Free _) = simple_ast_of t
    92       | ast_of (t as Const ("_tvar", _) $ Var _) = simple_ast_of t
    93       | ast_of (Const (a, _)) = trans a []
    94       | ast_of (t as _ $ _) =
    95           (case strip_comb t of
    96             (Const (a, _), args) => trans a args
    97           | (f, args) => Ast.Appl (map ast_of (f :: args)))
    98       | ast_of t = simple_ast_of t
    99     and trans a args =
   100       ast_of (apply_trans ctxt (apply_typed false dummyT (trf a)) args)
   101         handle Match => Ast.mk_appl (Ast.Constant a) (map ast_of args);
   102   in ast_of tm end;
   103 
   104 fun sort_to_ast ctxt trf S = ast_of_termT ctxt trf (Type_Ext.term_of_sort S);
   105 fun typ_to_ast ctxt trf T = ast_of_termT ctxt trf (Type_Ext.term_of_typ (! show_sorts) T);
   106 
   107 
   108 
   109 (** term_to_ast **)
   110 
   111 fun ast_of_term idents consts ctxt trf
   112     show_all_types no_freeTs show_types show_sorts show_structs tm =
   113   let
   114     val {structs, fixes} = idents;
   115 
   116     fun mark_atoms ((t as Const (c, T)) $ u) =
   117           if member (op =) Syn_Ext.standard_token_markers c
   118           then t $ u else mark_atoms t $ mark_atoms u
   119       | mark_atoms (t $ u) = mark_atoms t $ mark_atoms u
   120       | mark_atoms (Abs (x, T, t)) = Abs (x, T, mark_atoms t)
   121       | mark_atoms (t as Const (c, T)) =
   122           if member (op =) consts c then t
   123           else Const (Lexicon.mark_const c, T)
   124       | mark_atoms (t as Free (x, T)) =
   125           let val i = find_index (fn s => s = x) structs + 1 in
   126             if i = 0 andalso member (op =) fixes x then
   127               Const (Lexicon.mark_fixed x, T)
   128             else if i = 1 andalso not show_structs then
   129               Lexicon.const "_struct" $ Lexicon.const "_indexdefault"
   130             else Lexicon.const "_free" $ t
   131           end
   132       | mark_atoms (t as Var (xi, T)) =
   133           if xi = Syn_Ext.dddot_indexname then Const ("_DDDOT", T)
   134           else Lexicon.const "_var" $ t
   135       | mark_atoms a = a;
   136 
   137     fun prune_typs (t_seen as (Const _, _)) = t_seen
   138       | prune_typs (t as Free (x, ty), seen) =
   139           if ty = dummyT then (t, seen)
   140           else if no_freeTs orelse member (op aconv) seen t then (Lexicon.free x, seen)
   141           else (t, t :: seen)
   142       | prune_typs (t as Var (xi, ty), seen) =
   143           if ty = dummyT then (t, seen)
   144           else if no_freeTs orelse member (op aconv) seen t then (Lexicon.var xi, seen)
   145           else (t, t :: seen)
   146       | prune_typs (t_seen as (Bound _, _)) = t_seen
   147       | prune_typs (Abs (x, ty, t), seen) =
   148           let val (t', seen') = prune_typs (t, seen);
   149           in (Abs (x, ty, t'), seen') end
   150       | prune_typs (t1 $ t2, seen) =
   151           let
   152             val (t1', seen') = prune_typs (t1, seen);
   153             val (t2', seen'') = prune_typs (t2, seen');
   154           in (t1' $ t2', seen'') end;
   155 
   156     fun ast_of tm =
   157       (case strip_comb tm of
   158         (t as Abs _, ts) => Ast.mk_appl (ast_of (Syn_Trans.abs_tr' t)) (map ast_of ts)
   159       | ((c as Const ("_free", _)), Free (x, T) :: ts) =>
   160           Ast.mk_appl (constrain (c $ Lexicon.free x) T) (map ast_of ts)
   161       | ((c as Const ("_var", _)), Var (xi, T) :: ts) =>
   162           Ast.mk_appl (constrain (c $ Lexicon.var xi) T) (map ast_of ts)
   163       | ((c as Const ("_bound", _)), Free (x, T) :: ts) =>
   164           Ast.mk_appl (constrain (c $ Lexicon.free x) T) (map ast_of ts)
   165       | (Const ("_idtdummy", T), ts) =>
   166           Ast.mk_appl (constrain (Lexicon.const "_idtdummy") T) (map ast_of ts)
   167       | (const as Const (c, T), ts) =>
   168           if show_all_types
   169           then Ast.mk_appl (constrain const T) (map ast_of ts)
   170           else trans c T ts
   171       | (t, ts) => Ast.mk_appl (simple_ast_of t) (map ast_of ts))
   172 
   173     and trans a T args =
   174       ast_of (apply_trans ctxt (apply_typed show_sorts T (trf a)) args)
   175         handle Match => Ast.mk_appl (Ast.Constant a) (map ast_of args)
   176 
   177     and constrain t T =
   178       if show_types andalso T <> dummyT then
   179         Ast.Appl [Ast.Constant Syn_Ext.constrainC, simple_ast_of t,
   180           ast_of_termT ctxt trf (Type_Ext.term_of_typ show_sorts T)]
   181       else simple_ast_of t;
   182   in
   183     tm
   184     |> Syn_Trans.prop_tr'
   185     |> show_types ? (#1 o prune_typs o rpair [])
   186     |> mark_atoms
   187     |> ast_of
   188   end;
   189 
   190 fun term_to_ast idents consts ctxt trf tm =
   191   ast_of_term idents consts ctxt trf (! show_all_types) (! show_no_free_types)
   192     (! show_types orelse ! show_sorts orelse ! show_all_types) (! show_sorts) (! show_structs) tm;
   193 
   194 
   195 
   196 (** type prtabs **)
   197 
   198 datatype symb =
   199   Arg of int |
   200   TypArg of int |
   201   String of string |
   202   Space of string |
   203   Break of int |
   204   Block of int * symb list;
   205 
   206 type prtabs = (string * ((symb list * int * int) list) Symtab.table) list;
   207 
   208 fun mode_tab prtabs mode = the_default Symtab.empty (AList.lookup (op =) prtabs mode);
   209 fun mode_tabs prtabs modes = map_filter (AList.lookup (op =) prtabs) (modes @ [""]);
   210 
   211 
   212 (* xprod_to_fmt *)
   213 
   214 fun xprod_to_fmt (Syn_Ext.XProd (_, _, "", _)) = NONE
   215   | xprod_to_fmt (Syn_Ext.XProd (_, xsymbs, const, pri)) =
   216       let
   217         fun arg (s, p) =
   218           (if s = "type" then TypArg else Arg)
   219           (if Lexicon.is_terminal s then Syn_Ext.max_pri else p);
   220 
   221         fun xsyms_to_syms (Syn_Ext.Delim s :: xsyms) =
   222               apfst (cons (String s)) (xsyms_to_syms xsyms)
   223           | xsyms_to_syms (Syn_Ext.Argument s_p :: xsyms) =
   224               apfst (cons (arg s_p)) (xsyms_to_syms xsyms)
   225           | xsyms_to_syms (Syn_Ext.Space s :: xsyms) =
   226               apfst (cons (Space s)) (xsyms_to_syms xsyms)
   227           | xsyms_to_syms (Syn_Ext.Bg i :: xsyms) =
   228               let
   229                 val (bsyms, xsyms') = xsyms_to_syms xsyms;
   230                 val (syms, xsyms'') = xsyms_to_syms xsyms';
   231               in
   232                 (Block (i, bsyms) :: syms, xsyms'')
   233               end
   234           | xsyms_to_syms (Syn_Ext.Brk i :: xsyms) =
   235               apfst (cons (Break i)) (xsyms_to_syms xsyms)
   236           | xsyms_to_syms (Syn_Ext.En :: xsyms) = ([], xsyms)
   237           | xsyms_to_syms [] = ([], []);
   238 
   239         fun nargs (Arg _ :: syms) = nargs syms + 1
   240           | nargs (TypArg _ :: syms) = nargs syms + 1
   241           | nargs (String _ :: syms) = nargs syms
   242           | nargs (Space _ :: syms) = nargs syms
   243           | nargs (Break _ :: syms) = nargs syms
   244           | nargs (Block (_, bsyms) :: syms) = nargs syms + nargs bsyms
   245           | nargs [] = 0;
   246       in
   247         (case xsyms_to_syms xsymbs of
   248           (symbs, []) => SOME (const, (symbs, nargs symbs, pri))
   249         | _ => sys_error "xprod_to_fmt: unbalanced blocks")
   250       end;
   251 
   252 
   253 (* empty, extend, merge prtabs *)
   254 
   255 val empty_prtabs = [];
   256 
   257 fun update_prtabs mode xprods prtabs =
   258   let
   259     val fmts = map_filter xprod_to_fmt xprods;
   260     val tab' = fold (Symtab.update_list (op =)) fmts (mode_tab prtabs mode);
   261   in AList.update (op =) (mode, tab') prtabs end;
   262 
   263 fun remove_prtabs mode xprods prtabs =
   264   let
   265     val tab = mode_tab prtabs mode;
   266     val fmts = map_filter (fn xprod as Syn_Ext.XProd (_, _, c, _) =>
   267       if null (Symtab.lookup_list tab c) then NONE
   268       else xprod_to_fmt xprod) xprods;
   269     val tab' = fold (Symtab.remove_list (op =)) fmts tab;
   270   in AList.update (op =) (mode, tab') prtabs end;
   271 
   272 fun merge_prtabs prtabs1 prtabs2 =
   273   let
   274     val modes = distinct (op =) (map fst (prtabs1 @ prtabs2));
   275     fun merge m = (m, Symtab.merge_list (op =) (mode_tab prtabs1 m, mode_tab prtabs2 m));
   276   in map merge modes end;
   277 
   278 
   279 
   280 (** pretty term or typ asts **)
   281 
   282 fun is_chain [Block (_, pr)] = is_chain pr
   283   | is_chain [Arg _] = true
   284   | is_chain _  = false;
   285 
   286 fun pretty extern ctxt tabs trf tokentrf type_mode curried ast0 p0 =
   287   let
   288     val {extern_class, extern_type, extern_const} = extern;
   289 
   290     fun token_trans a x =
   291       (case tokentrf a of
   292         NONE =>
   293           if member (op =) Syn_Ext.standard_token_classes a
   294           then SOME (Pretty.str x) else NONE
   295       | SOME f => SOME (f ctxt x));
   296 
   297     (*default applications: prefix / postfix*)
   298     val appT =
   299       if type_mode then Type_Ext.tappl_ast_tr'
   300       else if curried then Syn_Trans.applC_ast_tr'
   301       else Syn_Trans.appl_ast_tr';
   302 
   303     fun synT _ ([], args) = ([], args)
   304       | synT markup (Arg p :: symbs, t :: args) =
   305           let val (Ts, args') = synT markup (symbs, args);
   306           in (astT (t, p) @ Ts, args') end
   307       | synT markup (TypArg p :: symbs, t :: args) =
   308           let
   309             val (Ts, args') = synT markup (symbs, args);
   310           in
   311             if type_mode then (astT (t, p) @ Ts, args')
   312             else (pretty extern ctxt tabs trf tokentrf true curried t p @ Ts, args')
   313           end
   314       | synT markup (String s :: symbs, args) =
   315           let val (Ts, args') = synT markup (symbs, args);
   316           in (markup (Pretty.str s) :: Ts, args') end
   317       | synT markup (Space s :: symbs, args) =
   318           let val (Ts, args') = synT markup (symbs, args);
   319           in (Pretty.str s :: Ts, args') end
   320       | synT markup (Block (i, bsymbs) :: symbs, args) =
   321           let
   322             val (bTs, args') = synT markup (bsymbs, args);
   323             val (Ts, args'') = synT markup (symbs, args');
   324             val T =
   325               if i < 0 then Pretty.unbreakable (Pretty.block bTs)
   326               else Pretty.blk (i, bTs);
   327           in (T :: Ts, args'') end
   328       | synT markup (Break i :: symbs, args) =
   329           let
   330             val (Ts, args') = synT markup (symbs, args);
   331             val T = if i < 0 then Pretty.fbrk else Pretty.brk i;
   332           in (T :: Ts, args') end
   333 
   334     and parT markup (pr, args, p, p': int) = #1 (synT markup
   335           (if p > p' orelse
   336             (! show_brackets andalso p' <> Syn_Ext.max_pri andalso not (is_chain pr))
   337             then [Block (1, Space "(" :: pr @ [Space ")"])]
   338             else pr, args))
   339 
   340     and atomT a = a |> Lexicon.unmark
   341      {case_class = fn c => Pretty.mark (Markup.tclass c) (Pretty.str (extern_class c)),
   342       case_type = fn c => Pretty.mark (Markup.tycon c) (Pretty.str (extern_type c)),
   343       case_const = fn c => Pretty.mark (Markup.const c) (Pretty.str (extern_const c)),
   344       case_fixed = fn x => the (token_trans "_free" x),
   345       case_default = Pretty.str}
   346 
   347     and prefixT (_, a, [], _) = [atomT a]
   348       | prefixT (c, _, args, p) = astT (appT (c, args), p)
   349 
   350     and splitT 0 ([x], ys) = (x, ys)
   351       | splitT 0 (rev_xs, ys) = (Ast.Appl (rev rev_xs), ys)
   352       | splitT n (rev_xs, y :: ys) = splitT (n - 1) (y :: rev_xs, ys)
   353 
   354     and combT (tup as (c, a, args, p)) =
   355       let
   356         val nargs = length args;
   357         val markup = a |> Lexicon.unmark
   358          {case_class = Pretty.mark o Markup.tclass,
   359           case_type = Pretty.mark o Markup.tycon,
   360           case_const = Pretty.mark o Markup.const,
   361           case_fixed = Pretty.mark o Markup.fixed,
   362           case_default = K I};
   363 
   364         (*find matching table entry, or print as prefix / postfix*)
   365         fun prnt ([], []) = prefixT tup
   366           | prnt ([], tb :: tbs) = prnt (Symtab.lookup_list tb a, tbs)
   367           | prnt ((pr, n, p') :: prnps, tbs) =
   368               if nargs = n then parT markup (pr, args, p, p')
   369               else if nargs > n andalso not type_mode then
   370                 astT (appT (splitT n ([c], args)), p)
   371               else prnt (prnps, tbs);
   372       in
   373         (case tokentrT a args of
   374           SOME prt => [prt]
   375         | NONE => astT (apply_trans ctxt (trf a) args, p) handle Match => prnt ([], tabs))
   376       end
   377 
   378     and tokentrT a [Ast.Variable x] = token_trans a x
   379       | tokentrT _ _ = NONE
   380 
   381     and astT (c as Ast.Constant a, p) = combT (c, a, [], p)
   382       | astT (Ast.Variable x, _) = [Pretty.str x]
   383       | astT (Ast.Appl ((c as Ast.Constant a) :: (args as _ :: _)), p) = combT (c, a, args, p)
   384       | astT (Ast.Appl (f :: (args as _ :: _)), p) = astT (appT (f, args), p)
   385       | astT (ast as Ast.Appl _, _) = raise Ast.AST ("pretty: malformed ast", [ast]);
   386   in astT (ast0, p0) end;
   387 
   388 
   389 (* pretty_term_ast *)
   390 
   391 fun pretty_term_ast extern ctxt curried prtabs trf tokentrf ast =
   392   pretty extern ctxt (mode_tabs prtabs (print_mode_value ()))
   393     trf tokentrf false curried ast 0;
   394 
   395 
   396 (* pretty_typ_ast *)
   397 
   398 fun pretty_typ_ast {extern_class, extern_type} ctxt _ prtabs trf tokentrf ast =
   399   pretty {extern_class = extern_class, extern_type = extern_type, extern_const = I}
   400     ctxt (mode_tabs prtabs (print_mode_value ()))
   401     trf tokentrf true false ast 0;
   402 
   403 end;