| author | wenzelm | 
| Thu, 14 Feb 2008 21:33:44 +0100 | |
| changeset 26069 | 321c4ca82923 | 
| parent 25393 | 0856e0141062 | 
| child 26707 | ddf6bab64b96 | 
| permissions | -rw-r--r-- | 
| 18 | 1 | (* Title: Pure/Syntax/printer.ML | 
| 0 | 2 | ID: $Id$ | 
| 3 | Author: Tobias Nipkow and Markus Wenzel, TU Muenchen | |
| 18 | 4 | |
| 5 | Pretty printing of asts, terms, types and print (ast) translation. | |
| 0 | 6 | *) | 
| 7 | ||
| 8 | signature PRINTER0 = | |
| 2384 | 9 | sig | 
| 504 
a4f09493d929
added flag show_brackets for printinmg fully bracketed terms.
 nipkow parents: 
381diff
changeset | 10 | val show_brackets: bool ref | 
| 
a4f09493d929
added flag show_brackets for printinmg fully bracketed terms.
 nipkow parents: 
381diff
changeset | 11 | val show_sorts: bool ref | 
| 0 | 12 | val show_types: bool ref | 
| 617 | 13 | val show_no_free_types: bool ref | 
| 14176 
716fec31f9ea
Added show_all_types flag, such that all type information in the term
 skalberg parents: 
12785diff
changeset | 14 | val show_all_types: bool ref | 
| 14837 
827c68f8267c
added pp_show_brackets; support unbreakable blocks;
 wenzelm parents: 
14783diff
changeset | 15 | val pp_show_brackets: Pretty.pp -> Pretty.pp | 
| 2384 | 16 | end; | 
| 0 | 17 | |
| 18 | signature PRINTER = | |
| 2384 | 19 | sig | 
| 0 | 20 | include PRINTER0 | 
| 21772 | 21 | val term_to_ast: Proof.context -> | 
| 22 | (string -> (Proof.context -> bool -> typ -> term list -> term) list) -> term -> Ast.ast | |
| 23 | val typ_to_ast: Proof.context -> | |
| 24 | (string -> (Proof.context -> bool -> typ -> term list -> term) list) -> typ -> Ast.ast | |
| 25 | val sort_to_ast: Proof.context -> | |
| 26 | (string -> (Proof.context -> bool -> typ -> term list -> term) list) -> sort -> Ast.ast | |
| 2200 
2538977e94fa
added print_mode: string list ref (order of printer tables);
 wenzelm parents: 
1509diff
changeset | 27 | type prtabs | 
| 
2538977e94fa
added print_mode: string list ref (order of printer tables);
 wenzelm parents: 
1509diff
changeset | 28 | val empty_prtabs: prtabs | 
| 25393 
0856e0141062
replaced extend_prtabs by update_prtabs (absorb duplicates);
 wenzelm parents: 
25386diff
changeset | 29 | val update_prtabs: string -> SynExt.xprod list -> prtabs -> prtabs | 
| 15753 | 30 | val remove_prtabs: string -> SynExt.xprod list -> prtabs -> prtabs | 
| 2200 
2538977e94fa
added print_mode: string list ref (order of printer tables);
 wenzelm parents: 
1509diff
changeset | 31 | val merge_prtabs: prtabs -> prtabs -> prtabs | 
| 21772 | 32 | val pretty_term_ast: (string -> xstring) -> Proof.context -> bool -> prtabs | 
| 33 | -> (string -> (Proof.context -> Ast.ast list -> Ast.ast) list) | |
| 23660 
18765718cf62
type output = string indicates raw system output;
 wenzelm parents: 
23630diff
changeset | 34 | -> (string -> (string -> output * int) option) -> Ast.ast -> Pretty.T list | 
| 21772 | 35 | val pretty_typ_ast: Proof.context -> bool -> prtabs | 
| 36 | -> (string -> (Proof.context -> Ast.ast list -> Ast.ast) list) | |
| 23660 
18765718cf62
type output = string indicates raw system output;
 wenzelm parents: 
23630diff
changeset | 37 | -> (string -> (string -> output * int) option) -> Ast.ast -> Pretty.T list | 
| 2384 | 38 | end; | 
| 0 | 39 | |
| 2365 | 40 | structure Printer: PRINTER = | 
| 0 | 41 | struct | 
| 2200 
2538977e94fa
added print_mode: string list ref (order of printer tables);
 wenzelm parents: 
1509diff
changeset | 42 | |
| 
2538977e94fa
added print_mode: string list ref (order of printer tables);
 wenzelm parents: 
1509diff
changeset | 43 | |
| 0 | 44 | (** options for printing **) | 
| 45 | ||
| 46 | val show_types = ref false; | |
| 47 | val show_sorts = ref false; | |
| 504 
a4f09493d929
added flag show_brackets for printinmg fully bracketed terms.
 nipkow parents: 
381diff
changeset | 48 | val show_brackets = ref false; | 
| 617 | 49 | val show_no_free_types = ref false; | 
| 14176 
716fec31f9ea
Added show_all_types flag, such that all type information in the term
 skalberg parents: 
12785diff
changeset | 50 | val show_all_types = ref false; | 
| 0 | 51 | |
| 14975 | 52 | fun pp_show_brackets pp = Pretty.pp (setmp show_brackets true (Pretty.term pp), | 
| 53 | Pretty.typ pp, Pretty.sort pp, Pretty.classrel pp, Pretty.arity pp); | |
| 14837 
827c68f8267c
added pp_show_brackets; support unbreakable blocks;
 wenzelm parents: 
14783diff
changeset | 54 | |
| 
827c68f8267c
added pp_show_brackets; support unbreakable blocks;
 wenzelm parents: 
14783diff
changeset | 55 | |
| 0 | 56 | |
| 2701 | 57 | (** misc utils **) | 
| 58 | ||
| 59 | (* apply print (ast) translation function *) | |
| 0 | 60 | |
| 23937 
66e1f24d655d
eliminated transform_failure (to avoid critical section for main transactions);
 wenzelm parents: 
23660diff
changeset | 61 | fun apply_trans ctxt fns args = | 
| 16611 
edb368e2878f
proper treatment of advanced trfuns: pass thy argument;
 wenzelm parents: 
15973diff
changeset | 62 | let | 
| 
edb368e2878f
proper treatment of advanced trfuns: pass thy argument;
 wenzelm parents: 
15973diff
changeset | 63 | fun app_first [] = raise Match | 
| 21772 | 64 | | app_first (f :: fs) = f ctxt args handle Match => app_first fs; | 
| 23937 
66e1f24d655d
eliminated transform_failure (to avoid critical section for main transactions);
 wenzelm parents: 
23660diff
changeset | 65 | in app_first fns end; | 
| 5691 | 66 | |
| 21772 | 67 | fun apply_typed x y fs = map (fn f => fn ctxt => f ctxt x y) fs; | 
| 4147 | 68 | |
| 0 | 69 | |
| 2701 | 70 | (* simple_ast_of *) | 
| 71 | ||
| 5691 | 72 | fun simple_ast_of (Const (c, _)) = Ast.Constant c | 
| 73 | | simple_ast_of (Free (x, _)) = Ast.Variable x | |
| 14783 | 74 | | simple_ast_of (Var (xi, _)) = Ast.Variable (Term.string_of_vname xi) | 
| 2701 | 75 | | simple_ast_of (t as _ $ _) = | 
| 76 | let val (f, args) = strip_comb t in | |
| 5691 | 77 | Ast.mk_appl (simple_ast_of f) (map simple_ast_of args) | 
| 2701 | 78 | end | 
| 5691 | 79 |   | simple_ast_of (Bound i) = Ast.Variable ("B." ^ string_of_int i)
 | 
| 2701 | 80 | | simple_ast_of (Abs _) = sys_error "simple_ast_of: Abs"; | 
| 81 | ||
| 82 | ||
| 83 | ||
| 3776 | 84 | (** sort_to_ast, typ_to_ast **) | 
| 2701 | 85 | |
| 21772 | 86 | fun ast_of_termT ctxt trf tm = | 
| 0 | 87 | let | 
| 2701 | 88 |     fun ast_of (t as Const ("_class", _) $ Free _) = simple_ast_of t
 | 
| 89 |       | ast_of (t as Const ("_tfree", _) $ Free _) = simple_ast_of t
 | |
| 90 |       | ast_of (t as Const ("_tvar", _) $ Var _) = simple_ast_of t
 | |
| 91 | | ast_of (Const (a, _)) = trans a [] | |
| 92 | | ast_of (t as _ $ _) = | |
| 93 | (case strip_comb t of | |
| 94 | (Const (a, _), args) => trans a args | |
| 5691 | 95 | | (f, args) => Ast.Appl (map ast_of (f :: args))) | 
| 96 | | ast_of t = simple_ast_of t | |
| 2701 | 97 | and trans a args = | 
| 23937 
66e1f24d655d
eliminated transform_failure (to avoid critical section for main transactions);
 wenzelm parents: 
23660diff
changeset | 98 | ast_of (apply_trans ctxt (apply_typed false dummyT (trf a)) args) | 
| 5691 | 99 | handle Match => Ast.mk_appl (Ast.Constant a) (map ast_of args); | 
| 100 | in ast_of tm end; | |
| 617 | 101 | |
| 21772 | 102 | fun sort_to_ast ctxt trf S = ast_of_termT ctxt trf (TypeExt.term_of_sort S); | 
| 103 | fun typ_to_ast ctxt trf T = ast_of_termT ctxt trf (TypeExt.term_of_typ (! show_sorts) T); | |
| 2701 | 104 | |
| 105 | ||
| 106 | ||
| 107 | (** term_to_ast **) | |
| 108 | ||
| 11795 | 109 | fun mark_freevars ((t as Const (c, _)) $ u) = | 
| 18139 | 110 | if member (op =) SynExt.standard_token_markers c then (t $ u) | 
| 11795 | 111 | else t $ mark_freevars u | 
| 112 | | mark_freevars (t $ u) = mark_freevars t $ mark_freevars u | |
| 2701 | 113 | | mark_freevars (Abs (x, T, t)) = Abs (x, T, mark_freevars t) | 
| 5691 | 114 | | mark_freevars (t as Free _) = Lexicon.const "_free" $ t | 
| 6767 | 115 | | mark_freevars (t as Var (xi, T)) = | 
| 116 |       if xi = SynExt.dddot_indexname then Const ("_DDDOT", T)
 | |
| 117 | else Lexicon.const "_var" $ t | |
| 2701 | 118 | | mark_freevars a = a; | 
| 119 | ||
| 21772 | 120 | fun ast_of_term ctxt trf show_all_types no_freeTs show_types show_sorts tm = | 
| 2701 | 121 | let | 
| 0 | 122 | fun prune_typs (t_seen as (Const _, _)) = t_seen | 
| 123 | | prune_typs (t as Free (x, ty), seen) = | |
| 124 | if ty = dummyT then (t, seen) | |
| 18139 | 125 | else if no_freeTs orelse member (op aconv) seen t then (Lexicon.free x, seen) | 
| 0 | 126 | else (t, t :: seen) | 
| 127 | | prune_typs (t as Var (xi, ty), seen) = | |
| 128 | if ty = dummyT then (t, seen) | |
| 18139 | 129 | else if no_freeTs orelse member (op aconv) seen t then (Lexicon.var xi, seen) | 
| 0 | 130 | else (t, t :: seen) | 
| 131 | | prune_typs (t_seen as (Bound _, _)) = t_seen | |
| 132 | | prune_typs (Abs (x, ty, t), seen) = | |
| 133 | let val (t', seen') = prune_typs (t, seen); | |
| 134 | in (Abs (x, ty, t'), seen') end | |
| 135 | | prune_typs (t1 $ t2, seen) = | |
| 136 | let | |
| 137 | val (t1', seen') = prune_typs (t1, seen); | |
| 138 | val (t2', seen'') = prune_typs (t2, seen'); | |
| 6767 | 139 | in (t1' $ t2', seen'') end; | 
| 0 | 140 | |
| 2701 | 141 | fun ast_of tm = | 
| 142 | (case strip_comb tm of | |
| 5691 | 143 | (t as Abs _, ts) => Ast.mk_appl (ast_of (SynTrans.abs_tr' t)) (map ast_of ts) | 
| 2701 | 144 |       | ((c as Const ("_free", _)), Free (x, T) :: ts) =>
 | 
| 5691 | 145 | Ast.mk_appl (constrain (c $ Lexicon.free x) T) (map ast_of ts) | 
| 21748 | 146 |       | ((c as Const ("_var", _)), Var (xi, T) :: ts) =>
 | 
| 147 | Ast.mk_appl (constrain (c $ Lexicon.var xi) T) (map ast_of ts) | |
| 2701 | 148 |       | ((c as Const ("_bound", _)), Free (x, T) :: ts) =>
 | 
| 5691 | 149 | Ast.mk_appl (constrain (c $ Lexicon.free x) T) (map ast_of ts) | 
| 21748 | 150 |       | (Const ("_idtdummy", T), ts) =>
 | 
| 151 | Ast.mk_appl (constrain (Lexicon.const "_idtdummy") T) (map ast_of ts) | |
| 14176 
716fec31f9ea
Added show_all_types flag, such that all type information in the term
 skalberg parents: 
12785diff
changeset | 152 | | (c' as Const (c, T), ts) => | 
| 14696 | 153 | if show_all_types | 
| 154 | then Ast.mk_appl (constrain c' T) (map ast_of ts) | |
| 155 | else trans c T ts | |
| 5691 | 156 | | (t, ts) => Ast.mk_appl (simple_ast_of t) (map ast_of ts)) | 
| 0 | 157 | |
| 2384 | 158 | and trans a T args = | 
| 23937 
66e1f24d655d
eliminated transform_failure (to avoid critical section for main transactions);
 wenzelm parents: 
23660diff
changeset | 159 | ast_of (apply_trans ctxt (apply_typed show_sorts T (trf a)) args) | 
| 5691 | 160 | handle Match => Ast.mk_appl (Ast.Constant a) (map ast_of args) | 
| 0 | 161 | |
| 2701 | 162 | and constrain t T = | 
| 163 | if show_types andalso T <> dummyT then | |
| 5691 | 164 | Ast.Appl [Ast.Constant SynExt.constrainC, simple_ast_of t, | 
| 21772 | 165 | ast_of_termT ctxt trf (TypeExt.term_of_typ show_sorts T)] | 
| 2701 | 166 | else simple_ast_of t | 
| 0 | 167 | in | 
| 2701 | 168 | tm | 
| 5691 | 169 | |> SynTrans.prop_tr' | 
| 2701 | 170 | |> (if show_types then #1 o prune_typs o rpair [] else I) | 
| 171 | |> mark_freevars | |
| 172 | |> ast_of | |
| 0 | 173 | end; | 
| 174 | ||
| 21772 | 175 | fun term_to_ast ctxt trf tm = | 
| 176 | ast_of_term ctxt trf (! show_all_types) (! show_no_free_types) | |
| 14696 | 177 | (! show_types orelse ! show_sorts orelse ! show_all_types) (! show_sorts) tm; | 
| 0 | 178 | |
| 179 | ||
| 180 | ||
| 2200 
2538977e94fa
added print_mode: string list ref (order of printer tables);
 wenzelm parents: 
1509diff
changeset | 181 | (** type prtabs **) | 
| 0 | 182 | |
| 183 | datatype symb = | |
| 184 | Arg of int | | |
| 185 | TypArg of int | | |
| 186 | String of string | | |
| 23630 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 187 | Space of string | | 
| 0 | 188 | Break of int | | 
| 189 | Block of int * symb list; | |
| 190 | ||
| 2200 
2538977e94fa
added print_mode: string list ref (order of printer tables);
 wenzelm parents: 
1509diff
changeset | 191 | type prtabs = (string * ((symb list * int * int) list) Symtab.table) list; | 
| 
2538977e94fa
added print_mode: string list ref (order of printer tables);
 wenzelm parents: 
1509diff
changeset | 192 | |
| 18957 | 193 | fun mode_tab prtabs mode = the_default Symtab.empty (AList.lookup (op =) prtabs mode); | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19374diff
changeset | 194 | fun mode_tabs prtabs modes = map_filter (AList.lookup (op =) prtabs) (modes @ [""]); | 
| 2200 
2538977e94fa
added print_mode: string list ref (order of printer tables);
 wenzelm parents: 
1509diff
changeset | 195 | |
| 18 | 196 | |
| 3816 | 197 | (* xprod_to_fmt *) | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 198 | |
| 15531 | 199 | fun xprod_to_fmt (SynExt.XProd (_, _, "", _)) = NONE | 
| 5691 | 200 | | xprod_to_fmt (SynExt.XProd (_, xsymbs, const, pri)) = | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 201 | let | 
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 202 | fun arg (s, p) = | 
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 203 | (if s = "type" then TypArg else Arg) | 
| 5691 | 204 | (if Lexicon.is_terminal s then SynExt.max_pri else p); | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 205 | |
| 5691 | 206 | fun xsyms_to_syms (SynExt.Delim s :: xsyms) = | 
| 23630 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 207 | apfst (cons (String s)) (xsyms_to_syms xsyms) | 
| 5691 | 208 | | xsyms_to_syms (SynExt.Argument s_p :: xsyms) = | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 209 | apfst (cons (arg s_p)) (xsyms_to_syms xsyms) | 
| 5691 | 210 | | xsyms_to_syms (SynExt.Space s :: xsyms) = | 
| 23630 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 211 | apfst (cons (Space s)) (xsyms_to_syms xsyms) | 
| 5691 | 212 | | xsyms_to_syms (SynExt.Bg i :: xsyms) = | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 213 | let | 
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 214 | val (bsyms, xsyms') = xsyms_to_syms xsyms; | 
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 215 | val (syms, xsyms'') = xsyms_to_syms xsyms'; | 
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 216 | in | 
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 217 | (Block (i, bsyms) :: syms, xsyms'') | 
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 218 | end | 
| 5691 | 219 | | xsyms_to_syms (SynExt.Brk i :: xsyms) = | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 220 | apfst (cons (Break i)) (xsyms_to_syms xsyms) | 
| 5691 | 221 | | xsyms_to_syms (SynExt.En :: xsyms) = ([], xsyms) | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 222 | | xsyms_to_syms [] = ([], []); | 
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 223 | |
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 224 | fun nargs (Arg _ :: syms) = nargs syms + 1 | 
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 225 | | nargs (TypArg _ :: syms) = nargs syms + 1 | 
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 226 | | nargs (String _ :: syms) = nargs syms | 
| 23630 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 227 | | nargs (Space _ :: syms) = nargs syms | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 228 | | nargs (Break _ :: syms) = nargs syms | 
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 229 | | nargs (Block (_, bsyms) :: syms) = nargs syms + nargs bsyms | 
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 230 | | nargs [] = 0; | 
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 231 | in | 
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 232 | (case xsyms_to_syms xsymbs of | 
| 15531 | 233 | (symbs, []) => SOME (const, (symbs, nargs symbs, pri)) | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 234 | | _ => sys_error "xprod_to_fmt: unbalanced blocks") | 
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 235 | end; | 
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 236 | |
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 237 | |
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 238 | (* empty, extend, merge prtabs *) | 
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 239 | |
| 2200 
2538977e94fa
added print_mode: string list ref (order of printer tables);
 wenzelm parents: 
1509diff
changeset | 240 | val empty_prtabs = []; | 
| 18 | 241 | |
| 25393 
0856e0141062
replaced extend_prtabs by update_prtabs (absorb duplicates);
 wenzelm parents: 
25386diff
changeset | 242 | fun update_prtabs mode xprods prtabs = | 
| 2200 
2538977e94fa
added print_mode: string list ref (order of printer tables);
 wenzelm parents: 
1509diff
changeset | 243 | let | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19374diff
changeset | 244 | val fmts = map_filter xprod_to_fmt xprods; | 
| 25393 
0856e0141062
replaced extend_prtabs by update_prtabs (absorb duplicates);
 wenzelm parents: 
25386diff
changeset | 245 | val tab' = fold (Symtab.update_list (op =)) fmts (mode_tab prtabs mode); | 
| 25386 | 246 | in AList.update (op =) (mode, tab') prtabs end; | 
| 15753 | 247 | |
| 25386 | 248 | fun remove_prtabs mode xprods prtabs = | 
| 249 | let | |
| 250 | val tab = mode_tab prtabs mode; | |
| 251 | val fmts = map_filter (fn xprod as SynExt.XProd (_, _, c, _) => | |
| 252 | if null (Symtab.lookup_list tab c) then NONE | |
| 253 | else xprod_to_fmt xprod) xprods; | |
| 25393 
0856e0141062
replaced extend_prtabs by update_prtabs (absorb duplicates);
 wenzelm parents: 
25386diff
changeset | 254 | val tab' = fold (Symtab.remove_list (op =)) fmts tab; | 
| 
0856e0141062
replaced extend_prtabs by update_prtabs (absorb duplicates);
 wenzelm parents: 
25386diff
changeset | 255 | in AList.update (op =) (mode, tab') prtabs end; | 
| 0 | 256 | |
| 2200 
2538977e94fa
added print_mode: string list ref (order of printer tables);
 wenzelm parents: 
1509diff
changeset | 257 | fun merge_prtabs prtabs1 prtabs2 = | 
| 
2538977e94fa
added print_mode: string list ref (order of printer tables);
 wenzelm parents: 
1509diff
changeset | 258 | let | 
| 19046 
bc5c6c9b114e
removed distinct, renamed gen_distinct to distinct;
 wenzelm parents: 
18977diff
changeset | 259 | val modes = distinct (op =) (map fst (prtabs1 @ prtabs2)); | 
| 18957 | 260 | fun merge m = (m, Symtab.merge_list (op =) (mode_tab prtabs1 m, mode_tab prtabs2 m)); | 
| 12292 | 261 | in map merge modes end; | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 262 | |
| 0 | 263 | |
| 264 | ||
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 265 | (** pretty term or typ asts **) | 
| 0 | 266 | |
| 2200 
2538977e94fa
added print_mode: string list ref (order of printer tables);
 wenzelm parents: 
1509diff
changeset | 267 | fun is_chain [Block (_, pr)] = is_chain pr | 
| 
2538977e94fa
added print_mode: string list ref (order of printer tables);
 wenzelm parents: 
1509diff
changeset | 268 | | is_chain [Arg _] = true | 
| 
2538977e94fa
added print_mode: string list ref (order of printer tables);
 wenzelm parents: 
1509diff
changeset | 269 | | is_chain _ = false; | 
| 506 
e0ca460d6e51
improved show_brackets again - Trueprop does not create () any more.
 nipkow parents: 
505diff
changeset | 270 | |
| 23630 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 271 | fun const_markup c = Pretty.markup (Markup.const c) o single; | 
| 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 272 | |
| 21772 | 273 | fun pretty extern_const ctxt tabs trf tokentrf type_mode curried ast0 p0 = | 
| 0 | 274 | let | 
| 5691 | 275 | fun token_trans c [Ast.Variable x] = | 
| 2701 | 276 | (case tokentrf c of | 
| 15531 | 277 | NONE => NONE | 
| 278 | | SOME f => SOME (f x)) | |
| 279 | | token_trans _ _ = NONE; | |
| 2701 | 280 | |
| 2200 
2538977e94fa
added print_mode: string list ref (order of printer tables);
 wenzelm parents: 
1509diff
changeset | 281 | (*default applications: prefix / postfix*) | 
| 
2538977e94fa
added print_mode: string list ref (order of printer tables);
 wenzelm parents: 
1509diff
changeset | 282 | val appT = | 
| 5691 | 283 | if type_mode then TypeExt.tappl_ast_tr' | 
| 284 | else if curried then SynTrans.applC_ast_tr' | |
| 285 | else SynTrans.appl_ast_tr'; | |
| 18 | 286 | |
| 23630 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 287 | fun synT _ ([], args) = ([], args) | 
| 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 288 | | synT markup (Arg p :: symbs, t :: args) = | 
| 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 289 | let val (Ts, args') = synT markup (symbs, args); | 
| 0 | 290 | in (astT (t, p) @ Ts, args') end | 
| 23630 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 291 | | synT markup (TypArg p :: symbs, t :: args) = | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 292 | let | 
| 23630 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 293 | val (Ts, args') = synT markup (symbs, args); | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 294 | in | 
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 295 | if type_mode then (astT (t, p) @ Ts, args') | 
| 21772 | 296 | else (pretty I ctxt tabs trf tokentrf true curried t p @ Ts, args') | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 297 | end | 
| 23630 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 298 | | synT markup (String s :: symbs, args) = | 
| 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 299 | let val (Ts, args') = synT markup (symbs, args); | 
| 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 300 | in (markup (Pretty.str s) :: Ts, args') end | 
| 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 301 | | synT markup (Space s :: symbs, args) = | 
| 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 302 | let val (Ts, args') = synT markup (symbs, args); | 
| 8457 | 303 | in (Pretty.str s :: Ts, args') end | 
| 23630 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 304 | | synT markup (Block (i, bsymbs) :: symbs, args) = | 
| 0 | 305 | let | 
| 23630 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 306 | val (bTs, args') = synT markup (bsymbs, args); | 
| 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 307 | val (Ts, args'') = synT markup (symbs, args'); | 
| 14837 
827c68f8267c
added pp_show_brackets; support unbreakable blocks;
 wenzelm parents: 
14783diff
changeset | 308 | val T = | 
| 
827c68f8267c
added pp_show_brackets; support unbreakable blocks;
 wenzelm parents: 
14783diff
changeset | 309 | if i < 0 then Pretty.unbreakable (Pretty.block bTs) | 
| 
827c68f8267c
added pp_show_brackets; support unbreakable blocks;
 wenzelm parents: 
14783diff
changeset | 310 | else Pretty.blk (i, bTs); | 
| 
827c68f8267c
added pp_show_brackets; support unbreakable blocks;
 wenzelm parents: 
14783diff
changeset | 311 | in (T :: Ts, args'') end | 
| 23630 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 312 | | synT markup (Break i :: symbs, args) = | 
| 14837 
827c68f8267c
added pp_show_brackets; support unbreakable blocks;
 wenzelm parents: 
14783diff
changeset | 313 | let | 
| 23630 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 314 | val (Ts, args') = synT markup (symbs, args); | 
| 14837 
827c68f8267c
added pp_show_brackets; support unbreakable blocks;
 wenzelm parents: 
14783diff
changeset | 315 | val T = if i < 0 then Pretty.fbrk else Pretty.brk i; | 
| 
827c68f8267c
added pp_show_brackets; support unbreakable blocks;
 wenzelm parents: 
14783diff
changeset | 316 | in (T :: Ts, args') end | 
| 23630 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 317 | | synT _ (_ :: _, []) = sys_error "synT" | 
| 0 | 318 | |
| 23630 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 319 | and parT markup (pr, args, p, p': int) = #1 (synT markup | 
| 554 | 320 | (if p > p' orelse | 
| 5691 | 321 | (! show_brackets andalso p' <> SynExt.max_pri andalso not (is_chain pr)) | 
| 23630 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 322 |             then [Block (1, Space "(" :: pr @ [Space ")"])]
 | 
| 554 | 323 | else pr, args)) | 
| 0 | 324 | |
| 19374 
ae4a225e0c1f
pretty: late externing of consts (support authentic syntax);
 wenzelm parents: 
19046diff
changeset | 325 | and atomT a = | 
| 
ae4a225e0c1f
pretty: late externing of consts (support authentic syntax);
 wenzelm parents: 
19046diff
changeset | 326 | (case try (unprefix Lexicon.constN) a of | 
| 23630 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 327 | SOME c => const_markup c (Pretty.str (extern_const c)) | 
| 19374 
ae4a225e0c1f
pretty: late externing of consts (support authentic syntax);
 wenzelm parents: 
19046diff
changeset | 328 | | NONE => | 
| 
ae4a225e0c1f
pretty: late externing of consts (support authentic syntax);
 wenzelm parents: 
19046diff
changeset | 329 | (case try (unprefix Lexicon.fixedN) a of | 
| 
ae4a225e0c1f
pretty: late externing of consts (support authentic syntax);
 wenzelm parents: 
19046diff
changeset | 330 | SOME x => | 
| 
ae4a225e0c1f
pretty: late externing of consts (support authentic syntax);
 wenzelm parents: 
19046diff
changeset | 331 | (case tokentrf "_free" of | 
| 
ae4a225e0c1f
pretty: late externing of consts (support authentic syntax);
 wenzelm parents: 
19046diff
changeset | 332 | SOME f => Pretty.raw_str (f x) | 
| 
ae4a225e0c1f
pretty: late externing of consts (support authentic syntax);
 wenzelm parents: 
19046diff
changeset | 333 | | NONE => Pretty.str x) | 
| 
ae4a225e0c1f
pretty: late externing of consts (support authentic syntax);
 wenzelm parents: 
19046diff
changeset | 334 | | NONE => Pretty.str a)) | 
| 
ae4a225e0c1f
pretty: late externing of consts (support authentic syntax);
 wenzelm parents: 
19046diff
changeset | 335 | |
| 
ae4a225e0c1f
pretty: late externing of consts (support authentic syntax);
 wenzelm parents: 
19046diff
changeset | 336 | and prefixT (_, a, [], _) = [atomT a] | 
| 16611 
edb368e2878f
proper treatment of advanced trfuns: pass thy argument;
 wenzelm parents: 
15973diff
changeset | 337 | | prefixT (c, _, args, p) = astT (appT (c, args), p) | 
| 0 | 338 | |
| 18 | 339 | and splitT 0 ([x], ys) = (x, ys) | 
| 5691 | 340 | | splitT 0 (rev_xs, ys) = (Ast.Appl (rev rev_xs), ys) | 
| 18 | 341 | | splitT n (rev_xs, y :: ys) = splitT (n - 1) (y :: rev_xs, ys) | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
62diff
changeset | 342 | | splitT _ _ = sys_error "splitT" | 
| 18 | 343 | |
| 0 | 344 | and combT (tup as (c, a, args, p)) = | 
| 345 | let | |
| 346 | val nargs = length args; | |
| 23630 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 347 | val markup = const_markup (unprefix Lexicon.constN a) | 
| 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 348 | handle Fail _ => I; | 
| 18 | 349 | |
| 2701 | 350 | (*find matching table entry, or print as prefix / postfix*) | 
| 6280 | 351 | fun prnt ([], []) = prefixT tup | 
| 18957 | 352 | | prnt ([], tb :: tbs) = prnt (Symtab.lookup_list tb a, tbs) | 
| 6280 | 353 | | prnt ((pr, n, p') :: prnps, tbs) = | 
| 23630 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 354 | if nargs = n then parT markup (pr, args, p, p') | 
| 2200 
2538977e94fa
added print_mode: string list ref (order of printer tables);
 wenzelm parents: 
1509diff
changeset | 355 | else if nargs > n andalso not type_mode then | 
| 
2538977e94fa
added print_mode: string list ref (order of printer tables);
 wenzelm parents: 
1509diff
changeset | 356 | astT (appT (splitT n ([c], args)), p) | 
| 6280 | 357 | else prnt (prnps, tbs); | 
| 2200 
2538977e94fa
added print_mode: string list ref (order of printer tables);
 wenzelm parents: 
1509diff
changeset | 358 | in | 
| 2701 | 359 | (case token_trans a args of (*try token translation function*) | 
| 19374 
ae4a225e0c1f
pretty: late externing of consts (support authentic syntax);
 wenzelm parents: 
19046diff
changeset | 360 | SOME s => [Pretty.raw_str s] | 
| 15531 | 361 | | NONE => (*try print translation functions*) | 
| 23937 
66e1f24d655d
eliminated transform_failure (to avoid critical section for main transactions);
 wenzelm parents: 
23660diff
changeset | 362 | astT (apply_trans ctxt (trf a) args, p) handle Match => prnt ([], tabs)) | 
| 2200 
2538977e94fa
added print_mode: string list ref (order of printer tables);
 wenzelm parents: 
1509diff
changeset | 363 | end | 
| 
2538977e94fa
added print_mode: string list ref (order of printer tables);
 wenzelm parents: 
1509diff
changeset | 364 | |
| 5691 | 365 | and astT (c as Ast.Constant a, p) = combT (c, a, [], p) | 
| 8457 | 366 | | astT (Ast.Variable x, _) = [Pretty.str x] | 
| 6280 | 367 | | astT (Ast.Appl ((c as Ast.Constant a) :: (args as _ :: _)), p) = combT (c, a, args, p) | 
| 5691 | 368 | | astT (Ast.Appl (f :: (args as _ :: _)), p) = astT (appT (f, args), p) | 
| 369 |       | astT (ast as Ast.Appl _, _) = raise Ast.AST ("pretty: malformed ast", [ast]);
 | |
| 6280 | 370 | in astT (ast0, p0) end; | 
| 0 | 371 | |
| 372 | ||
| 373 | (* pretty_term_ast *) | |
| 374 | ||
| 21772 | 375 | fun pretty_term_ast extern_const ctxt curried prtabs trf tokentrf ast = | 
| 24612 | 376 | pretty extern_const ctxt (mode_tabs prtabs (print_mode_value ())) | 
| 23630 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 377 | trf tokentrf false curried ast 0; | 
| 0 | 378 | |
| 379 | ||
| 380 | (* pretty_typ_ast *) | |
| 381 | ||
| 21772 | 382 | fun pretty_typ_ast ctxt _ prtabs trf tokentrf ast = | 
| 24612 | 383 | pretty I ctxt (mode_tabs prtabs (print_mode_value ())) | 
| 23630 
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
 wenzelm parents: 
23615diff
changeset | 384 | trf tokentrf true false ast 0; | 
| 0 | 385 | |
| 386 | end; |