author | paulson |
Thu, 04 Dec 2003 10:29:17 +0100 | |
changeset 14272 | 5efbb548107d |
parent 14176 | 716fec31f9ea |
child 14696 | e862cc138e9c |
permissions | -rw-r--r-- |
18 | 1 |
(* Title: Pure/Syntax/printer.ML |
0 | 2 |
ID: $Id$ |
3 |
Author: Tobias Nipkow and Markus Wenzel, TU Muenchen |
|
12785 | 4 |
License: GPL (GNU GENERAL PUBLIC LICENSE) |
18 | 5 |
|
6 |
Pretty printing of asts, terms, types and print (ast) translation. |
|
0 | 7 |
*) |
8 |
||
9 |
signature PRINTER0 = |
|
2384 | 10 |
sig |
504
a4f09493d929
added flag show_brackets for printinmg fully bracketed terms.
nipkow
parents:
381
diff
changeset
|
11 |
val show_brackets: bool ref |
a4f09493d929
added flag show_brackets for printinmg fully bracketed terms.
nipkow
parents:
381
diff
changeset
|
12 |
val show_sorts: bool ref |
0 | 13 |
val show_types: bool ref |
617 | 14 |
val show_no_free_types: bool ref |
14176
716fec31f9ea
Added show_all_types flag, such that all type information in the term
skalberg
parents:
12785
diff
changeset
|
15 |
val show_all_types: bool ref |
2384 | 16 |
end; |
0 | 17 |
|
18 |
signature PRINTER = |
|
2384 | 19 |
sig |
0 | 20 |
include PRINTER0 |
5691 | 21 |
val term_to_ast: (string -> (bool -> typ -> term list -> term) list) -> term -> Ast.ast |
22 |
val typ_to_ast: (string -> (bool -> typ -> term list -> term) list) -> typ -> Ast.ast |
|
23 |
val sort_to_ast: (string -> (bool -> typ -> term list -> term) list) -> sort -> Ast.ast |
|
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
24 |
type prtabs |
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
25 |
val empty_prtabs: prtabs |
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
26 |
val extend_prtabs: prtabs -> string -> SynExt.xprod list -> prtabs |
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
27 |
val merge_prtabs: prtabs -> prtabs -> prtabs |
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
28 |
val pretty_term_ast: bool -> prtabs |
5691 | 29 |
-> (string -> (Ast.ast list -> Ast.ast) list) |
6322 | 30 |
-> (string -> (string -> string * real) option) -> Ast.ast -> Pretty.T |
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
31 |
val pretty_typ_ast: bool -> prtabs |
5691 | 32 |
-> (string -> (Ast.ast list -> Ast.ast) list) |
6322 | 33 |
-> (string -> (string -> string * real) option) -> Ast.ast -> Pretty.T |
2384 | 34 |
end; |
0 | 35 |
|
2365 | 36 |
structure Printer: PRINTER = |
0 | 37 |
struct |
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
38 |
|
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
39 |
|
0 | 40 |
(** options for printing **) |
41 |
||
42 |
val show_types = ref false; |
|
43 |
val show_sorts = ref false; |
|
504
a4f09493d929
added flag show_brackets for printinmg fully bracketed terms.
nipkow
parents:
381
diff
changeset
|
44 |
val show_brackets = ref false; |
617 | 45 |
val show_no_free_types = ref false; |
14176
716fec31f9ea
Added show_all_types flag, such that all type information in the term
skalberg
parents:
12785
diff
changeset
|
46 |
val show_all_types = ref false; |
0 | 47 |
|
48 |
||
2701 | 49 |
(** misc utils **) |
50 |
||
51 |
(* apply print (ast) translation function *) |
|
0 | 52 |
|
12292 | 53 |
fun apply_first [] x = raise Match |
54 |
| apply_first (f :: fs) x = f x handle Match => apply_first fs x; |
|
5691 | 55 |
|
56 |
fun apply_trans name a fs args = |
|
12292 | 57 |
(apply_first fs args handle |
0 | 58 |
Match => raise Match |
9910
f025cf787554
internalize error "insufficient syntax for prefix application";
wenzelm
parents:
8457
diff
changeset
|
59 |
| exn => (priority ("Error in " ^ name ^ " for " ^ quote a); raise exn)); |
0 | 60 |
|
5691 | 61 |
fun apply_typed x y fs = map (fn f => f x y) fs; |
4147 | 62 |
|
0 | 63 |
|
2701 | 64 |
(* simple_ast_of *) |
65 |
||
5691 | 66 |
fun simple_ast_of (Const (c, _)) = Ast.Constant c |
67 |
| simple_ast_of (Free (x, _)) = Ast.Variable x |
|
68 |
| simple_ast_of (Var (xi, _)) = Ast.Variable (Lexicon.string_of_vname xi) |
|
2701 | 69 |
| simple_ast_of (t as _ $ _) = |
70 |
let val (f, args) = strip_comb t in |
|
5691 | 71 |
Ast.mk_appl (simple_ast_of f) (map simple_ast_of args) |
2701 | 72 |
end |
5691 | 73 |
| simple_ast_of (Bound i) = Ast.Variable ("B." ^ string_of_int i) |
2701 | 74 |
| simple_ast_of (Abs _) = sys_error "simple_ast_of: Abs"; |
75 |
||
76 |
||
77 |
||
3776 | 78 |
(** sort_to_ast, typ_to_ast **) |
2701 | 79 |
|
80 |
fun ast_of_termT trf tm = |
|
0 | 81 |
let |
2701 | 82 |
fun ast_of (t as Const ("_class", _) $ Free _) = simple_ast_of t |
83 |
| ast_of (t as Const ("_tfree", _) $ Free _) = simple_ast_of t |
|
84 |
| ast_of (t as Const ("_tvar", _) $ Var _) = simple_ast_of t |
|
85 |
| ast_of (Const (a, _)) = trans a [] |
|
86 |
| ast_of (t as _ $ _) = |
|
87 |
(case strip_comb t of |
|
88 |
(Const (a, _), args) => trans a args |
|
5691 | 89 |
| (f, args) => Ast.Appl (map ast_of (f :: args))) |
90 |
| ast_of t = simple_ast_of t |
|
2701 | 91 |
and trans a args = |
5691 | 92 |
ast_of (apply_trans "print translation" a (apply_typed false dummyT (trf a)) args) |
93 |
handle Match => Ast.mk_appl (Ast.Constant a) (map ast_of args); |
|
94 |
in ast_of tm end; |
|
617 | 95 |
|
5691 | 96 |
fun sort_to_ast trf S = ast_of_termT trf (TypeExt.term_of_sort S); |
97 |
fun typ_to_ast trf T = ast_of_termT trf (TypeExt.term_of_typ (! show_sorts) T); |
|
2701 | 98 |
|
99 |
||
100 |
||
101 |
(** term_to_ast **) |
|
102 |
||
11795 | 103 |
fun mark_freevars ((t as Const (c, _)) $ u) = |
104 |
if c mem_string TokenTrans.standard_token_markers then (t $ u) |
|
105 |
else t $ mark_freevars u |
|
106 |
| mark_freevars (t $ u) = mark_freevars t $ mark_freevars u |
|
2701 | 107 |
| mark_freevars (Abs (x, T, t)) = Abs (x, T, mark_freevars t) |
5691 | 108 |
| mark_freevars (t as Free _) = Lexicon.const "_free" $ t |
6767 | 109 |
| mark_freevars (t as Var (xi, T)) = |
110 |
if xi = SynExt.dddot_indexname then Const ("_DDDOT", T) |
|
111 |
else Lexicon.const "_var" $ t |
|
2701 | 112 |
| mark_freevars a = a; |
113 |
||
14176
716fec31f9ea
Added show_all_types flag, such that all type information in the term
skalberg
parents:
12785
diff
changeset
|
114 |
fun ast_of_term trf show_const_types no_freeTs show_types show_sorts tm = |
2701 | 115 |
let |
0 | 116 |
fun prune_typs (t_seen as (Const _, _)) = t_seen |
117 |
| prune_typs (t as Free (x, ty), seen) = |
|
118 |
if ty = dummyT then (t, seen) |
|
5691 | 119 |
else if no_freeTs orelse t mem seen then (Lexicon.free x, seen) |
0 | 120 |
else (t, t :: seen) |
121 |
| prune_typs (t as Var (xi, ty), seen) = |
|
122 |
if ty = dummyT then (t, seen) |
|
5691 | 123 |
else if no_freeTs orelse t mem seen then (Lexicon.var xi, seen) |
0 | 124 |
else (t, t :: seen) |
125 |
| prune_typs (t_seen as (Bound _, _)) = t_seen |
|
126 |
| prune_typs (Abs (x, ty, t), seen) = |
|
127 |
let val (t', seen') = prune_typs (t, seen); |
|
128 |
in (Abs (x, ty, t'), seen') end |
|
129 |
| prune_typs (t1 $ t2, seen) = |
|
130 |
let |
|
131 |
val (t1', seen') = prune_typs (t1, seen); |
|
132 |
val (t2', seen'') = prune_typs (t2, seen'); |
|
6767 | 133 |
in (t1' $ t2', seen'') end; |
0 | 134 |
|
2701 | 135 |
fun ast_of tm = |
136 |
(case strip_comb tm of |
|
5691 | 137 |
(t as Abs _, ts) => Ast.mk_appl (ast_of (SynTrans.abs_tr' t)) (map ast_of ts) |
2701 | 138 |
| ((c as Const ("_free", _)), Free (x, T) :: ts) => |
5691 | 139 |
Ast.mk_appl (constrain (c $ Lexicon.free x) T) (map ast_of ts) |
2701 | 140 |
| ((c as Const ("_bound", _)), Free (x, T) :: ts) => |
5691 | 141 |
Ast.mk_appl (constrain (c $ Lexicon.free x) T) (map ast_of ts) |
2701 | 142 |
| ((c as Const ("_var", _)), Var (xi, T) :: ts) => |
5691 | 143 |
Ast.mk_appl (constrain (c $ Lexicon.var xi) T) (map ast_of ts) |
14176
716fec31f9ea
Added show_all_types flag, such that all type information in the term
skalberg
parents:
12785
diff
changeset
|
144 |
| (c' as Const (c, T), ts) => |
716fec31f9ea
Added show_all_types flag, such that all type information in the term
skalberg
parents:
12785
diff
changeset
|
145 |
if show_const_types |
716fec31f9ea
Added show_all_types flag, such that all type information in the term
skalberg
parents:
12785
diff
changeset
|
146 |
then Ast.mk_appl (constrain c' T) (map ast_of ts) |
716fec31f9ea
Added show_all_types flag, such that all type information in the term
skalberg
parents:
12785
diff
changeset
|
147 |
else trans c T ts |
5691 | 148 |
| (t, ts) => Ast.mk_appl (simple_ast_of t) (map ast_of ts)) |
0 | 149 |
|
2384 | 150 |
and trans a T args = |
5691 | 151 |
ast_of (apply_trans "print translation" a (apply_typed show_sorts T (trf a)) args) |
152 |
handle Match => Ast.mk_appl (Ast.Constant a) (map ast_of args) |
|
0 | 153 |
|
2701 | 154 |
and constrain t T = |
155 |
if show_types andalso T <> dummyT then |
|
5691 | 156 |
Ast.Appl [Ast.Constant SynExt.constrainC, simple_ast_of t, |
157 |
ast_of_termT trf (TypeExt.term_of_typ show_sorts T)] |
|
2701 | 158 |
else simple_ast_of t |
0 | 159 |
in |
2701 | 160 |
tm |
5691 | 161 |
|> SynTrans.prop_tr' |
2701 | 162 |
|> (if show_types then #1 o prune_typs o rpair [] else I) |
163 |
|> mark_freevars |
|
164 |
|> ast_of |
|
0 | 165 |
end; |
166 |
||
167 |
fun term_to_ast trf tm = |
|
14176
716fec31f9ea
Added show_all_types flag, such that all type information in the term
skalberg
parents:
12785
diff
changeset
|
168 |
ast_of_term trf (! show_all_types) (! show_no_free_types) (! show_types orelse ! show_sorts orelse ! show_all_types) |
2701 | 169 |
(! show_sorts) tm; |
0 | 170 |
|
171 |
||
172 |
||
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
173 |
(** type prtabs **) |
0 | 174 |
|
175 |
datatype symb = |
|
176 |
Arg of int | |
|
177 |
TypArg of int | |
|
178 |
String of string | |
|
179 |
Break of int | |
|
180 |
Block of int * symb list; |
|
181 |
||
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
182 |
type prtabs = (string * ((symb list * int * int) list) Symtab.table) list; |
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
183 |
|
3816 | 184 |
|
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
185 |
(*find tab for mode*) |
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
186 |
fun get_tab prtabs mode = |
4487 | 187 |
if_none (assoc (prtabs, mode)) Symtab.empty; |
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
188 |
|
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
189 |
(*collect tabs for mode hierarchy (default "")*) |
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
190 |
fun tabs_of prtabs modes = |
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
191 |
mapfilter (fn mode => assoc (prtabs, mode)) (modes @ [""]); |
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
192 |
|
18 | 193 |
|
3816 | 194 |
(* xprod_to_fmt *) |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
195 |
|
5691 | 196 |
fun xprod_to_fmt (SynExt.XProd (_, _, "", _)) = None |
197 |
| xprod_to_fmt (SynExt.XProd (_, xsymbs, const, pri)) = |
|
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
198 |
let |
2384 | 199 |
fun cons_str s (String s' :: syms) = String (s ^ s') :: syms |
200 |
| cons_str s syms = String s :: syms; |
|
201 |
||
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
202 |
fun arg (s, p) = |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
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:
62
diff
changeset
|
205 |
|
5691 | 206 |
fun xsyms_to_syms (SynExt.Delim s :: xsyms) = |
6273 | 207 |
apfst (cons_str 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:
62
diff
changeset
|
209 |
apfst (cons (arg s_p)) (xsyms_to_syms xsyms) |
5691 | 210 |
| xsyms_to_syms (SynExt.Space s :: xsyms) = |
6273 | 211 |
apfst (cons_str 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:
62
diff
changeset
|
213 |
let |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
214 |
val (bsyms, xsyms') = xsyms_to_syms xsyms; |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
215 |
val (syms, xsyms'') = xsyms_to_syms xsyms'; |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
216 |
in |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
217 |
(Block (i, bsyms) :: syms, xsyms'') |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
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:
62
diff
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:
62
diff
changeset
|
222 |
| xsyms_to_syms [] = ([], []); |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
223 |
|
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
224 |
fun nargs (Arg _ :: syms) = nargs syms + 1 |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
225 |
| nargs (TypArg _ :: syms) = nargs syms + 1 |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
226 |
| nargs (String _ :: syms) = nargs syms |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
227 |
| nargs (Break _ :: syms) = nargs syms |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
228 |
| nargs (Block (_, bsyms) :: syms) = nargs syms + nargs bsyms |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
229 |
| nargs [] = 0; |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
230 |
in |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
231 |
(case xsyms_to_syms xsymbs of |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
232 |
(symbs, []) => Some (const, (symbs, nargs symbs, pri)) |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
233 |
| _ => sys_error "xprod_to_fmt: unbalanced blocks") |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
234 |
end; |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
235 |
|
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
236 |
|
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
237 |
(* empty, extend, merge prtabs *) |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
238 |
|
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
239 |
val empty_prtabs = []; |
18 | 240 |
|
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
241 |
fun extend_prtabs prtabs mode xprods = |
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
242 |
let |
12292 | 243 |
val fmts = mapfilter xprod_to_fmt xprods; |
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
244 |
val tab = get_tab prtabs mode; |
12292 | 245 |
val new_tab = foldr Symtab.update_multi (rev fmts, tab); |
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
246 |
in overwrite (prtabs, (mode, new_tab)) end; |
0 | 247 |
|
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
248 |
fun merge_prtabs prtabs1 prtabs2 = |
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
249 |
let |
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
250 |
val modes = distinct (map fst (prtabs1 @ prtabs2)); |
12292 | 251 |
fun merge m = (m, Symtab.merge_multi' (op =) (get_tab prtabs1 m, get_tab prtabs2 m)); |
252 |
in map merge modes end; |
|
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
253 |
|
0 | 254 |
|
255 |
||
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
256 |
(** pretty term or typ asts **) |
0 | 257 |
|
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
258 |
fun is_chain [Block (_, pr)] = is_chain pr |
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
259 |
| is_chain [Arg _] = true |
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
260 |
| is_chain _ = false; |
506
e0ca460d6e51
improved show_brackets again - Trueprop does not create () any more.
nipkow
parents:
505
diff
changeset
|
261 |
|
2701 | 262 |
fun pretty tabs trf tokentrf type_mode curried ast0 p0 = |
0 | 263 |
let |
18 | 264 |
val trans = apply_trans "print ast translation"; |
265 |
||
5691 | 266 |
fun token_trans c [Ast.Variable x] = |
2701 | 267 |
(case tokentrf c of |
268 |
None => None |
|
269 |
| Some f => Some (f x)) |
|
270 |
| token_trans _ _ = None; |
|
271 |
||
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
272 |
(*default applications: prefix / postfix*) |
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
273 |
val appT = |
5691 | 274 |
if type_mode then TypeExt.tappl_ast_tr' |
275 |
else if curried then SynTrans.applC_ast_tr' |
|
276 |
else SynTrans.appl_ast_tr'; |
|
18 | 277 |
|
0 | 278 |
fun synT ([], args) = ([], args) |
279 |
| synT (Arg p :: symbs, t :: args) = |
|
18 | 280 |
let val (Ts, args') = synT (symbs, args); |
0 | 281 |
in (astT (t, p) @ Ts, args') end |
282 |
| synT (TypArg p :: symbs, t :: args) = |
|
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
283 |
let |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
284 |
val (Ts, args') = synT (symbs, args); |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
285 |
in |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
286 |
if type_mode then (astT (t, p) @ Ts, args') |
2701 | 287 |
else (pretty tabs trf tokentrf true curried t p @ Ts, args') |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
288 |
end |
0 | 289 |
| synT (String s :: symbs, args) = |
18 | 290 |
let val (Ts, args') = synT (symbs, args); |
8457 | 291 |
in (Pretty.str s :: Ts, args') end |
0 | 292 |
| synT (Block (i, bsymbs) :: symbs, args) = |
293 |
let |
|
294 |
val (bTs, args') = synT (bsymbs, args); |
|
295 |
val (Ts, args'') = synT (symbs, args'); |
|
296 |
in (Pretty.blk (i, bTs) :: Ts, args'') end |
|
297 |
| synT (Break i :: symbs, args) = |
|
18 | 298 |
let val (Ts, args') = synT (symbs, args); |
299 |
in ((if i < 0 then Pretty.fbrk else Pretty.brk i) :: Ts, args') end |
|
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
300 |
| synT (_ :: _, []) = sys_error "synT" |
0 | 301 |
|
554 | 302 |
and parT (pr, args, p, p': int) = #1 (synT |
303 |
(if p > p' orelse |
|
5691 | 304 |
(! show_brackets andalso p' <> SynExt.max_pri andalso not (is_chain pr)) |
554 | 305 |
then [Block (1, String "(" :: pr @ [String ")"])] |
306 |
else pr, args)) |
|
0 | 307 |
|
8457 | 308 |
and prefixT (_, a, [], _) = [Pretty.str a] |
2384 | 309 |
| prefixT (c, _, args, p) = |
5691 | 310 |
if c = Ast.Constant "_appl" orelse c = Ast.Constant "_applC" then |
9910
f025cf787554
internalize error "insufficient syntax for prefix application";
wenzelm
parents:
8457
diff
changeset
|
311 |
[Pretty.str "*** INSUFFICIENT SYNTAX FOR PREFIX APPLICATION ***"] |
2384 | 312 |
else astT (appT (c, args), p) |
0 | 313 |
|
18 | 314 |
and splitT 0 ([x], ys) = (x, ys) |
5691 | 315 |
| splitT 0 (rev_xs, ys) = (Ast.Appl (rev rev_xs), ys) |
18 | 316 |
| 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:
62
diff
changeset
|
317 |
| splitT _ _ = sys_error "splitT" |
18 | 318 |
|
0 | 319 |
and combT (tup as (c, a, args, p)) = |
320 |
let |
|
321 |
val nargs = length args; |
|
18 | 322 |
|
2701 | 323 |
(*find matching table entry, or print as prefix / postfix*) |
6280 | 324 |
fun prnt ([], []) = prefixT tup |
325 |
| prnt ([], tb :: tbs) = prnt (Symtab.lookup_multi (tb, a), tbs) |
|
326 |
| prnt ((pr, n, p') :: prnps, tbs) = |
|
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
327 |
if nargs = n then parT (pr, args, p, p') |
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
328 |
else if nargs > n andalso not type_mode then |
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
329 |
astT (appT (splitT n ([c], args)), p) |
6280 | 330 |
else prnt (prnps, tbs); |
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
331 |
in |
2701 | 332 |
(case token_trans a args of (*try token translation function*) |
8457 | 333 |
Some s_len => [Pretty.raw_str s_len] |
5691 | 334 |
| None => (*try print translation functions*) |
6280 | 335 |
astT (trans a (trf a) args, p) handle Match => prnt ([], tabs)) |
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
336 |
end |
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
337 |
|
5691 | 338 |
and astT (c as Ast.Constant a, p) = combT (c, a, [], p) |
8457 | 339 |
| astT (Ast.Variable x, _) = [Pretty.str x] |
6280 | 340 |
| astT (Ast.Appl ((c as Ast.Constant a) :: (args as _ :: _)), p) = combT (c, a, args, p) |
5691 | 341 |
| astT (Ast.Appl (f :: (args as _ :: _)), p) = astT (appT (f, args), p) |
342 |
| astT (ast as Ast.Appl _, _) = raise Ast.AST ("pretty: malformed ast", [ast]); |
|
6280 | 343 |
in astT (ast0, p0) end; |
0 | 344 |
|
345 |
||
346 |
(* pretty_term_ast *) |
|
347 |
||
2701 | 348 |
fun pretty_term_ast curried prtabs trf tokentrf ast = |
349 |
Pretty.blk (0, pretty (tabs_of prtabs (! print_mode)) |
|
350 |
trf tokentrf false curried ast 0); |
|
0 | 351 |
|
352 |
||
353 |
(* pretty_typ_ast *) |
|
354 |
||
2701 | 355 |
fun pretty_typ_ast _ prtabs trf tokentrf ast = |
356 |
Pretty.blk (0, pretty (tabs_of prtabs (! print_mode)) |
|
357 |
trf tokentrf true false ast 0); |
|
0 | 358 |
|
359 |
end; |