| author | desharna | 
| Thu, 07 Oct 2021 10:34:48 +0200 | |
| changeset 74474 | 253c98aa935a | 
| parent 74231 | b3c65c984210 | 
| child 75353 | 05f7f5454cb6 | 
| permissions | -rw-r--r-- | 
| 37744 | 1  | 
(* Title: Tools/Code/code_printer.ML  | 
| 28060 | 2  | 
Author: Florian Haftmann, TU Muenchen  | 
3  | 
||
4  | 
Generic operations for pretty printing of target language code.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
signature CODE_PRINTER =  | 
|
8  | 
sig  | 
|
| 32908 | 9  | 
type itype = Code_Thingol.itype  | 
10  | 
type iterm = Code_Thingol.iterm  | 
|
11  | 
type const = Code_Thingol.const  | 
|
12  | 
type dict = Code_Thingol.dict  | 
|
13  | 
||
| 55236 | 14  | 
val eqn_error: theory -> thm option -> string -> 'a  | 
| 28060 | 15  | 
|
16  | 
val @@ : 'a * 'a -> 'a list  | 
|
17  | 
val @| : 'a list * 'a -> 'a list  | 
|
18  | 
val str: string -> Pretty.T  | 
|
19  | 
val concat: Pretty.T list -> Pretty.T  | 
|
20  | 
val brackets: Pretty.T list -> Pretty.T  | 
|
| 
34178
 
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
 
haftmann 
parents: 
34176 
diff
changeset
 | 
21  | 
val enclose: string -> string -> Pretty.T list -> Pretty.T  | 
| 38778 | 22  | 
val commas: Pretty.T list -> Pretty.T list  | 
| 
34178
 
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
 
haftmann 
parents: 
34176 
diff
changeset
 | 
23  | 
val enum: string -> string -> string -> Pretty.T list -> Pretty.T  | 
| 
 
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
 
haftmann 
parents: 
34176 
diff
changeset
 | 
24  | 
val enum_default: string -> string -> string -> string -> Pretty.T list -> Pretty.T  | 
| 28060 | 25  | 
val semicolon: Pretty.T list -> Pretty.T  | 
| 
33989
 
cb136b5f6050
more speaking function names for Code_Printer; added doublesemicolon
 
haftmann 
parents: 
32924 
diff
changeset
 | 
26  | 
val doublesemicolon: Pretty.T list -> Pretty.T  | 
| 
34178
 
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
 
haftmann 
parents: 
34176 
diff
changeset
 | 
27  | 
val indent: int -> Pretty.T -> Pretty.T  | 
| 55150 | 28  | 
val markup_stmt: Code_Symbol.T -> Pretty.T -> Pretty.T  | 
29  | 
val format: Code_Symbol.T list -> int -> Pretty.T -> string  | 
|
| 28060 | 30  | 
|
| 
30648
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30161 
diff
changeset
 | 
31  | 
type var_ctxt  | 
| 
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30161 
diff
changeset
 | 
32  | 
val make_vars: string list -> var_ctxt  | 
| 
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30161 
diff
changeset
 | 
33  | 
val intro_vars: string list -> var_ctxt -> var_ctxt  | 
| 
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30161 
diff
changeset
 | 
34  | 
val lookup_var: var_ctxt -> string -> string  | 
| 32913 | 35  | 
val intro_base_names: (string -> bool) -> (string -> string)  | 
36  | 
-> string list -> var_ctxt -> var_ctxt  | 
|
| 55150 | 37  | 
val intro_base_names_for: (string -> bool) -> (Code_Symbol.T -> string)  | 
| 
55145
 
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
 
haftmann 
parents: 
52433 
diff
changeset
 | 
38  | 
-> iterm list -> var_ctxt -> var_ctxt  | 
| 32908 | 39  | 
val aux_params: var_ctxt -> iterm list list -> string list  | 
| 
30648
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30161 
diff
changeset
 | 
40  | 
|
| 
31056
 
01ac77eb660b
robustifed infrastructure for complex term syntax during code generation
 
haftmann 
parents: 
30648 
diff
changeset
 | 
41  | 
type literals  | 
| 68028 | 42  | 
  val Literals: { literal_string: string -> string,
 | 
| 
37958
 
9728342bcd56
another refinement chapter in the neverending numeral story
 
haftmann 
parents: 
37899 
diff
changeset
 | 
43  | 
literal_numeral: int -> string,  | 
| 
31056
 
01ac77eb660b
robustifed infrastructure for complex term syntax during code generation
 
haftmann 
parents: 
30648 
diff
changeset
 | 
44  | 
literal_list: Pretty.T list -> Pretty.T, infix_cons: int * string }  | 
| 
 
01ac77eb660b
robustifed infrastructure for complex term syntax during code generation
 
haftmann 
parents: 
30648 
diff
changeset
 | 
45  | 
-> literals  | 
| 
 
01ac77eb660b
robustifed infrastructure for complex term syntax during code generation
 
haftmann 
parents: 
30648 
diff
changeset
 | 
46  | 
val literal_string: literals -> string -> string  | 
| 
34944
 
970e1466028d
code literals: distinguish numeral classes by different entries
 
haftmann 
parents: 
34247 
diff
changeset
 | 
47  | 
val literal_numeral: literals -> int -> string  | 
| 
31056
 
01ac77eb660b
robustifed infrastructure for complex term syntax during code generation
 
haftmann 
parents: 
30648 
diff
changeset
 | 
48  | 
val literal_list: literals -> Pretty.T list -> Pretty.T  | 
| 
 
01ac77eb660b
robustifed infrastructure for complex term syntax during code generation
 
haftmann 
parents: 
30648 
diff
changeset
 | 
49  | 
val infix_cons: literals -> int * string  | 
| 
 
01ac77eb660b
robustifed infrastructure for complex term syntax during code generation
 
haftmann 
parents: 
30648 
diff
changeset
 | 
50  | 
|
| 28060 | 51  | 
type lrx  | 
52  | 
val L: lrx  | 
|
53  | 
val R: lrx  | 
|
54  | 
val X: lrx  | 
|
55  | 
type fixity  | 
|
56  | 
val BR: fixity  | 
|
57  | 
val NOBR: fixity  | 
|
58  | 
val INFX: int * lrx -> fixity  | 
|
59  | 
val APP: fixity  | 
|
60  | 
val brackify: fixity -> Pretty.T list -> Pretty.T  | 
|
| 
37242
 
97097e589715
brackify_infix etc.: no break before infix operator -- eases survival in Scala
 
haftmann 
parents: 
37146 
diff
changeset
 | 
61  | 
val brackify_infix: int * lrx -> fixity -> Pretty.T * Pretty.T * Pretty.T -> Pretty.T  | 
| 31665 | 62  | 
val brackify_block: fixity -> Pretty.T -> Pretty.T list -> Pretty.T -> Pretty.T  | 
| 
50626
 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 
haftmann 
parents: 
50618 
diff
changeset
 | 
63  | 
  val gen_applify: bool -> string -> string -> ('a -> Pretty.T) -> fixity -> Pretty.T -> 'a list -> Pretty.T
 | 
| 37638 | 64  | 
  val applify: string -> string -> ('a -> Pretty.T) -> fixity -> Pretty.T -> 'a list -> Pretty.T
 | 
| 38922 | 65  | 
val tuplify: (fixity -> 'a -> Pretty.T) -> fixity -> 'a list -> Pretty.T option  | 
66  | 
||
| 
55153
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
67  | 
type raw_const_syntax  | 
| 
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
68  | 
val plain_const_syntax: string -> raw_const_syntax  | 
| 34152 | 69  | 
type simple_const_syntax  | 
| 
55153
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
70  | 
val simple_const_syntax: simple_const_syntax -> raw_const_syntax  | 
| 37881 | 71  | 
type complex_const_syntax  | 
| 
55153
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
72  | 
val complex_const_syntax: complex_const_syntax -> raw_const_syntax  | 
| 
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
73  | 
val parse_const_syntax: raw_const_syntax parser  | 
| 
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
74  | 
val requires_args: raw_const_syntax -> int  | 
| 
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
75  | 
datatype const_printer = Plain_printer of string  | 
| 
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
76  | 
| Complex_printer of (var_ctxt -> fixity -> iterm -> Pretty.T)  | 
| 
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
77  | 
-> thm option -> var_ctxt -> fixity -> (iterm * itype) list -> Pretty.T  | 
| 
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
78  | 
type const_syntax = int * const_printer  | 
| 
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
79  | 
val prep_const_syntax: theory -> literals  | 
| 
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
80  | 
-> string -> raw_const_syntax -> const_syntax  | 
| 
52070
 
fd497099f5f5
infrastructure for generic data for code symbols (constants, type constructors, classes, instances)
 
haftmann 
parents: 
52069 
diff
changeset
 | 
81  | 
type tyco_syntax  | 
| 52079 | 82  | 
val parse_tyco_syntax: tyco_syntax parser  | 
| 35228 | 83  | 
val gen_print_app: (thm option -> var_ctxt -> const * iterm list -> Pretty.T list)  | 
84  | 
-> (thm option -> var_ctxt -> fixity -> iterm -> Pretty.T)  | 
|
| 
55153
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
85  | 
-> (string -> const_syntax option)  | 
| 35228 | 86  | 
-> thm option -> var_ctxt -> fixity -> const * iterm list -> Pretty.T  | 
87  | 
val gen_print_bind: (thm option -> var_ctxt -> fixity -> iterm -> Pretty.T)  | 
|
88  | 
-> thm option -> fixity  | 
|
| 31889 | 89  | 
-> iterm -> var_ctxt -> Pretty.T * var_ctxt  | 
| 59103 | 90  | 
|
91  | 
type identifiers  | 
|
92  | 
type printings  | 
|
93  | 
type data  | 
|
94  | 
val empty_data: data  | 
|
95  | 
val map_data: (string list * identifiers * printings  | 
|
96  | 
-> string list * identifiers * printings)  | 
|
97  | 
-> data -> data  | 
|
98  | 
val merge_data: data * data -> data  | 
|
99  | 
val the_reserved: data -> string list;  | 
|
100  | 
val the_identifiers: data -> identifiers;  | 
|
101  | 
val the_printings: data -> printings;  | 
|
| 28060 | 102  | 
end;  | 
103  | 
||
104  | 
structure Code_Printer : CODE_PRINTER =  | 
|
105  | 
struct  | 
|
106  | 
||
| 55150 | 107  | 
open Basic_Code_Symbol;  | 
| 28060 | 108  | 
open Code_Thingol;  | 
109  | 
||
| 
39034
 
ebeb48fd653b
formal framework for presentation of selected statements
 
haftmann 
parents: 
38923 
diff
changeset
 | 
110  | 
(** generic nonsense *)  | 
| 
 
ebeb48fd653b
formal framework for presentation of selected statements
 
haftmann 
parents: 
38923 
diff
changeset
 | 
111  | 
|
| 55236 | 112  | 
fun eqn_error thy (SOME thm) s =  | 
| 61268 | 113  | 
error (s ^ ",\nin equation " ^ Thm.string_of_thm_global thy thm)  | 
| 55236 | 114  | 
| eqn_error _ NONE s = error s;  | 
| 28060 | 115  | 
|
| 
39034
 
ebeb48fd653b
formal framework for presentation of selected statements
 
haftmann 
parents: 
38923 
diff
changeset
 | 
116  | 
val code_presentationN = "code_presentation";  | 
| 
39057
 
c6d146ed07ae
manage statement selection for presentation wholly through markup
 
haftmann 
parents: 
39056 
diff
changeset
 | 
117  | 
val stmt_nameN = "stmt_name";  | 
| 
39034
 
ebeb48fd653b
formal framework for presentation of selected statements
 
haftmann 
parents: 
38923 
diff
changeset
 | 
118  | 
val _ = Markup.add_mode code_presentationN YXML.output_markup;  | 
| 
 
ebeb48fd653b
formal framework for presentation of selected statements
 
haftmann 
parents: 
38923 
diff
changeset
 | 
119  | 
|
| 
 
ebeb48fd653b
formal framework for presentation of selected statements
 
haftmann 
parents: 
38923 
diff
changeset
 | 
120  | 
|
| 34071 | 121  | 
(** assembling and printing text pieces **)  | 
| 28060 | 122  | 
|
123  | 
infixr 5 @@;  | 
|
124  | 
infixr 5 @|;  | 
|
125  | 
fun x @@ y = [x, y];  | 
|
126  | 
fun xs @| y = xs @ [y];  | 
|
| 69659 | 127  | 
fun with_no_print_mode f = Print_Mode.setmp [] f;  | 
128  | 
val str = with_no_print_mode Pretty.str;  | 
|
| 28060 | 129  | 
val concat = Pretty.block o Pretty.breaks;  | 
| 69659 | 130  | 
val commas = with_no_print_mode Pretty.commas;  | 
131  | 
fun enclose l r = with_no_print_mode (Pretty.enclose l r);  | 
|
| 
34178
 
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
 
haftmann 
parents: 
34176 
diff
changeset
 | 
132  | 
val brackets = enclose "(" ")" o Pretty.breaks;
 | 
| 69659 | 133  | 
fun enum sep l r = with_no_print_mode (Pretty.enum sep l r);  | 
| 
34178
 
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
 
haftmann 
parents: 
34176 
diff
changeset
 | 
134  | 
fun enum_default default sep l r [] = str default  | 
| 
 
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
 
haftmann 
parents: 
34176 
diff
changeset
 | 
135  | 
| enum_default default sep l r xs = enum sep l r xs;  | 
| 28060 | 136  | 
fun semicolon ps = Pretty.block [concat ps, str ";"];  | 
| 
33989
 
cb136b5f6050
more speaking function names for Code_Printer; added doublesemicolon
 
haftmann 
parents: 
32924 
diff
changeset
 | 
137  | 
fun doublesemicolon ps = Pretty.block [concat ps, str ";;"];  | 
| 69659 | 138  | 
fun indent i = with_no_print_mode (Pretty.indent i);  | 
| 28060 | 139  | 
|
| 69659 | 140  | 
fun with_presentation_marker f = Print_Mode.setmp [code_presentationN] f;  | 
141  | 
||
142  | 
fun markup_stmt sym = with_presentation_marker  | 
|
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
143  | 
(Pretty.mark (code_presentationN, [(stmt_nameN, Code_Symbol.marker sym)]));  | 
| 39062 | 144  | 
|
| 74231 | 145  | 
fun filter_presentation [] xml =  | 
146  | 
Buffer.build (fold XML.add_content xml)  | 
|
147  | 
| filter_presentation presentation_syms xml =  | 
|
| 
39057
 
c6d146ed07ae
manage statement selection for presentation wholly through markup
 
haftmann 
parents: 
39056 
diff
changeset
 | 
148  | 
let  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
149  | 
val presentation_idents = map Code_Symbol.marker presentation_syms  | 
| 
43947
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
43326 
diff
changeset
 | 
150  | 
fun is_selected (name, attrs) =  | 
| 39558 | 151  | 
name = code_presentationN  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
152  | 
andalso member (op =) presentation_idents (the (Properties.get attrs stmt_nameN));  | 
| 39558 | 153  | 
fun add_content_with_space tree (is_first, buf) =  | 
154  | 
buf  | 
|
155  | 
|> not is_first ? Buffer.add "\n\n"  | 
|
156  | 
|> XML.add_content tree  | 
|
157  | 
|> pair false;  | 
|
158  | 
fun filter (XML.Elem (name_attrs, xs)) =  | 
|
159  | 
fold (if is_selected name_attrs then add_content_with_space else filter) xs  | 
|
| 47576 | 160  | 
| filter (XML.Text _) = I;  | 
| 74231 | 161  | 
in snd (fold filter xml (true, Buffer.empty)) end;  | 
| 39062 | 162  | 
|
163  | 
fun format presentation_names width =  | 
|
| 69659 | 164  | 
with_presentation_marker (Pretty.string_of_margin width)  | 
| 39062 | 165  | 
#> YXML.parse_body  | 
| 39558 | 166  | 
#> filter_presentation presentation_names  | 
| 39678 | 167  | 
#> Buffer.add "\n"  | 
| 39558 | 168  | 
#> Buffer.content;  | 
| 34071 | 169  | 
|
| 28060 | 170  | 
|
| 
30648
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30161 
diff
changeset
 | 
171  | 
(** names and variable name contexts **)  | 
| 
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30161 
diff
changeset
 | 
172  | 
|
| 
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30161 
diff
changeset
 | 
173  | 
type var_ctxt = string Symtab.table * Name.context;  | 
| 
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30161 
diff
changeset
 | 
174  | 
|
| 
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30161 
diff
changeset
 | 
175  | 
fun make_vars names = (fold (fn name => Symtab.update_new (name, name)) names Symtab.empty,  | 
| 
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30161 
diff
changeset
 | 
176  | 
Name.make_context names);  | 
| 
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30161 
diff
changeset
 | 
177  | 
|
| 
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30161 
diff
changeset
 | 
178  | 
fun intro_vars names (namemap, namectxt) =  | 
| 
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30161 
diff
changeset
 | 
179  | 
let  | 
| 
43326
 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 
wenzelm 
parents: 
43324 
diff
changeset
 | 
180  | 
val (names', namectxt') = fold_map Name.variant names namectxt;  | 
| 
30648
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30161 
diff
changeset
 | 
181  | 
val namemap' = fold2 (curry Symtab.update) names names' namemap;  | 
| 
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30161 
diff
changeset
 | 
182  | 
in (namemap', namectxt') end;  | 
| 
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30161 
diff
changeset
 | 
183  | 
|
| 
43947
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
43326 
diff
changeset
 | 
184  | 
fun lookup_var (namemap, _) name =  | 
| 
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
43326 
diff
changeset
 | 
185  | 
case Symtab.lookup namemap name of  | 
| 
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
43326 
diff
changeset
 | 
186  | 
SOME name' => name'  | 
| 
30648
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30161 
diff
changeset
 | 
187  | 
  | NONE => error ("Invalid name in context: " ^ quote name);
 | 
| 
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30161 
diff
changeset
 | 
188  | 
|
| 32908 | 189  | 
fun aux_params vars lhss =  | 
190  | 
let  | 
|
191  | 
fun fish_param _ (w as SOME _) = w  | 
|
192  | 
| fish_param (IVar (SOME v)) NONE = SOME v  | 
|
193  | 
| fish_param _ NONE = NONE;  | 
|
194  | 
fun fillup_param _ (_, SOME v) = v  | 
|
195  | 
| fillup_param x (i, NONE) = x ^ string_of_int i;  | 
|
196  | 
val fished1 = fold (map2 fish_param) lhss (replicate (length (hd lhss)) NONE);  | 
|
| 
43324
 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
197  | 
val x = singleton (Name.variant_list (map_filter I fished1)) "x";  | 
| 32908 | 198  | 
val fished2 = map_index (fillup_param x) fished1;  | 
| 
43326
 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 
wenzelm 
parents: 
43324 
diff
changeset
 | 
199  | 
val (fished3, _) = fold_map Name.variant fished2 Name.context;  | 
| 32908 | 200  | 
val vars' = intro_vars fished3 vars;  | 
201  | 
in map (lookup_var vars') fished3 end;  | 
|
202  | 
||
| 
55145
 
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
 
haftmann 
parents: 
52433 
diff
changeset
 | 
203  | 
fun intro_base_names no_syntax deresolve =  | 
| 
 
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
 
haftmann 
parents: 
52433 
diff
changeset
 | 
204  | 
map_filter (fn name => if no_syntax name then  | 
| 32913 | 205  | 
let val name' = deresolve name in  | 
206  | 
if Long_Name.is_qualified name' then NONE else SOME name'  | 
|
207  | 
end else NONE)  | 
|
| 
55145
 
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
 
haftmann 
parents: 
52433 
diff
changeset
 | 
208  | 
#> intro_vars;  | 
| 
 
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
 
haftmann 
parents: 
52433 
diff
changeset
 | 
209  | 
|
| 
 
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
 
haftmann 
parents: 
52433 
diff
changeset
 | 
210  | 
fun intro_base_names_for no_syntax deresolve ts =  | 
| 
 
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
 
haftmann 
parents: 
52433 
diff
changeset
 | 
211  | 
[]  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
212  | 
|> fold Code_Thingol.add_constsyms ts  | 
| 55150 | 213  | 
|> intro_base_names (fn Constant const => no_syntax const | _ => true) deresolve;  | 
| 32913 | 214  | 
|
| 
30648
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30161 
diff
changeset
 | 
215  | 
|
| 
31056
 
01ac77eb660b
robustifed infrastructure for complex term syntax during code generation
 
haftmann 
parents: 
30648 
diff
changeset
 | 
216  | 
(** pretty literals **)  | 
| 
 
01ac77eb660b
robustifed infrastructure for complex term syntax during code generation
 
haftmann 
parents: 
30648 
diff
changeset
 | 
217  | 
|
| 
 
01ac77eb660b
robustifed infrastructure for complex term syntax during code generation
 
haftmann 
parents: 
30648 
diff
changeset
 | 
218  | 
datatype literals = Literals of {
 | 
| 
 
01ac77eb660b
robustifed infrastructure for complex term syntax during code generation
 
haftmann 
parents: 
30648 
diff
changeset
 | 
219  | 
literal_string: string -> string,  | 
| 
34944
 
970e1466028d
code literals: distinguish numeral classes by different entries
 
haftmann 
parents: 
34247 
diff
changeset
 | 
220  | 
literal_numeral: int -> string,  | 
| 
31056
 
01ac77eb660b
robustifed infrastructure for complex term syntax during code generation
 
haftmann 
parents: 
30648 
diff
changeset
 | 
221  | 
literal_list: Pretty.T list -> Pretty.T,  | 
| 
 
01ac77eb660b
robustifed infrastructure for complex term syntax during code generation
 
haftmann 
parents: 
30648 
diff
changeset
 | 
222  | 
infix_cons: int * string  | 
| 
 
01ac77eb660b
robustifed infrastructure for complex term syntax during code generation
 
haftmann 
parents: 
30648 
diff
changeset
 | 
223  | 
};  | 
| 
 
01ac77eb660b
robustifed infrastructure for complex term syntax during code generation
 
haftmann 
parents: 
30648 
diff
changeset
 | 
224  | 
|
| 
 
01ac77eb660b
robustifed infrastructure for complex term syntax during code generation
 
haftmann 
parents: 
30648 
diff
changeset
 | 
225  | 
fun dest_Literals (Literals lits) = lits;  | 
| 
 
01ac77eb660b
robustifed infrastructure for complex term syntax during code generation
 
haftmann 
parents: 
30648 
diff
changeset
 | 
226  | 
|
| 
 
01ac77eb660b
robustifed infrastructure for complex term syntax during code generation
 
haftmann 
parents: 
30648 
diff
changeset
 | 
227  | 
val literal_string = #literal_string o dest_Literals;  | 
| 
 
01ac77eb660b
robustifed infrastructure for complex term syntax during code generation
 
haftmann 
parents: 
30648 
diff
changeset
 | 
228  | 
val literal_numeral = #literal_numeral o dest_Literals;  | 
| 
 
01ac77eb660b
robustifed infrastructure for complex term syntax during code generation
 
haftmann 
parents: 
30648 
diff
changeset
 | 
229  | 
val literal_list = #literal_list o dest_Literals;  | 
| 
 
01ac77eb660b
robustifed infrastructure for complex term syntax during code generation
 
haftmann 
parents: 
30648 
diff
changeset
 | 
230  | 
val infix_cons = #infix_cons o dest_Literals;  | 
| 
 
01ac77eb660b
robustifed infrastructure for complex term syntax during code generation
 
haftmann 
parents: 
30648 
diff
changeset
 | 
231  | 
|
| 
 
01ac77eb660b
robustifed infrastructure for complex term syntax during code generation
 
haftmann 
parents: 
30648 
diff
changeset
 | 
232  | 
|
| 28060 | 233  | 
(** syntax printer **)  | 
234  | 
||
235  | 
(* binding priorities *)  | 
|
236  | 
||
237  | 
datatype lrx = L | R | X;  | 
|
238  | 
||
239  | 
datatype fixity =  | 
|
240  | 
BR  | 
|
241  | 
| NOBR  | 
|
242  | 
| INFX of (int * lrx);  | 
|
243  | 
||
244  | 
val APP = INFX (~1, L);  | 
|
245  | 
||
246  | 
fun fixity_lrx L L = false  | 
|
247  | 
| fixity_lrx R R = false  | 
|
248  | 
| fixity_lrx _ _ = true;  | 
|
249  | 
||
250  | 
fun fixity NOBR _ = false  | 
|
251  | 
| fixity _ NOBR = false  | 
|
| 52433 | 252  | 
| fixity (INFX (pr, lr)) (INFX (pr_ctxt, lr_ctxt)) =  | 
253  | 
pr < pr_ctxt  | 
|
254  | 
orelse pr = pr_ctxt  | 
|
| 28060 | 255  | 
andalso fixity_lrx lr lr_ctxt  | 
| 52433 | 256  | 
orelse pr_ctxt = ~1  | 
| 28060 | 257  | 
| fixity BR (INFX _) = false  | 
258  | 
| fixity _ _ = true;  | 
|
259  | 
||
260  | 
fun gen_brackify _ [p] = p  | 
|
| 
34178
 
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
 
haftmann 
parents: 
34176 
diff
changeset
 | 
261  | 
  | gen_brackify true (ps as _::_) = enclose "(" ")" ps
 | 
| 28060 | 262  | 
| gen_brackify false (ps as _::_) = Pretty.block ps;  | 
263  | 
||
264  | 
fun brackify fxy_ctxt =  | 
|
265  | 
gen_brackify (fixity BR fxy_ctxt) o Pretty.breaks;  | 
|
266  | 
||
| 
37242
 
97097e589715
brackify_infix etc.: no break before infix operator -- eases survival in Scala
 
haftmann 
parents: 
37146 
diff
changeset
 | 
267  | 
fun brackify_infix infx fxy_ctxt (l, m, r) =  | 
| 50627 | 268  | 
gen_brackify (fixity (INFX infx) fxy_ctxt) [l, str " ", m, Pretty.brk 1, r];  | 
| 28060 | 269  | 
|
| 31665 | 270  | 
fun brackify_block fxy_ctxt p1 ps p2 =  | 
271  | 
let val p = Pretty.block_enclose (p1, p2) ps  | 
|
272  | 
in if fixity BR fxy_ctxt  | 
|
| 
34178
 
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
 
haftmann 
parents: 
34176 
diff
changeset
 | 
273  | 
    then enclose "(" ")" [p]
 | 
| 31665 | 274  | 
else p  | 
275  | 
end;  | 
|
276  | 
||
| 
50626
 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 
haftmann 
parents: 
50618 
diff
changeset
 | 
277  | 
fun gen_applify strict opn cls f fxy_ctxt p [] =  | 
| 
 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 
haftmann 
parents: 
50618 
diff
changeset
 | 
278  | 
if strict  | 
| 50628 | 279  | 
then gen_brackify (fixity BR fxy_ctxt) [p, str (opn ^ cls)]  | 
| 
50626
 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 
haftmann 
parents: 
50618 
diff
changeset
 | 
280  | 
else p  | 
| 
 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 
haftmann 
parents: 
50618 
diff
changeset
 | 
281  | 
| gen_applify strict opn cls f fxy_ctxt p ps =  | 
| 50627 | 282  | 
gen_brackify (fixity BR fxy_ctxt) (p @@ enum "," opn cls (map f ps));  | 
| 34247 | 283  | 
|
| 
50626
 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 
haftmann 
parents: 
50618 
diff
changeset
 | 
284  | 
fun applify opn = gen_applify false opn;  | 
| 
 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 
haftmann 
parents: 
50618 
diff
changeset
 | 
285  | 
|
| 38922 | 286  | 
fun tuplify _ _ [] = NONE  | 
287  | 
| tuplify print fxy [x] = SOME (print fxy x)  | 
|
288  | 
  | tuplify print _ xs = SOME (enum "," "(" ")" (map (print NOBR) xs));
 | 
|
289  | 
||
| 28060 | 290  | 
|
| 
52070
 
fd497099f5f5
infrastructure for generic data for code symbols (constants, type constructors, classes, instances)
 
haftmann 
parents: 
52069 
diff
changeset
 | 
291  | 
(* generic syntax *)  | 
| 34152 | 292  | 
|
293  | 
type simple_const_syntax = int * ((fixity -> iterm -> Pretty.T)  | 
|
294  | 
-> fixity -> (iterm * itype) list -> Pretty.T);  | 
|
| 37881 | 295  | 
|
| 
55148
 
7e1b7cb54114
avoid (now superfluous) indirect passing of constant names
 
haftmann 
parents: 
55147 
diff
changeset
 | 
296  | 
type complex_const_syntax = int * (literals  | 
| 
31056
 
01ac77eb660b
robustifed infrastructure for complex term syntax during code generation
 
haftmann 
parents: 
30648 
diff
changeset
 | 
297  | 
-> (var_ctxt -> fixity -> iterm -> Pretty.T)  | 
| 
55148
 
7e1b7cb54114
avoid (now superfluous) indirect passing of constant names
 
haftmann 
parents: 
55147 
diff
changeset
 | 
298  | 
-> thm option -> var_ctxt -> fixity -> (iterm * itype) list -> Pretty.T);  | 
| 37881 | 299  | 
|
| 
55153
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
300  | 
datatype raw_const_syntax = plain_const_syntax of string  | 
| 37881 | 301  | 
| complex_const_syntax of complex_const_syntax;  | 
302  | 
||
| 34176 | 303  | 
fun simple_const_syntax syn =  | 
| 
43947
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
43326 
diff
changeset
 | 
304  | 
complex_const_syntax  | 
| 
55148
 
7e1b7cb54114
avoid (now superfluous) indirect passing of constant names
 
haftmann 
parents: 
55147 
diff
changeset
 | 
305  | 
(apsnd (fn f => fn _ => fn print => fn _ => fn vars => f (print vars)) syn);  | 
| 28060 | 306  | 
|
| 
55153
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
307  | 
fun requires_args (plain_const_syntax _) = 0  | 
| 
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
308  | 
| requires_args (complex_const_syntax (k, _)) = k;  | 
| 
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
309  | 
|
| 
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
310  | 
datatype const_printer = Plain_printer of string  | 
| 
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
311  | 
| Complex_printer of (var_ctxt -> fixity -> iterm -> Pretty.T)  | 
| 
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
312  | 
-> thm option -> var_ctxt -> fixity -> (iterm * itype) list -> Pretty.T;  | 
| 37881 | 313  | 
|
| 
55153
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
314  | 
type const_syntax = int * const_printer;  | 
| 
31056
 
01ac77eb660b
robustifed infrastructure for complex term syntax during code generation
 
haftmann 
parents: 
30648 
diff
changeset
 | 
315  | 
|
| 
55153
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
316  | 
fun prep_const_syntax thy literals c (plain_const_syntax s) =  | 
| 
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
317  | 
(Code.args_number thy c, Plain_printer s)  | 
| 
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
318  | 
| prep_const_syntax thy literals c (complex_const_syntax (n, f))=  | 
| 
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
319  | 
(n, Complex_printer (f literals));  | 
| 37881 | 320  | 
|
| 
43947
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
43326 
diff
changeset
 | 
321  | 
fun gen_print_app print_app_expr print_term const_syntax some_thm vars fxy  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
322  | 
    (app as ({ sym, dom, ... }, ts)) =
 | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
323  | 
case sym of  | 
| 55150 | 324  | 
Constant const => (case const_syntax const of  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
325  | 
NONE => brackify fxy (print_app_expr some_thm vars app)  | 
| 
55153
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
326  | 
| SOME (_, Plain_printer s) =>  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
327  | 
brackify fxy (str s :: map (print_term some_thm vars BR) ts)  | 
| 
55153
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
328  | 
| SOME (k, Complex_printer print) =>  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
329  | 
let  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
330  | 
fun print' fxy ts =  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
331  | 
print (print_term some_thm) some_thm vars fxy (ts ~~ take k dom);  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
332  | 
in  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
333  | 
if k = length ts  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
334  | 
then print' fxy ts  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
335  | 
else if k < length ts  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
336  | 
then case chop k ts of (ts1, ts2) =>  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
337  | 
brackify fxy (print' APP ts1 :: map (print_term some_thm vars BR) ts2)  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
338  | 
else print_term some_thm vars fxy (Code_Thingol.eta_expand k app)  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
339  | 
end)  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55145 
diff
changeset
 | 
340  | 
| _ => brackify fxy (print_app_expr some_thm vars app);  | 
| 28060 | 341  | 
|
| 
33989
 
cb136b5f6050
more speaking function names for Code_Printer; added doublesemicolon
 
haftmann 
parents: 
32924 
diff
changeset
 | 
342  | 
fun gen_print_bind print_term thm (fxy : fixity) pat vars =  | 
| 28060 | 343  | 
let  | 
| 31889 | 344  | 
val vs = Code_Thingol.fold_varnames (insert (op =)) pat [];  | 
| 31874 | 345  | 
val vars' = intro_vars vs vars;  | 
| 
33989
 
cb136b5f6050
more speaking function names for Code_Printer; added doublesemicolon
 
haftmann 
parents: 
32924 
diff
changeset
 | 
346  | 
in (print_term thm vars' fxy pat, vars') end;  | 
| 28060 | 347  | 
|
| 
52070
 
fd497099f5f5
infrastructure for generic data for code symbols (constants, type constructors, classes, instances)
 
haftmann 
parents: 
52069 
diff
changeset
 | 
348  | 
type tyco_syntax = int * ((fixity -> itype -> Pretty.T)  | 
| 
 
fd497099f5f5
infrastructure for generic data for code symbols (constants, type constructors, classes, instances)
 
haftmann 
parents: 
52069 
diff
changeset
 | 
349  | 
-> fixity -> itype list -> Pretty.T);  | 
| 
 
fd497099f5f5
infrastructure for generic data for code symbols (constants, type constructors, classes, instances)
 
haftmann 
parents: 
52069 
diff
changeset
 | 
350  | 
|
| 28060 | 351  | 
|
352  | 
(* mixfix syntax *)  | 
|
353  | 
||
354  | 
datatype 'a mixfix =  | 
|
355  | 
Arg of fixity  | 
|
| 37881 | 356  | 
| String of string  | 
357  | 
| Break;  | 
|
| 28060 | 358  | 
|
| 52069 | 359  | 
fun printer_of_mixfix prep_arg (fixity_this, mfx) =  | 
| 28060 | 360  | 
let  | 
361  | 
fun is_arg (Arg _) = true  | 
|
362  | 
| is_arg _ = false;  | 
|
363  | 
val i = (length o filter is_arg) mfx;  | 
|
364  | 
fun fillin _ [] [] =  | 
|
365  | 
[]  | 
|
| 
33989
 
cb136b5f6050
more speaking function names for Code_Printer; added doublesemicolon
 
haftmann 
parents: 
32924 
diff
changeset
 | 
366  | 
| fillin print (Arg fxy :: mfx) (a :: args) =  | 
| 
 
cb136b5f6050
more speaking function names for Code_Printer; added doublesemicolon
 
haftmann 
parents: 
32924 
diff
changeset
 | 
367  | 
(print fxy o prep_arg) a :: fillin print mfx args  | 
| 37881 | 368  | 
| fillin print (String s :: mfx) args =  | 
369  | 
str s :: fillin print mfx args  | 
|
370  | 
| fillin print (Break :: mfx) args =  | 
|
371  | 
Pretty.brk 1 :: fillin print mfx args;  | 
|
| 28060 | 372  | 
in  | 
| 
33989
 
cb136b5f6050
more speaking function names for Code_Printer; added doublesemicolon
 
haftmann 
parents: 
32924 
diff
changeset
 | 
373  | 
(i, fn print => fn fixity_ctxt => fn args =>  | 
| 
 
cb136b5f6050
more speaking function names for Code_Printer; added doublesemicolon
 
haftmann 
parents: 
32924 
diff
changeset
 | 
374  | 
gen_brackify (fixity fixity_this fixity_ctxt) (fillin print mfx args))  | 
| 28060 | 375  | 
end;  | 
376  | 
||
| 52069 | 377  | 
fun read_infix (fixity_this, i) s =  | 
| 28060 | 378  | 
let  | 
| 52069 | 379  | 
val l = case fixity_this of L => INFX (i, L) | _ => INFX (i, X);  | 
380  | 
val r = case fixity_this of R => INFX (i, R) | _ => INFX (i, X);  | 
|
| 28060 | 381  | 
in  | 
| 52069 | 382  | 
(INFX (i, fixity_this), [Arg l, String " ", String s, Break, Arg r])  | 
| 28060 | 383  | 
end;  | 
384  | 
||
| 52069 | 385  | 
fun read_mixfix s =  | 
| 28060 | 386  | 
let  | 
| 58854 | 387  | 
val sym_any = Scan.one Symbol.not_eof;  | 
| 52069 | 388  | 
val parse = Scan.optional ($$ "!" >> K NOBR) BR -- Scan.repeat (  | 
| 28060 | 389  | 
         ($$ "(" -- $$ "_" -- $$ ")" >> K (Arg NOBR))
 | 
390  | 
|| ($$ "_" >> K (Arg BR))  | 
|
| 37881 | 391  | 
|| ($$ "/" |-- Scan.repeat ($$ " ") >> (K Break))  | 
| 28060 | 392  | 
|| (Scan.repeat1  | 
393  | 
( $$ "'" |-- sym_any  | 
|
394  | 
            || Scan.unless ($$ "_" || $$ "/" || $$ "(" |-- $$ "_" |-- $$ ")")
 | 
|
| 37881 | 395  | 
sym_any) >> (String o implode)));  | 
| 
43947
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
43326 
diff
changeset
 | 
396  | 
fun err s (_, NONE) = (fn () => "malformed mixfix annotation: " ^ quote s)  | 
| 
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
43326 
diff
changeset
 | 
397  | 
| err _ (_, SOME msg) = msg;  | 
| 
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
43326 
diff
changeset
 | 
398  | 
in  | 
| 
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
43326 
diff
changeset
 | 
399  | 
case Scan.finite Symbol.stopper parse (Symbol.explode s) of  | 
| 52069 | 400  | 
(fixity_mixfix, []) => fixity_mixfix  | 
401  | 
| _ => Scan.!! (err s) Scan.fail ()  | 
|
| 28060 | 402  | 
end;  | 
403  | 
||
| 52069 | 404  | 
val parse_fixity =  | 
| 69593 | 405  | 
(\<^keyword>\<open>infix\<close> >> K X) || (\<^keyword>\<open>infixl\<close> >> K L) || (\<^keyword>\<open>infixr\<close> >> K R)  | 
| 52069 | 406  | 
|
| 52099 | 407  | 
fun parse_mixfix x =  | 
408  | 
(Parse.string >> read_mixfix  | 
|
| 52069 | 409  | 
|| parse_fixity -- Parse.nat -- Parse.string  | 
| 52099 | 410  | 
>> (fn ((fixity, i), s) => read_infix (fixity, i) s)) x;  | 
| 28060 | 411  | 
|
| 52069 | 412  | 
fun syntax_of_mixfix of_plain of_printer prep_arg (BR, [String s]) = of_plain s  | 
413  | 
| syntax_of_mixfix of_plain of_printer prep_arg (fixity, mfx) =  | 
|
414  | 
of_printer (printer_of_mixfix prep_arg (fixity, mfx));  | 
|
| 37881 | 415  | 
|
| 52099 | 416  | 
fun parse_tyco_syntax x =  | 
417  | 
(parse_mixfix >> syntax_of_mixfix (fn s => (0, (K o K o K o str) s)) I I) x;  | 
|
| 52069 | 418  | 
|
419  | 
val parse_const_syntax =  | 
|
420  | 
parse_mixfix >> syntax_of_mixfix plain_const_syntax simple_const_syntax fst;  | 
|
| 37881 | 421  | 
|
| 59103 | 422  | 
|
423  | 
(** custom data structure **)  | 
|
424  | 
||
425  | 
type identifiers = (string list * string, string list * string, string list * string, string list * string,  | 
|
426  | 
string list * string, string list * string) Code_Symbol.data;  | 
|
427  | 
type printings = (const_syntax, tyco_syntax, string, unit, unit,  | 
|
| 
69485
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
68028 
diff
changeset
 | 
428  | 
(Pretty.T * Code_Symbol.T list)) Code_Symbol.data;  | 
| 59103 | 429  | 
|
430  | 
datatype data = Data of { reserved: string list, identifiers: identifiers,
 | 
|
431  | 
printings: printings };  | 
|
432  | 
||
433  | 
fun make_data (reserved, identifiers, printings) =  | 
|
434  | 
  Data { reserved = reserved, identifiers = identifiers, printings = printings };
 | 
|
435  | 
val empty_data = make_data ([], Code_Symbol.empty_data, Code_Symbol.empty_data);  | 
|
436  | 
fun map_data f (Data { reserved, identifiers, printings }) =
 | 
|
437  | 
make_data (f (reserved, identifiers, printings));  | 
|
438  | 
fun merge_data (Data { reserved = reserved1, identifiers = identifiers1, printings = printings1 },
 | 
|
439  | 
    Data { reserved = reserved2, identifiers = identifiers2, printings = printings2 }) =
 | 
|
440  | 
make_data (merge (op =) (reserved1, reserved2),  | 
|
441  | 
Code_Symbol.merge_data (identifiers1, identifiers2), Code_Symbol.merge_data (printings1, printings2));  | 
|
442  | 
||
443  | 
fun the_reserved (Data { reserved, ... }) = reserved;
 | 
|
444  | 
fun the_identifiers (Data { identifiers , ... }) = identifiers;
 | 
|
445  | 
fun the_printings (Data { printings, ... }) = printings;
 | 
|
446  | 
||
447  | 
||
| 28060 | 448  | 
end; (*struct*)  |