author | wenzelm |
Fri, 03 Sep 2010 16:36:33 +0200 | |
changeset 39117 | 399977145846 |
parent 39116 | f14735a88886 |
child 39118 | 12f3788be67b |
permissions | -rw-r--r-- |
18 | 1 |
(* Title: Pure/Syntax/printer.ML |
0 | 2 |
Author: Tobias Nipkow and Markus Wenzel, TU Muenchen |
18 | 3 |
|
4 |
Pretty printing of asts, terms, types and print (ast) translation. |
|
0 | 5 |
*) |
6 |
||
7 |
signature PRINTER0 = |
|
2384 | 8 |
sig |
32738 | 9 |
val show_brackets: bool Unsynchronized.ref |
10 |
val show_sorts: bool Unsynchronized.ref |
|
11 |
val show_types: bool Unsynchronized.ref |
|
39115
a00da1674c1c
turned show_no_free_types into proper configuration option show_free_types, with flipped polarity;
wenzelm
parents:
38980
diff
changeset
|
12 |
val show_free_types: bool Config.T |
32738 | 13 |
val show_all_types: bool Unsynchronized.ref |
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
14 |
val show_structs: bool Unsynchronized.ref |
38980
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
37852
diff
changeset
|
15 |
val show_question_marks_default: bool Unsynchronized.ref |
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
37852
diff
changeset
|
16 |
val show_question_marks_value: Config.value Config.T |
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
37852
diff
changeset
|
17 |
val show_question_marks: bool Config.T |
14837
827c68f8267c
added pp_show_brackets; support unbreakable blocks;
wenzelm
parents:
14783
diff
changeset
|
18 |
val pp_show_brackets: Pretty.pp -> Pretty.pp |
2384 | 19 |
end; |
0 | 20 |
|
21 |
signature PRINTER = |
|
2384 | 22 |
sig |
0 | 23 |
include PRINTER0 |
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
24 |
val sort_to_ast: Proof.context -> |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
25 |
(string -> (Proof.context -> bool -> typ -> term list -> term) list) -> sort -> Ast.ast |
21772 | 26 |
val typ_to_ast: Proof.context -> |
27 |
(string -> (Proof.context -> bool -> typ -> term list -> term) list) -> typ -> Ast.ast |
|
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
28 |
val term_to_ast: {structs: string list, fixes: string list} -> string list -> Proof.context -> |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
29 |
(string -> (Proof.context -> bool -> typ -> term list -> term) list) -> term -> Ast.ast |
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
30 |
type prtabs |
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
31 |
val empty_prtabs: prtabs |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35435
diff
changeset
|
32 |
val update_prtabs: string -> Syn_Ext.xprod list -> prtabs -> prtabs |
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35435
diff
changeset
|
33 |
val remove_prtabs: string -> Syn_Ext.xprod list -> prtabs -> prtabs |
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
34 |
val merge_prtabs: prtabs -> prtabs -> prtabs |
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
35 |
val pretty_term_ast: {extern_class: string -> xstring, extern_type: string -> xstring, |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
36 |
extern_const: string -> xstring} -> Proof.context -> bool -> prtabs -> |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
37 |
(string -> (Proof.context -> Ast.ast list -> Ast.ast) list) -> |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
38 |
(string -> (Proof.context -> string -> Pretty.T) option) -> Ast.ast -> Pretty.T list |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
39 |
val pretty_typ_ast: {extern_class: string -> xstring, extern_type: string -> xstring} -> |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
40 |
Proof.context -> bool -> prtabs -> |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
41 |
(string -> (Proof.context -> Ast.ast list -> Ast.ast) list) -> |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
42 |
(string -> (Proof.context -> string -> Pretty.T) option) -> Ast.ast -> Pretty.T list |
2384 | 43 |
end; |
0 | 44 |
|
2365 | 45 |
structure Printer: PRINTER = |
0 | 46 |
struct |
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
47 |
|
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
48 |
|
0 | 49 |
(** options for printing **) |
50 |
||
32738 | 51 |
val show_types = Unsynchronized.ref false; |
52 |
val show_sorts = Unsynchronized.ref false; |
|
53 |
val show_brackets = Unsynchronized.ref false; |
|
54 |
val show_all_types = Unsynchronized.ref false; |
|
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
55 |
val show_structs = Unsynchronized.ref false; |
0 | 56 |
|
39117
399977145846
treat show_free_types as plain ML option, without the extras of global default and registration in the attribute name space -- NB: 'print_configs' only shows the latter;
wenzelm
parents:
39116
diff
changeset
|
57 |
val show_free_types = Config.bool (Config.declare "show_free_types" (fn _ => Config.Bool true)); |
39115
a00da1674c1c
turned show_no_free_types into proper configuration option show_free_types, with flipped polarity;
wenzelm
parents:
38980
diff
changeset
|
58 |
|
38980
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
37852
diff
changeset
|
59 |
val show_question_marks_default = Unsynchronized.ref true; |
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
37852
diff
changeset
|
60 |
val show_question_marks_value = |
39116
f14735a88886
more explicit Config.declare vs. Config.declare_global;
wenzelm
parents:
39115
diff
changeset
|
61 |
Config.declare "show_question_marks" (fn _ => Config.Bool (! show_question_marks_default)); |
38980
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
37852
diff
changeset
|
62 |
val show_question_marks = Config.bool show_question_marks_value; |
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
37852
diff
changeset
|
63 |
|
32966
5b21661fe618
indicate CRITICAL nature of various setmp combinators;
wenzelm
parents:
32738
diff
changeset
|
64 |
fun pp_show_brackets pp = Pretty.pp (setmp_CRITICAL show_brackets true (Pretty.term pp), |
14975 | 65 |
Pretty.typ pp, Pretty.sort pp, Pretty.classrel pp, Pretty.arity pp); |
14837
827c68f8267c
added pp_show_brackets; support unbreakable blocks;
wenzelm
parents:
14783
diff
changeset
|
66 |
|
827c68f8267c
added pp_show_brackets; support unbreakable blocks;
wenzelm
parents:
14783
diff
changeset
|
67 |
|
0 | 68 |
|
2701 | 69 |
(** misc utils **) |
70 |
||
71 |
(* apply print (ast) translation function *) |
|
0 | 72 |
|
23937
66e1f24d655d
eliminated transform_failure (to avoid critical section for main transactions);
wenzelm
parents:
23660
diff
changeset
|
73 |
fun apply_trans ctxt fns args = |
16611
edb368e2878f
proper treatment of advanced trfuns: pass thy argument;
wenzelm
parents:
15973
diff
changeset
|
74 |
let |
edb368e2878f
proper treatment of advanced trfuns: pass thy argument;
wenzelm
parents:
15973
diff
changeset
|
75 |
fun app_first [] = raise Match |
21772 | 76 |
| 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:
23660
diff
changeset
|
77 |
in app_first fns end; |
5691 | 78 |
|
21772 | 79 |
fun apply_typed x y fs = map (fn f => fn ctxt => f ctxt x y) fs; |
4147 | 80 |
|
0 | 81 |
|
2701 | 82 |
(* simple_ast_of *) |
83 |
||
38980
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
37852
diff
changeset
|
84 |
fun simple_ast_of ctxt = |
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
37852
diff
changeset
|
85 |
let |
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
37852
diff
changeset
|
86 |
val tune_var = if Config.get ctxt show_question_marks then I else unprefix "?"; |
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
37852
diff
changeset
|
87 |
fun ast_of (Const (c, _)) = Ast.Constant c |
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
37852
diff
changeset
|
88 |
| ast_of (Free (x, _)) = Ast.Variable x |
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
37852
diff
changeset
|
89 |
| ast_of (Var (xi, _)) = Ast.Variable (tune_var (Term.string_of_vname xi)) |
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
37852
diff
changeset
|
90 |
| ast_of (t as _ $ _) = |
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
37852
diff
changeset
|
91 |
let val (f, args) = strip_comb t |
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
37852
diff
changeset
|
92 |
in Ast.mk_appl (ast_of f) (map ast_of args) end |
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
37852
diff
changeset
|
93 |
| ast_of (Bound i) = Ast.Variable ("B." ^ string_of_int i) |
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
37852
diff
changeset
|
94 |
| ast_of (Abs _) = raise Fail "simple_ast_of: Abs"; |
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
37852
diff
changeset
|
95 |
in ast_of end; |
2701 | 96 |
|
97 |
||
98 |
||
3776 | 99 |
(** sort_to_ast, typ_to_ast **) |
2701 | 100 |
|
21772 | 101 |
fun ast_of_termT ctxt trf tm = |
0 | 102 |
let |
38980
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
37852
diff
changeset
|
103 |
fun ast_of (t as Const ("_tfree", _) $ Free _) = simple_ast_of ctxt t |
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
37852
diff
changeset
|
104 |
| ast_of (t as Const ("_tvar", _) $ Var _) = simple_ast_of ctxt t |
2701 | 105 |
| ast_of (Const (a, _)) = trans a [] |
106 |
| ast_of (t as _ $ _) = |
|
107 |
(case strip_comb t of |
|
108 |
(Const (a, _), args) => trans a args |
|
5691 | 109 |
| (f, args) => Ast.Appl (map ast_of (f :: args))) |
38980
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
37852
diff
changeset
|
110 |
| ast_of t = simple_ast_of ctxt t |
2701 | 111 |
and trans a args = |
23937
66e1f24d655d
eliminated transform_failure (to avoid critical section for main transactions);
wenzelm
parents:
23660
diff
changeset
|
112 |
ast_of (apply_trans ctxt (apply_typed false dummyT (trf a)) args) |
5691 | 113 |
handle Match => Ast.mk_appl (Ast.Constant a) (map ast_of args); |
114 |
in ast_of tm end; |
|
617 | 115 |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35435
diff
changeset
|
116 |
fun sort_to_ast ctxt trf S = ast_of_termT ctxt trf (Type_Ext.term_of_sort S); |
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35435
diff
changeset
|
117 |
fun typ_to_ast ctxt trf T = ast_of_termT ctxt trf (Type_Ext.term_of_typ (! show_sorts) T); |
2701 | 118 |
|
119 |
||
120 |
||
121 |
(** term_to_ast **) |
|
122 |
||
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
123 |
fun ast_of_term idents consts ctxt trf |
39115
a00da1674c1c
turned show_no_free_types into proper configuration option show_free_types, with flipped polarity;
wenzelm
parents:
38980
diff
changeset
|
124 |
show_all_types show_types show_sorts show_structs tm = |
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
125 |
let |
39115
a00da1674c1c
turned show_no_free_types into proper configuration option show_free_types, with flipped polarity;
wenzelm
parents:
38980
diff
changeset
|
126 |
val show_free_types = Config.get ctxt show_free_types; |
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
127 |
val {structs, fixes} = idents; |
2701 | 128 |
|
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
129 |
fun mark_atoms ((t as Const (c, T)) $ u) = |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35435
diff
changeset
|
130 |
if member (op =) Syn_Ext.standard_token_markers c |
35435
e6c03f397eb8
mark_atoms: more precise treatment of SynExt.standard_token_markers vs. syntax consts;
wenzelm
parents:
35429
diff
changeset
|
131 |
then t $ u else mark_atoms t $ mark_atoms u |
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
132 |
| mark_atoms (t $ u) = mark_atoms t $ mark_atoms u |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
133 |
| mark_atoms (Abs (x, T, t)) = Abs (x, T, mark_atoms t) |
35435
e6c03f397eb8
mark_atoms: more precise treatment of SynExt.standard_token_markers vs. syntax consts;
wenzelm
parents:
35429
diff
changeset
|
134 |
| mark_atoms (t as Const (c, T)) = |
e6c03f397eb8
mark_atoms: more precise treatment of SynExt.standard_token_markers vs. syntax consts;
wenzelm
parents:
35429
diff
changeset
|
135 |
if member (op =) consts c then t |
e6c03f397eb8
mark_atoms: more precise treatment of SynExt.standard_token_markers vs. syntax consts;
wenzelm
parents:
35429
diff
changeset
|
136 |
else Const (Lexicon.mark_const c, T) |
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
137 |
| mark_atoms (t as Free (x, T)) = |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
138 |
let val i = find_index (fn s => s = x) structs + 1 in |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
139 |
if i = 0 andalso member (op =) fixes x then |
35435
e6c03f397eb8
mark_atoms: more precise treatment of SynExt.standard_token_markers vs. syntax consts;
wenzelm
parents:
35429
diff
changeset
|
140 |
Const (Lexicon.mark_fixed x, T) |
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
141 |
else if i = 1 andalso not show_structs then |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
142 |
Lexicon.const "_struct" $ Lexicon.const "_indexdefault" |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
143 |
else Lexicon.const "_free" $ t |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
144 |
end |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
145 |
| mark_atoms (t as Var (xi, T)) = |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35435
diff
changeset
|
146 |
if xi = Syn_Ext.dddot_indexname then Const ("_DDDOT", T) |
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
147 |
else Lexicon.const "_var" $ t |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
148 |
| mark_atoms a = a; |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
149 |
|
0 | 150 |
fun prune_typs (t_seen as (Const _, _)) = t_seen |
151 |
| prune_typs (t as Free (x, ty), seen) = |
|
152 |
if ty = dummyT then (t, seen) |
|
39115
a00da1674c1c
turned show_no_free_types into proper configuration option show_free_types, with flipped polarity;
wenzelm
parents:
38980
diff
changeset
|
153 |
else if not show_free_types orelse member (op aconv) seen t then (Lexicon.free x, seen) |
0 | 154 |
else (t, t :: seen) |
155 |
| prune_typs (t as Var (xi, ty), seen) = |
|
156 |
if ty = dummyT then (t, seen) |
|
39115
a00da1674c1c
turned show_no_free_types into proper configuration option show_free_types, with flipped polarity;
wenzelm
parents:
38980
diff
changeset
|
157 |
else if not show_free_types orelse member (op aconv) seen t then (Lexicon.var xi, seen) |
0 | 158 |
else (t, t :: seen) |
159 |
| prune_typs (t_seen as (Bound _, _)) = t_seen |
|
160 |
| prune_typs (Abs (x, ty, t), seen) = |
|
161 |
let val (t', seen') = prune_typs (t, seen); |
|
162 |
in (Abs (x, ty, t'), seen') end |
|
163 |
| prune_typs (t1 $ t2, seen) = |
|
164 |
let |
|
165 |
val (t1', seen') = prune_typs (t1, seen); |
|
166 |
val (t2', seen'') = prune_typs (t2, seen'); |
|
6767 | 167 |
in (t1' $ t2', seen'') end; |
0 | 168 |
|
2701 | 169 |
fun ast_of tm = |
170 |
(case strip_comb tm of |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35435
diff
changeset
|
171 |
(t as Abs _, ts) => Ast.mk_appl (ast_of (Syn_Trans.abs_tr' t)) (map ast_of ts) |
2701 | 172 |
| ((c as Const ("_free", _)), Free (x, T) :: ts) => |
5691 | 173 |
Ast.mk_appl (constrain (c $ Lexicon.free x) T) (map ast_of ts) |
21748 | 174 |
| ((c as Const ("_var", _)), Var (xi, T) :: ts) => |
175 |
Ast.mk_appl (constrain (c $ Lexicon.var xi) T) (map ast_of ts) |
|
2701 | 176 |
| ((c as Const ("_bound", _)), Free (x, T) :: ts) => |
5691 | 177 |
Ast.mk_appl (constrain (c $ Lexicon.free x) T) (map ast_of ts) |
21748 | 178 |
| (Const ("_idtdummy", T), ts) => |
179 |
Ast.mk_appl (constrain (Lexicon.const "_idtdummy") T) (map ast_of ts) |
|
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
180 |
| (const as Const (c, T), ts) => |
14696 | 181 |
if show_all_types |
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
182 |
then Ast.mk_appl (constrain const T) (map ast_of ts) |
14696 | 183 |
else trans c T ts |
38980
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
37852
diff
changeset
|
184 |
| (t, ts) => Ast.mk_appl (simple_ast_of ctxt t) (map ast_of ts)) |
0 | 185 |
|
2384 | 186 |
and trans a T args = |
23937
66e1f24d655d
eliminated transform_failure (to avoid critical section for main transactions);
wenzelm
parents:
23660
diff
changeset
|
187 |
ast_of (apply_trans ctxt (apply_typed show_sorts T (trf a)) args) |
5691 | 188 |
handle Match => Ast.mk_appl (Ast.Constant a) (map ast_of args) |
0 | 189 |
|
2701 | 190 |
and constrain t T = |
191 |
if show_types andalso T <> dummyT then |
|
38980
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
37852
diff
changeset
|
192 |
Ast.Appl [Ast.Constant Syn_Ext.constrainC, simple_ast_of ctxt t, |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35435
diff
changeset
|
193 |
ast_of_termT ctxt trf (Type_Ext.term_of_typ show_sorts T)] |
38980
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
37852
diff
changeset
|
194 |
else simple_ast_of ctxt t; |
0 | 195 |
in |
2701 | 196 |
tm |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35435
diff
changeset
|
197 |
|> Syn_Trans.prop_tr' |
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
198 |
|> show_types ? (#1 o prune_typs o rpair []) |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
199 |
|> mark_atoms |
2701 | 200 |
|> ast_of |
0 | 201 |
end; |
202 |
||
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
203 |
fun term_to_ast idents consts ctxt trf tm = |
39115
a00da1674c1c
turned show_no_free_types into proper configuration option show_free_types, with flipped polarity;
wenzelm
parents:
38980
diff
changeset
|
204 |
ast_of_term idents consts ctxt trf (! show_all_types) |
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
205 |
(! show_types orelse ! show_sorts orelse ! show_all_types) (! show_sorts) (! show_structs) tm; |
0 | 206 |
|
207 |
||
208 |
||
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
209 |
(** type prtabs **) |
0 | 210 |
|
211 |
datatype symb = |
|
212 |
Arg of int | |
|
213 |
TypArg of int | |
|
214 |
String of string | |
|
23630
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
wenzelm
parents:
23615
diff
changeset
|
215 |
Space of string | |
0 | 216 |
Break of int | |
217 |
Block of int * symb list; |
|
218 |
||
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
219 |
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
|
220 |
|
18957 | 221 |
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:
19374
diff
changeset
|
222 |
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:
1509
diff
changeset
|
223 |
|
18 | 224 |
|
3816 | 225 |
(* xprod_to_fmt *) |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
226 |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35435
diff
changeset
|
227 |
fun xprod_to_fmt (Syn_Ext.XProd (_, _, "", _)) = NONE |
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35435
diff
changeset
|
228 |
| xprod_to_fmt (Syn_Ext.XProd (_, xsymbs, const, pri)) = |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
229 |
let |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
230 |
fun arg (s, p) = |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
231 |
(if s = "type" then TypArg else Arg) |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35435
diff
changeset
|
232 |
(if Lexicon.is_terminal s then Syn_Ext.max_pri else p); |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
233 |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35435
diff
changeset
|
234 |
fun xsyms_to_syms (Syn_Ext.Delim s :: xsyms) = |
23630
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
wenzelm
parents:
23615
diff
changeset
|
235 |
apfst (cons (String s)) (xsyms_to_syms xsyms) |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35435
diff
changeset
|
236 |
| xsyms_to_syms (Syn_Ext.Argument s_p :: xsyms) = |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
237 |
apfst (cons (arg s_p)) (xsyms_to_syms xsyms) |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35435
diff
changeset
|
238 |
| xsyms_to_syms (Syn_Ext.Space s :: xsyms) = |
23630
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
wenzelm
parents:
23615
diff
changeset
|
239 |
apfst (cons (Space s)) (xsyms_to_syms xsyms) |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35435
diff
changeset
|
240 |
| xsyms_to_syms (Syn_Ext.Bg i :: xsyms) = |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
241 |
let |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
242 |
val (bsyms, xsyms') = xsyms_to_syms xsyms; |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
243 |
val (syms, xsyms'') = xsyms_to_syms xsyms'; |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
244 |
in |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
245 |
(Block (i, bsyms) :: syms, xsyms'') |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
246 |
end |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35435
diff
changeset
|
247 |
| xsyms_to_syms (Syn_Ext.Brk i :: xsyms) = |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
248 |
apfst (cons (Break i)) (xsyms_to_syms xsyms) |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35435
diff
changeset
|
249 |
| xsyms_to_syms (Syn_Ext.En :: xsyms) = ([], xsyms) |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
250 |
| xsyms_to_syms [] = ([], []); |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
251 |
|
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
252 |
fun nargs (Arg _ :: syms) = nargs syms + 1 |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
253 |
| nargs (TypArg _ :: syms) = nargs syms + 1 |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
254 |
| nargs (String _ :: syms) = nargs syms |
23630
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
wenzelm
parents:
23615
diff
changeset
|
255 |
| nargs (Space _ :: syms) = nargs syms |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
256 |
| nargs (Break _ :: syms) = nargs syms |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
257 |
| nargs (Block (_, bsyms) :: syms) = nargs syms + nargs bsyms |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
258 |
| nargs [] = 0; |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
259 |
in |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
260 |
(case xsyms_to_syms xsymbs of |
15531 | 261 |
(symbs, []) => SOME (const, (symbs, nargs symbs, pri)) |
37852
a902f158b4fc
eliminated old-style sys_error/SYS_ERROR in favour of exception Fail -- after careful checking that there is no overlap with existing handling of that;
wenzelm
parents:
37216
diff
changeset
|
262 |
| _ => raise Fail "Unbalanced pretty-printing blocks") |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
263 |
end; |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
264 |
|
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
265 |
|
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
266 |
(* empty, extend, merge prtabs *) |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
267 |
|
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
268 |
val empty_prtabs = []; |
18 | 269 |
|
25393
0856e0141062
replaced extend_prtabs by update_prtabs (absorb duplicates);
wenzelm
parents:
25386
diff
changeset
|
270 |
fun update_prtabs mode xprods prtabs = |
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
271 |
let |
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19374
diff
changeset
|
272 |
val fmts = map_filter xprod_to_fmt xprods; |
25393
0856e0141062
replaced extend_prtabs by update_prtabs (absorb duplicates);
wenzelm
parents:
25386
diff
changeset
|
273 |
val tab' = fold (Symtab.update_list (op =)) fmts (mode_tab prtabs mode); |
25386 | 274 |
in AList.update (op =) (mode, tab') prtabs end; |
15753 | 275 |
|
25386 | 276 |
fun remove_prtabs mode xprods prtabs = |
277 |
let |
|
278 |
val tab = mode_tab prtabs mode; |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35435
diff
changeset
|
279 |
val fmts = map_filter (fn xprod as Syn_Ext.XProd (_, _, c, _) => |
25386 | 280 |
if null (Symtab.lookup_list tab c) then NONE |
281 |
else xprod_to_fmt xprod) xprods; |
|
25393
0856e0141062
replaced extend_prtabs by update_prtabs (absorb duplicates);
wenzelm
parents:
25386
diff
changeset
|
282 |
val tab' = fold (Symtab.remove_list (op =)) fmts tab; |
0856e0141062
replaced extend_prtabs by update_prtabs (absorb duplicates);
wenzelm
parents:
25386
diff
changeset
|
283 |
in AList.update (op =) (mode, tab') prtabs end; |
0 | 284 |
|
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
285 |
fun merge_prtabs prtabs1 prtabs2 = |
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
286 |
let |
19046
bc5c6c9b114e
removed distinct, renamed gen_distinct to distinct;
wenzelm
parents:
18977
diff
changeset
|
287 |
val modes = distinct (op =) (map fst (prtabs1 @ prtabs2)); |
18957 | 288 |
fun merge m = (m, Symtab.merge_list (op =) (mode_tab prtabs1 m, mode_tab prtabs2 m)); |
12292 | 289 |
in map merge modes end; |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
290 |
|
0 | 291 |
|
292 |
||
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
293 |
(** pretty term or typ asts **) |
0 | 294 |
|
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
295 |
fun is_chain [Block (_, pr)] = is_chain pr |
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
296 |
| is_chain [Arg _] = true |
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
297 |
| is_chain _ = false; |
506
e0ca460d6e51
improved show_brackets again - Trueprop does not create () any more.
nipkow
parents:
505
diff
changeset
|
298 |
|
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
299 |
fun pretty extern ctxt tabs trf tokentrf type_mode curried ast0 p0 = |
0 | 300 |
let |
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
301 |
val {extern_class, extern_type, extern_const} = extern; |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
302 |
|
26707
ddf6bab64b96
token translations: context dependent, result Pretty.T;
wenzelm
parents:
25393
diff
changeset
|
303 |
fun token_trans a x = |
ddf6bab64b96
token translations: context dependent, result Pretty.T;
wenzelm
parents:
25393
diff
changeset
|
304 |
(case tokentrf a of |
ddf6bab64b96
token translations: context dependent, result Pretty.T;
wenzelm
parents:
25393
diff
changeset
|
305 |
NONE => |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35435
diff
changeset
|
306 |
if member (op =) Syn_Ext.standard_token_classes a |
26707
ddf6bab64b96
token translations: context dependent, result Pretty.T;
wenzelm
parents:
25393
diff
changeset
|
307 |
then SOME (Pretty.str x) else NONE |
ddf6bab64b96
token translations: context dependent, result Pretty.T;
wenzelm
parents:
25393
diff
changeset
|
308 |
| SOME f => SOME (f ctxt x)); |
2701 | 309 |
|
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
310 |
(*default applications: prefix / postfix*) |
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
311 |
val appT = |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35435
diff
changeset
|
312 |
if type_mode then Type_Ext.tappl_ast_tr' |
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35435
diff
changeset
|
313 |
else if curried then Syn_Trans.applC_ast_tr' |
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35435
diff
changeset
|
314 |
else Syn_Trans.appl_ast_tr'; |
18 | 315 |
|
23630
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
wenzelm
parents:
23615
diff
changeset
|
316 |
fun synT _ ([], args) = ([], args) |
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
wenzelm
parents:
23615
diff
changeset
|
317 |
| synT markup (Arg p :: symbs, t :: args) = |
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
wenzelm
parents:
23615
diff
changeset
|
318 |
let val (Ts, args') = synT markup (symbs, args); |
0 | 319 |
in (astT (t, p) @ Ts, args') end |
23630
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
wenzelm
parents:
23615
diff
changeset
|
320 |
| synT markup (TypArg p :: symbs, t :: args) = |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
321 |
let |
23630
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
wenzelm
parents:
23615
diff
changeset
|
322 |
val (Ts, args') = synT markup (symbs, args); |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
323 |
in |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
62
diff
changeset
|
324 |
if type_mode then (astT (t, p) @ Ts, args') |
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
325 |
else (pretty extern ctxt 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
|
326 |
end |
23630
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
wenzelm
parents:
23615
diff
changeset
|
327 |
| synT markup (String s :: symbs, args) = |
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
wenzelm
parents:
23615
diff
changeset
|
328 |
let val (Ts, args') = synT markup (symbs, args); |
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
wenzelm
parents:
23615
diff
changeset
|
329 |
in (markup (Pretty.str s) :: Ts, args') end |
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
wenzelm
parents:
23615
diff
changeset
|
330 |
| synT markup (Space s :: symbs, args) = |
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
wenzelm
parents:
23615
diff
changeset
|
331 |
let val (Ts, args') = synT markup (symbs, args); |
8457 | 332 |
in (Pretty.str s :: Ts, args') end |
23630
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
wenzelm
parents:
23615
diff
changeset
|
333 |
| synT markup (Block (i, bsymbs) :: symbs, args) = |
0 | 334 |
let |
23630
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
wenzelm
parents:
23615
diff
changeset
|
335 |
val (bTs, args') = synT markup (bsymbs, args); |
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
wenzelm
parents:
23615
diff
changeset
|
336 |
val (Ts, args'') = synT markup (symbs, args'); |
14837
827c68f8267c
added pp_show_brackets; support unbreakable blocks;
wenzelm
parents:
14783
diff
changeset
|
337 |
val T = |
827c68f8267c
added pp_show_brackets; support unbreakable blocks;
wenzelm
parents:
14783
diff
changeset
|
338 |
if i < 0 then Pretty.unbreakable (Pretty.block bTs) |
827c68f8267c
added pp_show_brackets; support unbreakable blocks;
wenzelm
parents:
14783
diff
changeset
|
339 |
else Pretty.blk (i, bTs); |
827c68f8267c
added pp_show_brackets; support unbreakable blocks;
wenzelm
parents:
14783
diff
changeset
|
340 |
in (T :: Ts, args'') end |
23630
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
wenzelm
parents:
23615
diff
changeset
|
341 |
| synT markup (Break i :: symbs, args) = |
14837
827c68f8267c
added pp_show_brackets; support unbreakable blocks;
wenzelm
parents:
14783
diff
changeset
|
342 |
let |
23630
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
wenzelm
parents:
23615
diff
changeset
|
343 |
val (Ts, args') = synT markup (symbs, args); |
14837
827c68f8267c
added pp_show_brackets; support unbreakable blocks;
wenzelm
parents:
14783
diff
changeset
|
344 |
val T = if i < 0 then Pretty.fbrk else Pretty.brk i; |
827c68f8267c
added pp_show_brackets; support unbreakable blocks;
wenzelm
parents:
14783
diff
changeset
|
345 |
in (T :: Ts, args') end |
0 | 346 |
|
23630
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
wenzelm
parents:
23615
diff
changeset
|
347 |
and parT markup (pr, args, p, p': int) = #1 (synT markup |
554 | 348 |
(if p > p' orelse |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35435
diff
changeset
|
349 |
(! show_brackets andalso p' <> Syn_Ext.max_pri andalso not (is_chain pr)) |
23630
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
wenzelm
parents:
23615
diff
changeset
|
350 |
then [Block (1, Space "(" :: pr @ [Space ")"])] |
554 | 351 |
else pr, args)) |
0 | 352 |
|
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
353 |
and atomT a = a |> Lexicon.unmark |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
354 |
{case_class = fn c => Pretty.mark (Markup.tclass c) (Pretty.str (extern_class c)), |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
355 |
case_type = fn c => Pretty.mark (Markup.tycon c) (Pretty.str (extern_type c)), |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
356 |
case_const = fn c => Pretty.mark (Markup.const c) (Pretty.str (extern_const c)), |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
357 |
case_fixed = fn x => the (token_trans "_free" x), |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
358 |
case_default = Pretty.str} |
19374
ae4a225e0c1f
pretty: late externing of consts (support authentic syntax);
wenzelm
parents:
19046
diff
changeset
|
359 |
|
ae4a225e0c1f
pretty: late externing of consts (support authentic syntax);
wenzelm
parents:
19046
diff
changeset
|
360 |
and prefixT (_, a, [], _) = [atomT a] |
16611
edb368e2878f
proper treatment of advanced trfuns: pass thy argument;
wenzelm
parents:
15973
diff
changeset
|
361 |
| prefixT (c, _, args, p) = astT (appT (c, args), p) |
0 | 362 |
|
18 | 363 |
and splitT 0 ([x], ys) = (x, ys) |
5691 | 364 |
| splitT 0 (rev_xs, ys) = (Ast.Appl (rev rev_xs), ys) |
18 | 365 |
| splitT n (rev_xs, y :: ys) = splitT (n - 1) (y :: rev_xs, ys) |
366 |
||
0 | 367 |
and combT (tup as (c, a, args, p)) = |
368 |
let |
|
369 |
val nargs = length args; |
|
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
370 |
val markup = a |> Lexicon.unmark |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
371 |
{case_class = Pretty.mark o Markup.tclass, |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
372 |
case_type = Pretty.mark o Markup.tycon, |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
373 |
case_const = Pretty.mark o Markup.const, |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
374 |
case_fixed = Pretty.mark o Markup.fixed, |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
375 |
case_default = K I}; |
18 | 376 |
|
2701 | 377 |
(*find matching table entry, or print as prefix / postfix*) |
6280 | 378 |
fun prnt ([], []) = prefixT tup |
18957 | 379 |
| prnt ([], tb :: tbs) = prnt (Symtab.lookup_list tb a, tbs) |
6280 | 380 |
| prnt ((pr, n, p') :: prnps, tbs) = |
23630
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
wenzelm
parents:
23615
diff
changeset
|
381 |
if nargs = n then parT markup (pr, args, p, p') |
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
382 |
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
|
383 |
astT (appT (splitT n ([c], args)), p) |
6280 | 384 |
else prnt (prnps, tbs); |
2200
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
385 |
in |
26707
ddf6bab64b96
token translations: context dependent, result Pretty.T;
wenzelm
parents:
25393
diff
changeset
|
386 |
(case tokentrT a args of |
ddf6bab64b96
token translations: context dependent, result Pretty.T;
wenzelm
parents:
25393
diff
changeset
|
387 |
SOME prt => [prt] |
ddf6bab64b96
token translations: context dependent, result Pretty.T;
wenzelm
parents:
25393
diff
changeset
|
388 |
| NONE => 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:
1509
diff
changeset
|
389 |
end |
2538977e94fa
added print_mode: string list ref (order of printer tables);
wenzelm
parents:
1509
diff
changeset
|
390 |
|
26707
ddf6bab64b96
token translations: context dependent, result Pretty.T;
wenzelm
parents:
25393
diff
changeset
|
391 |
and tokentrT a [Ast.Variable x] = token_trans a x |
ddf6bab64b96
token translations: context dependent, result Pretty.T;
wenzelm
parents:
25393
diff
changeset
|
392 |
| tokentrT _ _ = NONE |
ddf6bab64b96
token translations: context dependent, result Pretty.T;
wenzelm
parents:
25393
diff
changeset
|
393 |
|
5691 | 394 |
and astT (c as Ast.Constant a, p) = combT (c, a, [], p) |
8457 | 395 |
| astT (Ast.Variable x, _) = [Pretty.str x] |
6280 | 396 |
| astT (Ast.Appl ((c as Ast.Constant a) :: (args as _ :: _)), p) = combT (c, a, args, p) |
5691 | 397 |
| astT (Ast.Appl (f :: (args as _ :: _)), p) = astT (appT (f, args), p) |
398 |
| astT (ast as Ast.Appl _, _) = raise Ast.AST ("pretty: malformed ast", [ast]); |
|
6280 | 399 |
in astT (ast0, p0) end; |
0 | 400 |
|
401 |
||
402 |
(* pretty_term_ast *) |
|
403 |
||
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
404 |
fun pretty_term_ast extern ctxt curried prtabs trf tokentrf ast = |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
405 |
pretty extern ctxt (mode_tabs prtabs (print_mode_value ())) |
23630
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
wenzelm
parents:
23615
diff
changeset
|
406 |
trf tokentrf false curried ast 0; |
0 | 407 |
|
408 |
||
409 |
(* pretty_typ_ast *) |
|
410 |
||
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
411 |
fun pretty_typ_ast {extern_class, extern_type} ctxt _ prtabs trf tokentrf ast = |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
412 |
pretty {extern_class = extern_class, extern_type = extern_type, extern_const = I} |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35262
diff
changeset
|
413 |
ctxt (mode_tabs prtabs (print_mode_value ())) |
23630
bc22daeed49e
pretty: markup for syntax/name of authentic consts;
wenzelm
parents:
23615
diff
changeset
|
414 |
trf tokentrf true false ast 0; |
0 | 415 |
|
416 |
end; |