| author | wenzelm |
| Mon, 09 Oct 2006 02:19:54 +0200 | |
| changeset 20902 | a0034e545c13 |
| parent 20896 | 1484c7af6d68 |
| child 20926 | b2f67b947200 |
| permissions | -rw-r--r-- |
|
18169
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
1 |
(* Title: Pure/Tools/codegen_serializer.ML |
|
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
2 |
ID: $Id$ |
|
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
3 |
Author: Florian Haftmann, TU Muenchen |
|
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
4 |
|
|
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
5 |
Serializer from intermediate language ("Thin-gol") to
|
| 20699 | 6 |
target languages (like SML or Haskell). |
|
18169
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
7 |
*) |
|
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
8 |
|
|
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
9 |
signature CODEGEN_SERIALIZER = |
|
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
10 |
sig |
| 20456 | 11 |
include BASIC_CODEGEN_THINGOL; |
| 20896 | 12 |
val get_serializer: theory -> string * Args.T list |
13 |
-> string list option -> CodegenThingol.code -> unit; |
|
| 20699 | 14 |
val eval_verbose: bool ref; |
15 |
val eval_term: theory -> |
|
| 20896 | 16 |
(string * 'a option ref) * CodegenThingol.iterm -> CodegenThingol.code |
| 20699 | 17 |
-> 'a; |
18 |
val const_has_serialization: theory -> string list -> string -> bool; |
|
19 |
val tyco_has_serialization: theory -> string list -> string -> bool; |
|
20 |
val add_syntax_class: |
|
21 |
string -> string -> string * (string * string) list -> theory -> theory; |
|
22 |
val add_syntax_inst: string -> (string * string) -> theory -> theory; |
|
23 |
val parse_syntax_tyco: (theory |
|
24 |
-> CodegenConsts.const list * (string * typ) list |
|
25 |
-> string |
|
26 |
-> CodegenNames.tyco |
|
27 |
-> typ list -> CodegenThingol.itype list) |
|
28 |
-> Symtab.key |
|
29 |
-> xstring |
|
30 |
-> OuterParse.token list |
|
31 |
-> (theory -> theory) * OuterParse.token list; |
|
32 |
val parse_syntax_const: (theory |
|
33 |
-> CodegenConsts.const list * (string * typ) list |
|
34 |
-> string |
|
35 |
-> CodegenNames.const |
|
36 |
-> term list -> CodegenThingol.iterm list) |
|
37 |
-> Symtab.key |
|
38 |
-> string |
|
39 |
-> OuterParse.token list |
|
40 |
-> (theory -> theory) * OuterParse.token list; |
|
41 |
val add_pretty_list: string -> string -> string -> (Pretty.T list -> Pretty.T) |
|
42 |
-> ((string -> string) * (string -> string)) option -> int * string |
|
43 |
-> theory -> theory; |
|
44 |
val add_pretty_ml_string: string -> string -> string -> string |
|
45 |
-> (string -> string) -> (string -> string) -> string -> theory -> theory; |
|
46 |
val add_undefined: string -> string -> string -> theory -> theory; |
|
| 20896 | 47 |
|
48 |
type fixity; |
|
49 |
type serializer; |
|
50 |
val add_serializer : string * serializer -> theory -> theory; |
|
51 |
val ml_from_thingol: serializer; |
|
52 |
val hs_from_thingol: serializer; |
|
|
18169
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
53 |
end; |
|
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
54 |
|
|
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
55 |
structure CodegenSerializer: CODEGEN_SERIALIZER = |
|
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
56 |
struct |
|
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
57 |
|
|
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
58 |
open BasicCodegenThingol; |
| 20845 | 59 |
val tracing = CodegenThingol.tracing; |
| 18850 | 60 |
|
| 20896 | 61 |
(** syntax **) |
62 |
||
63 |
(* basics *) |
|
| 18702 | 64 |
|
| 18216 | 65 |
datatype lrx = L | R | X; |
66 |
||
| 18516 | 67 |
datatype fixity = |
| 18216 | 68 |
BR |
69 |
| NOBR |
|
70 |
| INFX of (int * lrx); |
|
71 |
||
| 20699 | 72 |
type 'a pretty_syntax = int * (fixity -> (fixity -> 'a -> Pretty.T) |
| 18865 | 73 |
-> 'a list -> Pretty.T); |
| 18516 | 74 |
|
| 18216 | 75 |
fun eval_lrx L L = false |
76 |
| eval_lrx R R = false |
|
77 |
| eval_lrx _ _ = true; |
|
78 |
||
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
79 |
fun eval_fxy NOBR _ = false |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
80 |
| eval_fxy _ BR = true |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
81 |
| eval_fxy _ NOBR = false |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
82 |
| eval_fxy (INFX (pr, lr)) (INFX (pr_ctxt, lr_ctxt)) = |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
83 |
pr < pr_ctxt |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
84 |
orelse pr = pr_ctxt |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
85 |
andalso eval_lrx lr lr_ctxt |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
86 |
| eval_fxy _ (INFX _) = false; |
| 18216 | 87 |
|
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
88 |
fun gen_brackify _ [p] = p |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
89 |
| gen_brackify true (ps as _::_) = Pretty.enclose "(" ")" ps
|
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
90 |
| gen_brackify false (ps as _::_) = Pretty.block ps; |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
91 |
|
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
92 |
fun brackify fxy_ctxt ps = |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
93 |
gen_brackify (eval_fxy BR fxy_ctxt) (Pretty.breaks ps); |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
94 |
|
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
95 |
fun brackify_infix infx fxy_ctxt ps = |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
96 |
gen_brackify (eval_fxy (INFX infx) fxy_ctxt) (Pretty.breaks ps); |
| 18216 | 97 |
|
| 20896 | 98 |
fun mk_app mk_app' from_expr const_syntax fxy (const as (c, (_, ty)), es) = |
| 20191 | 99 |
case const_syntax c |
| 20896 | 100 |
of NONE => brackify fxy (mk_app' c es) |
| 20699 | 101 |
| SOME (i, pr) => |
102 |
let |
|
103 |
val k = if i < 0 then (length o fst o CodegenThingol.unfold_fun) ty else i |
|
104 |
in if k <= length es |
|
105 |
then case chop i es of (es1, es2) => |
|
|
19607
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
106 |
brackify fxy (pr fxy from_expr es1 :: map (from_expr BR) es2) |
| 20699 | 107 |
else from_expr fxy (CodegenThingol.eta_expand (const, es) i) |
108 |
end; |
|
109 |
||
| 20896 | 110 |
val first_upper = implode o nth_map 0 Symbol.to_ascii_upper o explode; |
| 20699 | 111 |
|
112 |
||
| 20896 | 113 |
(* user-defined syntax *) |
| 20699 | 114 |
|
| 20896 | 115 |
val str = setmp print_mode [] Pretty.str; |
| 20699 | 116 |
|
117 |
val (atomK, infixK, infixlK, infixrK) = |
|
118 |
("target_atom", "infix", "infixl", "infixr");
|
|
119 |
val _ = OuterSyntax.add_keywords [atomK, infixK, infixlK, infixrK]; |
|
120 |
||
121 |
datatype 'a mixfix = |
|
122 |
Arg of fixity |
|
123 |
| Pretty of Pretty.T |
|
124 |
| Quote of 'a; |
|
| 18865 | 125 |
|
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
126 |
fun fillin_mixfix fxy_this ms fxy_ctxt pr args = |
| 18702 | 127 |
let |
128 |
fun fillin [] [] = |
|
| 19008 | 129 |
[] |
| 18702 | 130 |
| fillin (Arg fxy :: ms) (a :: args) = |
131 |
pr fxy a :: fillin ms args |
|
132 |
| fillin (Pretty p :: ms) args = |
|
133 |
p :: fillin ms args |
|
134 |
| fillin (Quote q :: ms) args = |
|
| 19008 | 135 |
pr BR q :: fillin ms args |
136 |
| fillin [] _ = |
|
| 20389 | 137 |
error ("Inconsistent mixfix: too many arguments")
|
| 19008 | 138 |
| fillin _ [] = |
| 20389 | 139 |
error ("Inconsistent mixfix: too less arguments");
|
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
140 |
in gen_brackify (eval_fxy fxy_this fxy_ctxt) (fillin ms args) end; |
| 18702 | 141 |
|
142 |
fun parse_infix (fixity as INFX (i, x)) s = |
|
143 |
let |
|
144 |
val l = case x of L => fixity |
|
145 |
| _ => INFX (i, X); |
|
146 |
val r = case x of R => fixity |
|
147 |
| _ => INFX (i, X); |
|
148 |
in |
|
149 |
pair [Arg l, (Pretty o Pretty.brk) 1, (Pretty o str) s, (Pretty o Pretty.brk) 1, Arg r] |
|
150 |
end; |
|
151 |
||
152 |
fun parse_mixfix reader s ctxt = |
|
| 18335 | 153 |
let |
| 18702 | 154 |
fun sym s = Scan.lift ($$ s); |
155 |
fun lift_reader ctxt s = |
|
156 |
ctxt |
|
157 |
|> reader s |
|
158 |
|-> (fn x => pair (Quote x)); |
|
159 |
val sym_any = Scan.lift (Scan.one Symbol.not_eof); |
|
160 |
val parse = Scan.repeat ( |
|
161 |
(sym "_" -- sym "_" >> K (Arg NOBR)) |
|
162 |
|| (sym "_" >> K (Arg BR)) |
|
163 |
|| (sym "/" |-- Scan.repeat (sym " ") >> (Pretty o Pretty.brk o length)) |
|
164 |
|| Scan.depend (fn ctxt => $$ "{" |-- $$ "*" |-- Scan.repeat1
|
|
165 |
( $$ "'" |-- Scan.one Symbol.not_eof |
|
166 |
|| Scan.unless ($$ "*" -- $$ "}") (Scan.one Symbol.not_eof)) --| |
|
167 |
$$ "*" --| $$ "}" >> (implode #> lift_reader ctxt #> swap)) |
|
168 |
|| (Scan.repeat1 |
|
169 |
( sym "'" |-- sym_any |
|
170 |
|| Scan.unless (sym "_" || sym "?" || sym "/" || sym "{" |-- sym "*")
|
|
171 |
sym_any) >> (Pretty o str o implode))); |
|
172 |
in case Scan.finite' Symbol.stopper parse (ctxt, Symbol.explode s) |
|
173 |
of (p, (ctxt, [])) => (p, ctxt) |
|
174 |
| _ => error ("Malformed mixfix annotation: " ^ quote s)
|
|
175 |
end; |
|
176 |
||
| 20439 | 177 |
fun parse_syntax num_args reader = |
| 18335 | 178 |
let |
| 18702 | 179 |
fun is_arg (Arg _) = true |
180 |
| is_arg _ = false; |
|
| 20439 | 181 |
fun parse_nonatomic s ctxt = |
182 |
case parse_mixfix reader s ctxt |
|
183 |
of ([Pretty _], _) => |
|
184 |
error ("Mixfix contains just one pretty element; either declare as "
|
|
185 |
^ quote atomK ^ " or consider adding a break") |
|
186 |
| x => x; |
|
| 20456 | 187 |
val parse = ( |
| 20439 | 188 |
OuterParse.$$$ infixK |-- OuterParse.nat |
189 |
>> (fn i => (parse_infix (INFX (i, X)), INFX (i, X))) |
|
190 |
|| OuterParse.$$$ infixlK |-- OuterParse.nat |
|
191 |
>> (fn i => (parse_infix (INFX (i, L)), INFX (i, L))) |
|
192 |
|| OuterParse.$$$ infixrK |-- OuterParse.nat |
|
193 |
>> (fn i => (parse_infix (INFX (i, R)), INFX (i, R))) |
|
194 |
|| OuterParse.$$$ atomK |-- pair (parse_mixfix reader, NOBR) |
|
195 |
|| pair (parse_nonatomic, BR) |
|
| 20456 | 196 |
) -- OuterParse.string >> (fn ((p, fxy), s) => (p s, fxy)); |
|
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
197 |
fun mk fixity mfx ctxt = |
| 18702 | 198 |
let |
|
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
199 |
val i = (length o List.filter is_arg) mfx; |
| 20600 | 200 |
val _ = if i > num_args ctxt then error "Too many arguments in code syntax" else (); |
| 20699 | 201 |
in ((i, fillin_mixfix fixity mfx), ctxt) end; |
| 18702 | 202 |
in |
| 20439 | 203 |
parse |
|
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
204 |
#-> (fn (mfx_reader, fixity) => |
| 20699 | 205 |
pair (mfx_reader #-> (fn mfx => (mk fixity mfx))) |
| 18702 | 206 |
) |
207 |
end; |
|
208 |
||
209 |
||
| 20896 | 210 |
(* list and string serializer *) |
| 20401 | 211 |
|
212 |
fun implode_list c_nil c_cons e = |
|
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
213 |
let |
| 19202 | 214 |
fun dest_cons (IConst (c, _) `$ e1 `$ e2) = |
| 20401 | 215 |
if c = c_cons |
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
216 |
then SOME (e1, e2) |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
217 |
else NONE |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
218 |
| dest_cons _ = NONE; |
| 20401 | 219 |
val (es, e') = CodegenThingol.unfoldr dest_cons e; |
220 |
in case e' |
|
221 |
of IConst (c, _) => if c = c_nil then SOME es else NONE |
|
222 |
| _ => NONE |
|
223 |
end; |
|
224 |
||
225 |
fun implode_string mk_char mk_string es = |
|
226 |
if forall (fn IChar _ => true | _ => false) es |
|
227 |
then (SOME o str o mk_string o implode o map (fn IChar (c, _) => mk_char c)) es |
|
228 |
else NONE; |
|
229 |
||
230 |
fun pretty_ml_string c_nil c_cons mk_char mk_string target_implode = |
|
231 |
let |
|
232 |
fun pretty fxy pr [e] = |
|
233 |
case implode_list c_nil c_cons e |
|
234 |
of SOME es => (case implode_string mk_char mk_string es |
|
235 |
of SOME p => p |
|
236 |
| NONE => Pretty.block [str target_implode, Pretty.brk 1, pr BR e]) |
|
237 |
| NONE => Pretty.block [str target_implode, Pretty.brk 1, pr BR e] |
|
| 20699 | 238 |
in (1, pretty) end; |
| 20401 | 239 |
|
240 |
fun pretty_list c_nil c_cons mk_list mk_char_string (target_fxy, target_cons) = |
|
241 |
let |
|
242 |
fun default fxy pr e1 e2 = |
|
243 |
brackify_infix (target_fxy, R) fxy [ |
|
244 |
pr (INFX (target_fxy, X)) e1, |
|
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
245 |
str target_cons, |
| 20401 | 246 |
pr (INFX (target_fxy, R)) e2 |
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
247 |
]; |
| 20401 | 248 |
fun pretty fxy pr [e1, e2] = |
249 |
case Option.map (cons e1) (implode_list c_nil c_cons e2) |
|
250 |
of SOME es => |
|
251 |
(case mk_char_string |
|
252 |
of SOME (mk_char, mk_string) => |
|
253 |
(case implode_string mk_char mk_string es |
|
254 |
of SOME p => p |
|
255 |
| NONE => mk_list (map (pr NOBR) es)) |
|
256 |
| NONE => mk_list (map (pr NOBR) es)) |
|
257 |
| NONE => default fxy pr e1 e2; |
|
| 20699 | 258 |
in (2, pretty) end; |
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
259 |
|
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
260 |
|
| 20896 | 261 |
(* variable name contexts *) |
| 18216 | 262 |
|
| 20896 | 263 |
(*FIXME could name.ML do th whole job?*) |
264 |
fun make_vars names = (fold (fn name => Symtab.update_new (name, name)) names Symtab.empty, |
|
265 |
Name.make_context names); |
|
| 18216 | 266 |
|
| 20896 | 267 |
fun intro_vars names (namemap, namectxt) = |
268 |
let |
|
269 |
val (names', namectxt') = Name.variants names namectxt; |
|
270 |
val namemap' = fold2 (curry Symtab.update) names names' namemap; |
|
271 |
in (namemap', namectxt') end; |
|
| 19150 | 272 |
|
| 20896 | 273 |
fun lookup_var (namemap, _) name = case Symtab.lookup namemap name |
274 |
of SOME name' => name' |
|
275 |
| NONE => error ("invalid name in context: " ^ quote name);
|
|
276 |
||
277 |
fun constructive_fun (name, (eqs, ty)) = |
|
| 18216 | 278 |
let |
| 20699 | 279 |
val is_cons = CodegenNames.has_nsp CodegenNames.nsp_dtco; |
| 20896 | 280 |
fun is_pat (IConst (c, _)) = is_cons c |
281 |
| is_pat (IVar _) = true |
|
282 |
| is_pat (t1 `$ t2) = |
|
283 |
is_pat t1 andalso is_pat t2 |
|
284 |
| is_pat (INum _) = true |
|
285 |
| is_pat (IChar _) = true |
|
286 |
| is_pat _ = false; |
|
287 |
fun check_eq (eq as (lhs, rhs)) = |
|
288 |
if forall is_pat lhs |
|
289 |
then SOME eq |
|
290 |
else (warning ("In function " ^ quote name ^ ", throwing away one "
|
|
291 |
^ "non-executable function clause"); NONE) |
|
292 |
in case map_filter check_eq eqs |
|
293 |
of [] => error ("In function " ^ quote name ^ ", no "
|
|
294 |
^ "executable function clauses found") |
|
295 |
| eqs => (name, (eqs, ty)) |
|
296 |
end; |
|
297 |
||
298 |
||
299 |
(** SML serializer **) |
|
300 |
||
301 |
datatype ml_def = |
|
302 |
MLFuns of (string * ((iterm list * iterm) list * typscheme)) list |
|
303 |
| MLDatas of (string * ((vname * sort) list * (string * itype list) list)) list |
|
304 |
| MLClass of string * (class list * (vname * (string * itype) list)) |
|
305 |
| MLClassinst of string * ((class * (string * (vname * sort) list)) |
|
306 |
* ((class * (string * inst list list)) list |
|
307 |
* (string * iterm) list)); |
|
308 |
||
309 |
fun pr_sml_def tyco_syntax const_syntax keyword_vars deresolv ml_def = |
|
310 |
let |
|
311 |
val is_cons = CodegenNames.has_nsp CodegenNames.nsp_dtco; |
|
312 |
fun dictvar v = |
|
| 20699 | 313 |
first_upper v ^ "_"; |
| 20896 | 314 |
val label = translate_string (fn "." => "__" | c => c) |
315 |
o NameSpace.pack o op @ o apsnd single o apfst (fst o split_last) o split_last o NameSpace.unpack; |
|
316 |
fun pr_tyvar (v, []) = |
|
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
317 |
str "()" |
| 20896 | 318 |
| pr_tyvar (v, sort) = |
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
319 |
let |
| 20896 | 320 |
fun pr_class class = |
321 |
str ("'" ^ v ^ " " ^ deresolv class);
|
|
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
322 |
in |
|
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
323 |
Pretty.block [ |
|
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
324 |
str "(",
|
| 20896 | 325 |
(str o dictvar) v, |
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
326 |
str ":", |
|
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
327 |
case sort |
| 20896 | 328 |
of [class] => pr_class class |
329 |
| _ => Pretty.enum " *" "" "" (map pr_class sort), |
|
330 |
str ")" |
|
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
331 |
] |
|
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
332 |
end; |
| 20896 | 333 |
fun pr_insts fxy iys = |
| 18885 | 334 |
let |
| 20896 | 335 |
fun pr_proj s = str ("#" ^ s);
|
336 |
fun pr_lookup [] p = |
|
| 20401 | 337 |
p |
| 20896 | 338 |
| pr_lookup [p'] p = |
339 |
brackify BR [p', p] |
|
340 |
| pr_lookup (ps as _ :: _) p = |
|
341 |
brackify BR [Pretty.enum " o" "(" ")" ps, p];
|
|
342 |
fun pr_inst fxy (Instance (inst, iss)) = |
|
| 18885 | 343 |
brackify fxy ( |
| 20896 | 344 |
(str o deresolv) inst |
345 |
:: map (pr_insts BR) iss |
|
| 18885 | 346 |
) |
| 20896 | 347 |
| pr_inst fxy (Context (classes, (v, i))) = |
348 |
pr_lookup (map (pr_proj o label) classes |
|
349 |
@ (if i = ~1 then [] else [(pr_proj o string_of_int) (i+1)]) |
|
350 |
) ((str o dictvar) v); |
|
351 |
in case iys |
|
| 19253 | 352 |
of [] => str "()" |
| 20896 | 353 |
| [iy] => pr_inst fxy iy |
354 |
| _ :: _ => (Pretty.list "(" ")" o map (pr_inst NOBR)) iys
|
|
| 18885 | 355 |
end; |
| 20896 | 356 |
fun pr_tycoexpr fxy (tyco, tys) = |
| 18963 | 357 |
let |
| 20896 | 358 |
val tyco' = (str o deresolv) tyco |
359 |
in case map (pr_typ BR) tys |
|
| 18963 | 360 |
of [] => tyco' |
361 |
| [p] => Pretty.block [p, Pretty.brk 1, tyco'] |
|
362 |
| (ps as _::_) => Pretty.block [Pretty.list "(" ")" ps, Pretty.brk 1, tyco']
|
|
363 |
end |
|
| 20896 | 364 |
and pr_typ fxy (tyco `%% tys) = |
| 18702 | 365 |
(case tyco_syntax tyco |
| 20896 | 366 |
of NONE => pr_tycoexpr fxy (tyco, tys) |
| 20699 | 367 |
| SOME (i, pr) => |
368 |
if not (i = length tys) |
|
| 20389 | 369 |
then error ("Number of argument mismatch in customary serialization: "
|
| 18865 | 370 |
^ (string_of_int o length) tys ^ " given, " |
| 20699 | 371 |
^ string_of_int i ^ " expected") |
| 20896 | 372 |
else pr fxy pr_typ tys) |
373 |
| pr_typ fxy (t1 `-> t2) = |
|
374 |
(gen_brackify (case fxy of BR => false | _ => eval_fxy (INFX (1, R)) fxy) |
|
375 |
o Pretty.breaks) [ |
|
376 |
pr_typ (INFX (1, X)) t1, |
|
| 18702 | 377 |
str "->", |
| 20896 | 378 |
pr_typ (INFX (1, R)) t2 |
| 18702 | 379 |
] |
| 20896 | 380 |
| pr_typ fxy (ITyVar v) = |
| 18885 | 381 |
str ("'" ^ v);
|
| 20896 | 382 |
fun pr_term vars fxy (IConst c) = |
383 |
pr_app vars fxy (c, []) |
|
384 |
| pr_term vars fxy (IVar v) = |
|
385 |
str (lookup_var vars v) |
|
386 |
| pr_term vars fxy (t as t1 `$ t2) = |
|
387 |
(case CodegenThingol.unfold_const_app t |
|
388 |
of SOME c_ts => pr_app vars fxy c_ts |
|
| 18865 | 389 |
| NONE => |
| 20896 | 390 |
brackify fxy [pr_term vars NOBR t1, pr_term vars BR t2 ]) |
391 |
| pr_term vars fxy (t as _ `|-> _) = |
|
| 20105 | 392 |
let |
| 20896 | 393 |
val (ts, t') = CodegenThingol.unfold_abs t; |
394 |
val vs = fold (CodegenThingol.fold_varnames (insert (op =)) o fst) ts []; |
|
395 |
val vars' = intro_vars vs vars; |
|
396 |
fun mk_abs (t, ty) = (Pretty.block o Pretty.breaks) |
|
397 |
[str "fn", pr_term vars' NOBR t, str "=>"]; |
|
398 |
in brackify BR (map mk_abs ts @ [pr_term vars' NOBR t']) end |
|
399 |
| pr_term vars fxy (INum (n, _)) = |
|
400 |
brackify BR [(str o IntInf.toString) n, str ":", str "IntInf.int"] |
|
401 |
| pr_term vars _ (IChar (c, _)) = |
|
|
19607
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
402 |
(str o prefix "#" o quote) |
| 20203 | 403 |
(let val i = ord c |
| 20191 | 404 |
in if i < 32 |
| 20105 | 405 |
then prefix "\\" (string_of_int i) |
|
19607
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
406 |
else c |
|
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
407 |
end) |
| 20896 | 408 |
| pr_term vars fxy (t as ICase ((_, [_]), _)) = |
| 18216 | 409 |
let |
| 20896 | 410 |
val (ts, t) = CodegenThingol.unfold_let t; |
411 |
fun mk ((p, _), t) vars = |
|
| 20699 | 412 |
let |
| 20896 | 413 |
val vs = CodegenThingol.fold_varnames (insert (op =)) p []; |
414 |
val vars' = intro_vars vs vars; |
|
| 20699 | 415 |
in |
416 |
(Pretty.block [ |
|
417 |
(Pretty.block o Pretty.breaks) [ |
|
418 |
str "val", |
|
| 20896 | 419 |
pr_term vars' NOBR p, |
| 20699 | 420 |
str "=", |
| 20896 | 421 |
pr_term vars NOBR t |
| 20699 | 422 |
], |
423 |
str ";" |
|
| 20896 | 424 |
], vars') |
| 20699 | 425 |
end |
| 20896 | 426 |
val (binds, vars') = fold_map mk ts vars; |
427 |
in |
|
428 |
Pretty.chunks [ |
|
429 |
[str ("let"), Pretty.fbrk, binds |> Pretty.chunks] |> Pretty.block,
|
|
430 |
[str ("in"), Pretty.fbrk, pr_term vars' NOBR t] |> Pretty.block,
|
|
431 |
str ("end")
|
|
432 |
] end |
|
433 |
| pr_term vars fxy (ICase (((td, ty), b::bs), _)) = |
|
| 18216 | 434 |
let |
| 20896 | 435 |
fun pr definer (p, t) = |
| 20699 | 436 |
let |
| 20896 | 437 |
val vs = CodegenThingol.fold_varnames (insert (op =)) p []; |
438 |
val vars' = intro_vars vs vars; |
|
| 20699 | 439 |
in |
440 |
(Pretty.block o Pretty.breaks) [ |
|
441 |
str definer, |
|
| 20896 | 442 |
pr_term vars' NOBR p, |
| 20699 | 443 |
str "=>", |
| 20896 | 444 |
pr_term vars' NOBR t |
| 20699 | 445 |
] |
| 20896 | 446 |
end; |
447 |
in |
|
448 |
(Pretty.enclose "(" ")" o single o brackify fxy) (
|
|
449 |
str "case" |
|
450 |
:: pr_term vars NOBR td |
|
451 |
:: pr "of" b |
|
452 |
:: map (pr "|") bs |
|
453 |
) |
|
454 |
end |
|
455 |
and pr_app' vars c ts = |
|
456 |
let |
|
457 |
val p = (str o deresolv) c; |
|
458 |
val ps = map (pr_term vars BR) ts; |
|
459 |
in if is_cons c andalso length ts > 1 then |
|
460 |
[p, Pretty.enum "," "(" ")" ps]
|
|
| 18702 | 461 |
else |
| 20896 | 462 |
p :: ps |
463 |
end |
|
464 |
and pr_app vars fxy (app as ((c, (iss, ty)), ts)) = |
|
465 |
case if is_cons c then [] else (map (pr_insts BR) o filter_out null) iss |
|
| 18885 | 466 |
of [] => |
| 20896 | 467 |
mk_app (pr_app' vars) (pr_term vars) const_syntax fxy app |
468 |
| ps => |
|
| 20456 | 469 |
if (is_none o const_syntax) c then |
| 20896 | 470 |
brackify fxy ((str o deresolv) c :: (ps @ map (pr_term vars BR) ts)) |
471 |
else |
|
472 |
error ("Cannot apply user defined serialization for function expecting dictionaries: " ^ quote c)
|
|
473 |
fun eta_expand_poly_fun (funn as (_, (_::_, _))) = |
|
474 |
funn |
|
475 |
| eta_expand_poly_fun (funn as (_, ([_], ([], _)))) = |
|
476 |
funn |
|
477 |
| eta_expand_poly_fun (funn as (_, ([(_::_, _)], _))) = |
|
478 |
funn |
|
479 |
| eta_expand_poly_fun (funn as (_, ([(_, _ `|-> _)], _))) = |
|
480 |
funn |
|
481 |
| eta_expand_poly_fun (funn as (name, ([([], t)], tysm as (vs, ty)))) = |
|
482 |
if (null o fst o CodegenThingol.unfold_fun) ty |
|
483 |
orelse (not o null o filter_out (null o snd)) vs |
|
484 |
then funn |
|
485 |
else (name, ([([IVar "x"], t `$ IVar "x")], tysm)); |
|
486 |
fun pr_def (MLFuns raw_funns) = |
|
487 |
let |
|
488 |
val funns as (funn :: funns') = map (eta_expand_poly_fun o constructive_fun) raw_funns; |
|
489 |
val definer = |
|
490 |
let |
|
491 |
fun mk [] [] = "val" |
|
492 |
| mk (_::_) _ = "fun" |
|
493 |
| mk [] vs = if (null o filter_out (null o snd)) vs then "val" else "fun"; |
|
494 |
fun chk (_, ((ts, _) :: _, (vs, _))) NONE = SOME (mk ts vs) |
|
495 |
| chk (_, ((ts, _) :: _, (vs, _))) (SOME defi) = |
|
496 |
if defi = mk ts vs then SOME defi |
|
497 |
else error ("Mixing simultaneous vals and funs not implemented");
|
|
498 |
in the (fold chk funns NONE) end; |
|
499 |
fun pr_funn definer (name, (eqs as eq::eqs', (raw_vs, ty))) = |
|
500 |
let |
|
501 |
val vs = filter_out (null o snd) raw_vs; |
|
502 |
val shift = if null eqs' then I else |
|
503 |
map (Pretty.block o single o Pretty.block o single); |
|
504 |
fun pr_eq definer (ts, t) = |
|
505 |
let |
|
506 |
val consts = map_filter |
|
507 |
(fn c => if (is_some o const_syntax) c |
|
508 |
then NONE else (SOME o NameSpace.base o deresolv) c) |
|
509 |
((fold o CodegenThingol.fold_constnames) (insert (op =)) (t :: ts) []); |
|
510 |
val vars = keyword_vars |
|
511 |
|> intro_vars consts |
|
512 |
|> intro_vars ((fold o CodegenThingol.fold_unbound_varnames) (insert (op =)) (t :: ts) []); |
|
513 |
in |
|
514 |
(Pretty.block o Pretty.breaks) ( |
|
515 |
[str definer, (str o deresolv) name] |
|
516 |
@ (if null ts andalso null vs |
|
517 |
andalso not (ty = ITyVar "_")(*for evaluation*) |
|
518 |
then [str ":", pr_typ NOBR ty] |
|
519 |
else |
|
520 |
map pr_tyvar vs |
|
521 |
@ map (pr_term vars BR) ts) |
|
522 |
@ [str "=", pr_term vars NOBR t] |
|
523 |
) |
|
524 |
end |
|
525 |
in |
|
526 |
(Pretty.block o Pretty.fbreaks o shift) ( |
|
527 |
pr_eq definer eq |
|
528 |
:: map (pr_eq "|") eqs' |
|
529 |
) |
|
530 |
end; |
|
531 |
val (ps, p) = split_last (pr_funn definer funn :: map (pr_funn "and") funns'); |
|
532 |
in Pretty.chunks (ps @ [Pretty.block ([p, str ";"])]) end |
|
533 |
| pr_def (MLDatas (datas as (data :: datas'))) = |
|
534 |
let |
|
535 |
fun pr_co (co, []) = |
|
536 |
str (deresolv co) |
|
537 |
| pr_co (co, tys) = |
|
538 |
(Pretty.block o Pretty.breaks) [ |
|
539 |
str (deresolv co), |
|
540 |
str "of", |
|
541 |
Pretty.enum " *" "" "" (map (pr_typ (INFX (2, L))) tys) |
|
542 |
]; |
|
543 |
fun pr_data definer (tyco, (vs, cos)) = |
|
544 |
(Pretty.block o Pretty.breaks) ( |
|
545 |
str definer |
|
546 |
:: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs) |
|
547 |
:: str "=" |
|
548 |
:: separate (str "|") (map pr_co cos) |
|
549 |
); |
|
550 |
val (ps, p) = split_last (pr_data "datatype" data :: map (pr_data "and") datas'); |
|
551 |
in Pretty.chunks (ps @ [Pretty.block ([p, str ";"])]) end |
|
552 |
| pr_def (MLClass (class, (superclasses, (v, classops)))) = |
|
553 |
let |
|
554 |
val w = dictvar v; |
|
555 |
fun pr_superclass class = |
|
556 |
(Pretty.block o Pretty.breaks o map str) [ |
|
557 |
label class, ":", "'" ^ v, deresolv class |
|
558 |
]; |
|
559 |
fun pr_classop (classop, ty) = |
|
560 |
(Pretty.block o Pretty.breaks) [ |
|
561 |
(str o suffix "_" o NameSpace.base) classop, str ":", pr_typ NOBR ty |
|
562 |
]; |
|
563 |
fun pr_classop_fun (classop, _) = |
|
564 |
(Pretty.block o Pretty.breaks) [ |
|
565 |
str "fun", |
|
566 |
(str o deresolv) classop, |
|
567 |
Pretty.enclose "(" ")" [str (w ^ ":'" ^ v ^ " " ^ deresolv class)],
|
|
568 |
str "=", |
|
569 |
str ("#" ^ (suffix "_" o NameSpace.base) classop),
|
|
570 |
str (w ^ ";") |
|
571 |
]; |
|
572 |
in |
|
573 |
Pretty.chunks ( |
|
574 |
(Pretty.block o Pretty.breaks) [ |
|
575 |
str ("type '" ^ v),
|
|
576 |
(str o deresolv) class, |
|
577 |
str "=", |
|
578 |
Pretty.enum "," "{" "};" (
|
|
579 |
map pr_superclass superclasses @ map pr_classop classops |
|
580 |
) |
|
581 |
] |
|
582 |
:: map pr_classop_fun classops |
|
583 |
) |
|
584 |
end |
|
585 |
| pr_def (MLClassinst (inst, ((class, (tyco, arity)), (superarities, classop_defs)))) = |
|
586 |
let |
|
587 |
fun pr_superclass (superclass, superinst_iss) = |
|
588 |
(Pretty.block o Pretty.breaks) [ |
|
589 |
(str o label) superclass, |
|
590 |
str "=", |
|
591 |
pr_insts NOBR [Instance superinst_iss] |
|
592 |
]; |
|
593 |
fun pr_classop_def (classop, t) = |
|
594 |
let |
|
595 |
val consts = map_filter |
|
596 |
(fn c => if (is_some o const_syntax) c |
|
597 |
then NONE else (SOME o NameSpace.base o deresolv) c) |
|
598 |
(CodegenThingol.fold_constnames (insert (op =)) t []); |
|
599 |
val vars = keyword_vars |
|
600 |
|> intro_vars consts; |
|
601 |
in |
|
602 |
(Pretty.block o Pretty.breaks) [ |
|
603 |
(str o suffix "_" o NameSpace.base) classop, |
|
604 |
str "=", |
|
605 |
pr_term vars NOBR t |
|
606 |
] |
|
607 |
end; |
|
608 |
in |
|
609 |
(Pretty.block o Pretty.breaks) ([ |
|
610 |
str (if null arity then "val" else "fun"), |
|
611 |
(str o deresolv) inst ] @ |
|
612 |
map pr_tyvar arity @ [ |
|
613 |
str "=", |
|
614 |
Pretty.enum "," "{" "}" (map pr_superclass superarities @ map pr_classop_def classop_defs),
|
|
615 |
str ":", |
|
616 |
pr_tycoexpr NOBR (class, [tyco `%% map (ITyVar o fst) arity]) |
|
617 |
]) |
|
618 |
end; |
|
619 |
in pr_def ml_def end; |
|
|
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
620 |
|
| 20896 | 621 |
|
622 |
||
623 |
(** Haskell serializer **) |
|
624 |
||
625 |
fun pr_haskell class_syntax tyco_syntax const_syntax keyword_vars deresolv_here deresolv def = |
|
|
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
626 |
let |
| 20699 | 627 |
val is_cons = CodegenNames.has_nsp CodegenNames.nsp_dtco; |
| 20896 | 628 |
fun class_name class = case class_syntax class |
629 |
of NONE => deresolv class |
|
630 |
| SOME (class, _) => class; |
|
631 |
fun classop_name class classop = case class_syntax class |
|
632 |
of NONE => NameSpace.base classop |
|
633 |
| SOME (_, classop_syntax) => case classop_syntax classop |
|
634 |
of NONE => NameSpace.base classop |
|
635 |
| SOME classop => classop |
|
636 |
fun pr_typparms tyvars vs = |
|
637 |
case maps (fn (v, sort) => map (pair v) sort) vs |
|
638 |
of [] => str "" |
|
639 |
| xs => Pretty.block [ |
|
640 |
Pretty.enum "," "(" ")" (
|
|
641 |
map (fn (v, class) => str |
|
642 |
(class_name class ^ " " ^ lookup_var tyvars v)) xs |
|
643 |
), |
|
644 |
str " => " |
|
645 |
]; |
|
646 |
fun pr_tycoexpr tyvars fxy (tyco, tys) = |
|
647 |
brackify fxy (str tyco :: map (pr_typ tyvars BR) tys) |
|
648 |
and pr_typ tyvars fxy (tycoexpr as tyco `%% tys) = |
|
649 |
(case tyco_syntax tyco |
|
650 |
of NONE => |
|
651 |
pr_tycoexpr tyvars fxy (deresolv tyco, tys) |
|
652 |
| SOME (i, pr) => |
|
653 |
if not (i = length tys) |
|
654 |
then error ("Number of argument mismatch in customary serialization: "
|
|
655 |
^ (string_of_int o length) tys ^ " given, " |
|
656 |
^ string_of_int i ^ " expected") |
|
657 |
else pr fxy (pr_typ tyvars) tys) |
|
658 |
| pr_typ tyvars fxy (t1 `-> t2) = |
|
659 |
brackify_infix (1, R) fxy [ |
|
660 |
pr_typ tyvars (INFX (1, X)) t1, |
|
661 |
str "->", |
|
662 |
pr_typ tyvars (INFX (1, R)) t2 |
|
663 |
] |
|
664 |
| pr_typ tyvars fxy (ITyVar v) = |
|
665 |
(str o lookup_var tyvars) v; |
|
666 |
fun pr_typscheme_expr tyvars (vs, tycoexpr) = |
|
667 |
Pretty.block [pr_typparms tyvars vs, pr_tycoexpr tyvars NOBR tycoexpr]; |
|
668 |
fun pr_typscheme tyvars (vs, ty) = |
|
669 |
Pretty.block [pr_typparms tyvars vs, pr_typ tyvars NOBR ty]; |
|
670 |
fun pr_term vars fxy (IConst c) = |
|
671 |
pr_app vars fxy (c, []) |
|
672 |
| pr_term vars fxy (t as (t1 `$ t2)) = |
|
673 |
(case CodegenThingol.unfold_const_app t |
|
674 |
of SOME app => pr_app vars fxy app |
|
675 |
| _ => |
|
676 |
brackify fxy [ |
|
677 |
pr_term vars NOBR t1, |
|
678 |
pr_term vars BR t2 |
|
679 |
]) |
|
680 |
| pr_term vars fxy (IVar v) = |
|
681 |
(str o lookup_var vars) v |
|
682 |
| pr_term vars fxy (t as _ `|-> _) = |
|
|
19607
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
683 |
let |
| 20896 | 684 |
val (ts, t') = CodegenThingol.unfold_abs t; |
685 |
val vs = fold (CodegenThingol.fold_varnames (insert (op =)) o fst) ts []; |
|
686 |
val vars' = intro_vars vs vars; |
|
687 |
in |
|
688 |
brackify BR ( |
|
689 |
str "\\" |
|
690 |
:: map (pr_term vars' BR o fst) ts @ [ |
|
691 |
str "->", |
|
692 |
pr_term vars' NOBR t' |
|
693 |
]) |
|
694 |
end |
|
695 |
| pr_term vars fxy (INum (n, _)) = |
|
696 |
if n > 0 then |
|
697 |
(str o IntInf.toString) n |
|
698 |
else |
|
699 |
brackify BR [(str o Library.prefix "-" o IntInf.toString o IntInf.~) n] |
|
700 |
| pr_term vars fxy (IChar (c, _)) = |
|
701 |
(str o enclose "'" "'") |
|
702 |
(let val i = (Char.ord o the o Char.fromString) c |
|
703 |
in if i < 32 |
|
704 |
then Library.prefix "\\" (string_of_int i) |
|
705 |
else c |
|
706 |
end) |
|
707 |
| pr_term vars fxy (t as ICase ((_, [_]), _)) = |
|
| 18216 | 708 |
let |
| 20896 | 709 |
val (ts, t) = CodegenThingol.unfold_let t; |
710 |
fun pr ((p, _), t) vars = |
|
711 |
let |
|
712 |
val vs = CodegenThingol.fold_varnames (insert (op =)) p []; |
|
713 |
val vars' = intro_vars vs vars; |
|
714 |
in |
|
715 |
((Pretty.block o Pretty.breaks) [ |
|
716 |
pr_term vars' BR p, |
|
717 |
str "=", |
|
718 |
pr_term vars NOBR t |
|
719 |
], vars') |
|
720 |
end; |
|
721 |
val (binds, vars') = fold_map pr ts vars; |
|
722 |
in Pretty.chunks [ |
|
723 |
[str ("let"), Pretty.fbrk, binds |> Pretty.chunks] |> Pretty.block,
|
|
724 |
[str ("in "), pr_term vars' NOBR t] |> Pretty.block
|
|
725 |
] end |
|
726 |
| pr_term vars fxy (ICase (((td, _), bs), _)) = |
|
727 |
let |
|
728 |
fun pr (p, t) = |
|
729 |
let |
|
730 |
val vs = CodegenThingol.fold_varnames (insert (op =)) p []; |
|
731 |
val vars' = intro_vars vs vars; |
|
732 |
in |
|
733 |
(Pretty.block o Pretty.breaks) [ |
|
734 |
pr_term vars' NOBR p, |
|
735 |
str "->", |
|
736 |
pr_term vars' NOBR t |
|
737 |
] |
|
738 |
end |
|
739 |
in |
|
740 |
(Pretty.enclose "(" ")" o Pretty.breaks) [
|
|
741 |
str "case", |
|
742 |
pr_term vars NOBR td, |
|
743 |
str "of", |
|
744 |
(Pretty.chunks o map pr) bs |
|
745 |
] |
|
746 |
end |
|
747 |
and pr_app' vars c ts = |
|
748 |
(str o deresolv) c :: map (pr_term vars BR) ts |
|
749 |
and pr_app vars fxy = |
|
750 |
mk_app (pr_app' vars) (pr_term vars) const_syntax fxy; |
|
751 |
fun pr_def (name, CodegenThingol.Fun (funn as (eqs, (vs, ty)))) = |
|
752 |
let |
|
753 |
val tyvars = intro_vars (map fst vs) keyword_vars; |
|
754 |
fun pr_eq (ts, t) = |
|
| 20699 | 755 |
let |
756 |
val consts = map_filter |
|
757 |
(fn c => if (is_some o const_syntax) c |
|
| 20896 | 758 |
then NONE else (SOME o NameSpace.base o deresolv) c) |
759 |
((fold o CodegenThingol.fold_constnames) (insert (op =)) (t :: ts) []); |
|
760 |
val vars = keyword_vars |
|
761 |
|> intro_vars consts |
|
762 |
|> intro_vars ((fold o CodegenThingol.fold_unbound_varnames) (insert (op =)) (t :: ts) []); |
|
| 20699 | 763 |
in |
764 |
(Pretty.block o Pretty.breaks) ( |
|
| 20896 | 765 |
(str o deresolv_here) name |
766 |
:: map (pr_term vars BR) ts |
|
767 |
@ [str "=", pr_term vars NOBR t] |
|
| 20699 | 768 |
) |
| 20896 | 769 |
end; |
770 |
in |
|
771 |
Pretty.chunks ( |
|
772 |
Pretty.block [ |
|
773 |
(str o suffix " ::" o deresolv_here) name, |
|
774 |
Pretty.brk 1, |
|
775 |
pr_typscheme tyvars (vs, ty) |
|
776 |
] |
|
777 |
:: (map pr_eq o fst o snd o constructive_fun) (name, funn) |
|
778 |
) |
|
779 |
end |> SOME |
|
780 |
| pr_def (name, CodegenThingol.Datatype (vs, [(co, [ty])])) = |
|
781 |
let |
|
782 |
val tyvars = intro_vars (map fst vs) keyword_vars; |
|
| 18216 | 783 |
in |
| 20896 | 784 |
(Pretty.block o Pretty.breaks) [ |
785 |
str "newtype", |
|
786 |
pr_typscheme_expr tyvars (vs, (deresolv_here name, map (ITyVar o fst) vs)), |
|
787 |
str "=", |
|
788 |
(str o deresolv_here) co, |
|
789 |
pr_typ tyvars BR ty |
|
790 |
] |
|
791 |
end |> SOME |
|
792 |
| pr_def (name, CodegenThingol.Datatype (vs, co :: cos)) = |
|
793 |
let |
|
794 |
val tyvars = intro_vars (map fst vs) keyword_vars; |
|
795 |
fun pr_co (co, tys) = |
|
796 |
(Pretty.block o Pretty.breaks) ( |
|
797 |
(str o deresolv_here) co |
|
798 |
:: map (pr_typ tyvars BR) tys |
|
799 |
) |
|
800 |
in |
|
801 |
(Pretty.block o Pretty.breaks) ( |
|
802 |
str "data" |
|
803 |
:: pr_typscheme_expr tyvars (vs, (deresolv_here name, map (ITyVar o fst) vs)) |
|
804 |
:: str "=" |
|
805 |
:: pr_co co |
|
806 |
:: map ((fn p => Pretty.block [str "| ", p]) o pr_co) cos |
|
| 18912 | 807 |
) |
| 20896 | 808 |
end |> SOME |
809 |
| pr_def (_, CodegenThingol.Datatypecons _) = |
|
810 |
NONE |
|
811 |
| pr_def (name, CodegenThingol.Class (superclasss, (v, classops))) = |
|
812 |
let |
|
813 |
val tyvars = intro_vars [v] keyword_vars; |
|
814 |
fun pr_classop (classop, ty) = |
|
| 19136 | 815 |
Pretty.block [ |
| 20896 | 816 |
str (deresolv_here classop ^ " ::"), |
| 19136 | 817 |
Pretty.brk 1, |
| 20896 | 818 |
pr_typ tyvars NOBR ty |
| 19136 | 819 |
] |
| 20896 | 820 |
in |
821 |
Pretty.block [ |
|
822 |
str "class ", |
|
823 |
pr_typparms tyvars [(v, superclasss)], |
|
824 |
str (deresolv_here name ^ " " ^ v), |
|
825 |
str " where", |
|
826 |
Pretty.fbrk, |
|
827 |
Pretty.chunks (map pr_classop classops) |
|
828 |
] |
|
829 |
end |> SOME |
|
830 |
| pr_def (_, CodegenThingol.Classmember _) = |
|
831 |
NONE |
|
832 |
| pr_def (_, CodegenThingol.Classinst ((class, (tyco, vs)), (_, classop_defs))) = |
|
833 |
let |
|
834 |
val tyvars = intro_vars (map fst vs) keyword_vars; |
|
835 |
in |
|
836 |
Pretty.block [ |
|
837 |
str "instance ", |
|
838 |
pr_typparms tyvars vs, |
|
839 |
str (class_name class ^ " "), |
|
840 |
pr_typ tyvars BR (tyco `%% map (ITyVar o fst) vs), |
|
841 |
str " where", |
|
842 |
Pretty.fbrk, |
|
843 |
Pretty.chunks (map (fn (classop, t) => |
|
844 |
let |
|
845 |
val consts = map_filter |
|
846 |
(fn c => if (is_some o const_syntax) c |
|
847 |
then NONE else (SOME o NameSpace.base o deresolv) c) |
|
848 |
(CodegenThingol.fold_constnames (insert (op =)) t []); |
|
849 |
val vars = keyword_vars |
|
850 |
|> intro_vars consts; |
|
851 |
in |
|
852 |
(Pretty.block o Pretty.breaks) [ |
|
853 |
(str o classop_name class) classop, |
|
854 |
str "=", |
|
855 |
pr_term vars NOBR t |
|
856 |
] |
|
857 |
end |
|
858 |
) classop_defs) |
|
859 |
] |
|
860 |
end |> SOME |
|
861 |
in pr_def def end; |
|
862 |
||
863 |
||
864 |
(** generic abstract serializer **) |
|
865 |
||
866 |
structure NameMangler = NameManglerFun ( |
|
867 |
type ctxt = (string * string -> string) * (string -> string option); |
|
868 |
type src = string * string; |
|
869 |
val ord = prod_ord string_ord string_ord; |
|
870 |
fun mk (postprocess, validate) ((shallow, name), 0) = |
|
871 |
let |
|
872 |
val name' = postprocess (shallow, name); |
|
873 |
in case validate name' |
|
874 |
of NONE => name' |
|
875 |
| _ => mk (postprocess, validate) ((shallow, name), 1) |
|
876 |
end |
|
877 |
| mk (postprocess, validate) (("", name), i) =
|
|
878 |
postprocess ("", name ^ replicate_string i "'")
|
|
879 |
|> perhaps validate |
|
880 |
| mk (postprocess, validate) ((shallow, name), 1) = |
|
881 |
postprocess (shallow, shallow ^ "_" ^ name) |
|
882 |
|> perhaps validate |
|
883 |
| mk (postprocess, validate) ((shallow, name), i) = |
|
884 |
postprocess (shallow, name ^ replicate_string i "'") |
|
885 |
|> perhaps validate; |
|
886 |
fun is_valid _ _ = true; |
|
887 |
fun maybe_unique _ _ = NONE; |
|
888 |
fun re_mangle _ dst = error ("No such definition name: " ^ quote dst);
|
|
889 |
); |
|
890 |
||
891 |
(*FIXME refactor this properly*) |
|
892 |
fun code_serialize seri_defs seri_module validate postprocess nsp_conn name_root |
|
893 |
(code : CodegenThingol.code) = |
|
894 |
let |
|
895 |
datatype node = Def of CodegenThingol.def | Module of node Graph.T; |
|
896 |
fun dest_modl (Module m) = m; |
|
897 |
fun dest_name name = |
|
898 |
let |
|
899 |
val (names, name_base) = (split_last o NameSpace.unpack) name; |
|
900 |
val (names_mod, name_shallow) = split_last names; |
|
901 |
in (names_mod, NameSpace.pack [name_shallow, name_base]) end; |
|
902 |
fun mk_deresolver module nsp_conn postprocess validate = |
|
903 |
let |
|
904 |
datatype tabnode = N of string * tabnode Symtab.table option; |
|
905 |
fun mk module manglers tab = |
|
906 |
let |
|
907 |
fun mk_name name = |
|
908 |
case NameSpace.unpack name |
|
909 |
of [n] => ("", n)
|
|
910 |
| [s, n] => (s, n); |
|
911 |
fun in_conn (shallow, conn) = |
|
912 |
member (op = : string * string -> bool) conn shallow; |
|
913 |
fun add_name name = |
|
914 |
let |
|
915 |
val n as (shallow, _) = mk_name name; |
|
916 |
in |
|
917 |
AList.map_entry_yield in_conn shallow ( |
|
918 |
NameMangler.declare (postprocess, validate) n |
|
919 |
#-> (fn n' => pair (name, n')) |
|
920 |
) #> apfst the |
|
921 |
end; |
|
922 |
val (renamings, manglers') = |
|
923 |
fold_map add_name (Graph.keys module) manglers; |
|
924 |
fun extend_tab (n, n') = |
|
925 |
if (length o NameSpace.unpack) n = 1 |
|
926 |
then |
|
927 |
Symtab.update_new |
|
928 |
(n, N (n', SOME (mk ((dest_modl o Graph.get_node module) n) manglers' Symtab.empty))) |
|
929 |
else |
|
930 |
Symtab.update_new (n, N (n', NONE)); |
|
931 |
in fold extend_tab renamings tab end; |
|
932 |
fun get_path_name [] tab = |
|
933 |
([], SOME tab) |
|
934 |
| get_path_name [p] tab = |
|
935 |
let |
|
936 |
val SOME (N (p', tab')) = Symtab.lookup tab p |
|
937 |
in ([p'], tab') end |
|
938 |
| get_path_name [p1, p2] tab = |
|
939 |
(case Symtab.lookup tab p1 |
|
940 |
of SOME (N (p', SOME tab')) => |
|
941 |
let |
|
942 |
val (ps', tab'') = get_path_name [p2] tab' |
|
943 |
in (p' :: ps', tab'') end |
|
944 |
| NONE => |
|
945 |
let |
|
946 |
val SOME (N (p', NONE)) = Symtab.lookup tab (NameSpace.pack [p1, p2]) |
|
947 |
in ([p'], NONE) end) |
|
948 |
| get_path_name (p::ps) tab = |
|
949 |
let |
|
950 |
val SOME (N (p', SOME tab')) = Symtab.lookup tab p |
|
951 |
val (ps', tab'') = get_path_name ps tab' |
|
952 |
in (p' :: ps', tab'') end; |
|
953 |
fun deresolv tab prefix name = |
|
954 |
let |
|
955 |
val (common, (_, rem)) = chop_prefix (op =) (prefix, NameSpace.unpack name); |
|
956 |
val (_, SOME tab') = get_path_name common tab; |
|
957 |
val (name', _) = get_path_name rem tab'; |
|
958 |
in NameSpace.pack name' end handle BIND => (error ("Missing name: " ^ quote name ^ ", in " ^ quote (NameSpace.pack prefix)));
|
|
959 |
in deresolv (mk module (AList.make (K NameMangler.empty) nsp_conn) Symtab.empty) end; |
|
960 |
fun allimports_of modl = |
|
961 |
let |
|
962 |
fun imps_of prfx (Module modl) imps tab = |
|
963 |
let |
|
964 |
val this = NameSpace.pack prfx; |
|
965 |
val name_con = (rev o Graph.strong_conn) modl; |
|
966 |
in |
|
967 |
tab |
|
968 |
|> pair [] |
|
969 |
|> fold (fn names => fn (imps', tab) => |
|
970 |
tab |
|
971 |
|> fold_map (fn name => |
|
972 |
imps_of (prfx @ [name]) (Graph.get_node modl name) (imps' @ imps)) names |
|
973 |
|-> (fn imps'' => pair (flat imps'' @ imps'))) name_con |
|
974 |
|-> (fn imps' => |
|
975 |
Symtab.update_new (this, imps' @ imps) |
|
976 |
#> pair (this :: imps')) |
|
977 |
end |
|
978 |
| imps_of prfx (Def _) imps tab = |
|
979 |
([], tab); |
|
980 |
in snd (imps_of [] (Module modl) [] Symtab.empty) end; |
|
981 |
fun add_def ((names_mod, name_id), def) = |
|
982 |
let |
|
983 |
fun add [] = |
|
984 |
Graph.new_node (name_id, Def def) |
|
985 |
| add (m::ms) = |
|
986 |
Graph.default_node (m, Module Graph.empty) |
|
987 |
#> Graph.map_node m (Module o add ms o dest_modl) |
|
988 |
in add names_mod end; |
|
989 |
fun add_dep (name1, name2) = |
|
990 |
if name1 = name2 then I |
|
991 |
else |
|
992 |
let |
|
993 |
val m1 = dest_name name1 |> apsnd single |> (op @); |
|
994 |
val m2 = dest_name name2 |> apsnd single |> (op @); |
|
995 |
val (ms, (r1, r2)) = chop_prefix (op =) (m1, m2); |
|
996 |
val (ms, (s1::r1, s2::r2)) = chop_prefix (op =) (m1, m2); |
|
997 |
val add_edge = |
|
998 |
if null r1 andalso null r2 |
|
999 |
then Graph.add_edge |
|
1000 |
else fn edge => fn gr => (Graph.add_edge_acyclic edge gr |
|
1001 |
handle Graph.CYCLES _ => |
|
1002 |
error ("Module dependency "
|
|
1003 |
^ quote name1 ^ " -> " ^ quote name2 ^ " would result in module dependency cycle")) |
|
1004 |
fun add [] node = |
|
1005 |
node |
|
1006 |
|> add_edge (s1, s2) |
|
1007 |
| add (m::ms) node = |
|
1008 |
node |
|
1009 |
|> Graph.map_node m (Module o add ms o dest_modl); |
|
1010 |
in add ms end; |
|
1011 |
val root_module = |
|
1012 |
Graph.empty |
|
1013 |
|> Graph.fold (fn (name, (def, _)) => add_def (dest_name name, def)) code |
|
1014 |
|> Graph.fold (fn (name, (_, (_, deps))) => |
|
1015 |
fold (curry add_dep name) deps) code; |
|
1016 |
val names = map fst (Graph.dest root_module); |
|
1017 |
val imptab = allimports_of root_module; |
|
1018 |
val resolver = mk_deresolver root_module nsp_conn postprocess validate; |
|
1019 |
fun sresolver s = (resolver o NameSpace.unpack) s |
|
1020 |
fun mk_name prfx name = |
|
1021 |
let |
|
1022 |
val name_qual = NameSpace.pack (prfx @ [name]) |
|
1023 |
in (name_qual, resolver prfx name_qual) end; |
|
1024 |
fun is_bot (_, (Def Bot)) = true |
|
1025 |
| is_bot _ = false; |
|
1026 |
fun mk_contents prfx module = |
|
1027 |
map_filter (seri prfx) |
|
1028 |
((map (AList.make (Graph.get_node module)) o rev o Graph.strong_conn) module) |
|
1029 |
and seri prfx [(name, Module modl)] = |
|
1030 |
seri_module (resolver []) (map (resolver []) ((the o Symtab.lookup imptab) (NameSpace.pack (prfx @ [name])))) |
|
1031 |
(mk_name prfx name, mk_contents (prfx @ [name]) modl) |
|
1032 |
| seri prfx ds = |
|
1033 |
seri_defs sresolver (NameSpace.pack prfx) |
|
1034 |
(map (fn (name, Def def) => (fst (mk_name prfx name), def)) ds) |
|
1035 |
in |
|
1036 |
seri_module (resolver []) (map (resolver []) ((the o Symtab.lookup imptab) "")) |
|
1037 |
(("", name_root), (mk_contents [] root_module))
|
|
1038 |
end; |
|
1039 |
||
1040 |
fun abstract_serializer (target, nspgrp) name_root (from_defs, from_module, validator, postproc) |
|
1041 |
postprocess (class_syntax, tyco_syntax, const_syntax) |
|
1042 |
(drop, select) code = |
|
1043 |
let |
|
1044 |
fun project NONE code = code |
|
1045 |
| project (SOME names) code = |
|
1046 |
let |
|
1047 |
fun check name = if member (op =) drop name |
|
1048 |
then error ("shadowed definition " ^ quote name ^ " selected for serialization")
|
|
1049 |
else if can (Graph.get_node code) name |
|
1050 |
then () |
|
1051 |
else error ("dropped definition " ^ quote name ^ " selected for serialization")
|
|
1052 |
val names' = (map o tap) check names; |
|
1053 |
in CodegenThingol.project_code names code end; |
|
1054 |
fun from_module' resolv imps ((name_qual, name), defs) = |
|
1055 |
from_module resolv imps ((name_qual, name), defs) |
|
1056 |
|> postprocess (resolv name_qual); |
|
1057 |
in |
|
1058 |
code |
|
1059 |
|> tracing (fn _ => "dropping shadowed definitions...") |
|
1060 |
|> CodegenThingol.delete_garbage drop |
|
1061 |
|> tracing (fn _ => "projecting...") |
|
1062 |
|> project select |
|
1063 |
|> tracing (fn _ => "serializing...") |
|
1064 |
|> code_serialize (from_defs (class_syntax, tyco_syntax, const_syntax)) |
|
1065 |
from_module' validator postproc nspgrp name_root |
|
1066 |
|> K () |
|
1067 |
end; |
|
1068 |
||
1069 |
fun abstract_validator keywords name = |
|
1070 |
let |
|
1071 |
fun replace_invalid c = (*FIXME*) |
|
1072 |
if Symbol.is_ascii_letter c orelse Symbol.is_ascii_digit c orelse c = "'" |
|
1073 |
andalso not (NameSpace.separator = c) |
|
1074 |
then c |
|
1075 |
else "_" |
|
1076 |
fun suffix_it name= |
|
1077 |
name |
|
1078 |
|> member (op =) keywords ? suffix "'" |
|
1079 |
|> (fn name' => if name = name' then name else suffix_it name') |
|
1080 |
in |
|
1081 |
name |
|
1082 |
|> translate_string replace_invalid |
|
1083 |
|> suffix_it |
|
1084 |
|> (fn name' => if name = name' then NONE else SOME name') |
|
1085 |
end; |
|
1086 |
||
1087 |
fun write_file mkdir path p = ( |
|
1088 |
if mkdir |
|
1089 |
then |
|
1090 |
File.mkdir (Path.dir path) |
|
1091 |
else (); |
|
1092 |
File.write path (Pretty.output p ^ "\n"); |
|
1093 |
p |
|
1094 |
); |
|
1095 |
||
1096 |
fun mk_module_file postprocess_module ext path name p = |
|
1097 |
let |
|
1098 |
val prfx = Path.dir path; |
|
1099 |
val name' = case name |
|
1100 |
of "" => Path.base path |
|
1101 |
| _ => (Path.ext ext o Path.unpack o implode o separate "/" o NameSpace.unpack) name; |
|
1102 |
in |
|
1103 |
p |
|
1104 |
|> write_file true (Path.append prfx name') |
|
1105 |
|> postprocess_module name |
|
1106 |
end; |
|
1107 |
||
1108 |
fun parse_args f args = |
|
1109 |
case f args |
|
1110 |
of (x, []) => x |
|
1111 |
| (_, _) => error "bad serializer arguments"; |
|
1112 |
||
1113 |
fun parse_single_file serializer = |
|
1114 |
parse_args (Args.name |
|
1115 |
>> (fn path => serializer |
|
1116 |
(fn "" => write_file false (Path.unpack path) #> K NONE |
|
1117 |
| _ => SOME))); |
|
1118 |
||
1119 |
fun parse_multi_file postprocess_module ext serializer = |
|
1120 |
parse_args (Args.name |
|
1121 |
>> (fn path => (serializer o mk_module_file postprocess_module ext) (Path.unpack path))); |
|
1122 |
||
1123 |
fun parse_internal serializer = |
|
1124 |
parse_args (Args.name |
|
1125 |
>> (fn "-" => serializer |
|
1126 |
(fn "" => (fn p => (use_text Context.ml_output false (Pretty.output p); NONE)) |
|
1127 |
| _ => SOME) |
|
1128 |
| _ => Scan.fail ())); |
|
1129 |
||
1130 |
fun parse_stdout serializer = |
|
1131 |
parse_args (Args.name |
|
1132 |
>> (fn "_" => serializer |
|
1133 |
(fn "" => (fn p => (Pretty.writeln p; NONE)) |
|
1134 |
| _ => SOME) |
|
1135 |
| _ => Scan.fail ())); |
|
1136 |
||
1137 |
val nsp_module = CodegenNames.nsp_module; |
|
1138 |
val nsp_class = CodegenNames.nsp_class; |
|
1139 |
val nsp_tyco = CodegenNames.nsp_tyco; |
|
1140 |
val nsp_inst = CodegenNames.nsp_inst; |
|
1141 |
val nsp_fun = CodegenNames.nsp_fun; |
|
1142 |
val nsp_classop = CodegenNames.nsp_classop; |
|
1143 |
val nsp_dtco = CodegenNames.nsp_dtco; |
|
1144 |
val nsp_eval = CodegenNames.nsp_eval; |
|
1145 |
||
1146 |
||
1147 |
(** ML serializer **) |
|
1148 |
||
1149 |
local |
|
1150 |
||
1151 |
val reserved_ml' = [ |
|
1152 |
"bool", "int", "list", "unit", "option", "true", "false", "not", |
|
1153 |
"NONE", "SOME", "o", "string", "char", "String", "Term" |
|
1154 |
]; |
|
1155 |
||
1156 |
fun ml_from_defs (_, tyco_syntax, const_syntax) deresolver prefx defs = |
|
1157 |
let |
|
1158 |
val seri = pr_sml_def tyco_syntax const_syntax |
|
1159 |
(make_vars (ThmDatabase.ml_reserved @ reserved_ml')) |
|
1160 |
(deresolver prefx) #> SOME; |
|
1161 |
val filter_funs = |
|
1162 |
map |
|
1163 |
(fn (name, CodegenThingol.Fun info) => (name, info) |
|
1164 |
| (name, def) => error ("Function block containing illegal def: " ^ quote name)
|
|
|
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1165 |
) |
| 20896 | 1166 |
#> MLFuns; |
|
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1167 |
val filter_datatype = |
|
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19466
diff
changeset
|
1168 |
map_filter |
|
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
1169 |
(fn (name, CodegenThingol.Datatype info) => SOME (name, info) |
|
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
1170 |
| (name, CodegenThingol.Datatypecons _) => NONE |
| 20896 | 1171 |
| (name, def) => error ("Datatype block containing illegal def: " ^ quote name)
|
1172 |
) |
|
1173 |
#> MLDatas; |
|
| 20191 | 1174 |
fun filter_class defs = |
|
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19466
diff
changeset
|
1175 |
case map_filter |
|
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
1176 |
(fn (name, CodegenThingol.Class info) => SOME (name, info) |
|
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
1177 |
| (name, CodegenThingol.Classmember _) => NONE |
| 20896 | 1178 |
| (name, def) => error ("Class block containing illegal def: " ^ quote name)
|
1179 |
) defs |
|
1180 |
of [class] => MLClass class |
|
| 20389 | 1181 |
| _ => error ("Class block without class: " ^ (commas o map (quote o fst)) defs)
|
| 18850 | 1182 |
in case defs |
| 20896 | 1183 |
of (_, CodegenThingol.Fun _)::_ => (seri o filter_funs) defs |
1184 |
| (_, CodegenThingol.Datatypecons _)::_ => (seri o filter_datatype) defs |
|
1185 |
| (_, CodegenThingol.Datatype _)::_ => (seri o filter_datatype) defs |
|
1186 |
| (_, CodegenThingol.Class _)::_ => (seri o filter_class) defs |
|
1187 |
| (_, CodegenThingol.Classmember _)::_ => (seri o filter_class) defs |
|
1188 |
| [(inst, CodegenThingol.Classinst info)] => seri (MLClassinst (inst, info)) |
|
| 20389 | 1189 |
| defs => error ("Illegal mutual dependencies: " ^ (commas o map fst) defs)
|
| 18216 | 1190 |
end; |
1191 |
||
| 20699 | 1192 |
fun ml_serializer root_name target nspgrp = |
| 18216 | 1193 |
let |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
1194 |
fun ml_from_module resolv _ ((_, name), ps) = |
| 18756 | 1195 |
Pretty.chunks ([ |
| 18702 | 1196 |
str ("structure " ^ name ^ " = "),
|
1197 |
str "struct", |
|
1198 |
str "" |
|
1199 |
] @ separate (str "") ps @ [ |
|
1200 |
str "", |
|
1201 |
str ("end; (* struct " ^ name ^ " *)")
|
|
| 18756 | 1202 |
]); |
| 20401 | 1203 |
fun postproc (shallow, n) = |
1204 |
let |
|
1205 |
fun ch_first f = String.implode o nth_map 0 f o String.explode; |
|
| 20699 | 1206 |
in if shallow = CodegenNames.nsp_dtco |
| 20401 | 1207 |
then ch_first Char.toUpper n |
1208 |
else n |
|
1209 |
end; |
|
| 20191 | 1210 |
in abstract_serializer (target, nspgrp) |
| 20896 | 1211 |
root_name (ml_from_defs, ml_from_module, |
| 20699 | 1212 |
abstract_validator (ThmDatabase.ml_reserved @ reserved_ml'), postproc) end; |
| 20191 | 1213 |
|
1214 |
in |
|
1215 |
||
| 20896 | 1216 |
fun ml_from_thingol target args = |
| 20699 | 1217 |
let |
1218 |
val serializer = ml_serializer "ROOT" target [[nsp_module], [nsp_class, nsp_tyco], |
|
1219 |
[nsp_fun, nsp_dtco, nsp_class, nsp_classop, nsp_inst]] |
|
| 18850 | 1220 |
val parse_multi = |
| 20896 | 1221 |
Args.name |
| 20191 | 1222 |
#-> (fn "dir" => |
| 18850 | 1223 |
parse_multi_file |
1224 |
(K o SOME o str o suffix ";" o prefix "val _ = use " |
|
1225 |
o quote o suffix ".ML" o translate_string (fn "." => "/" | s => s)) "ML" serializer |
|
1226 |
| _ => Scan.fail ()); |
|
| 18282 | 1227 |
in |
| 20896 | 1228 |
(parse_multi |
| 19884 | 1229 |
|| parse_internal serializer |
| 20105 | 1230 |
|| parse_stdout serializer |
| 20896 | 1231 |
|| parse_single_file serializer) args |
| 18216 | 1232 |
end; |
1233 |
||
| 20389 | 1234 |
val eval_verbose = ref false; |
1235 |
||
| 20896 | 1236 |
fun eval_term_proto thy data hidden ((ref_name, reff), e) code = |
| 20191 | 1237 |
let |
| 20699 | 1238 |
val (val_name, code') = CodegenThingol.add_eval_def (nsp_eval, e) code; |
| 20191 | 1239 |
val struct_name = "EVAL"; |
| 20401 | 1240 |
fun output p = if !eval_verbose then (Pretty.writeln p; Pretty.output p) |
1241 |
else Pretty.output p; |
|
| 20699 | 1242 |
val serializer = ml_serializer struct_name "SML" [[nsp_module], [nsp_class, nsp_tyco], |
1243 |
[nsp_fun, nsp_dtco, nsp_class, nsp_classop, nsp_inst], [nsp_eval]] |
|
| 20401 | 1244 |
(fn "" => (fn p => (use_text Context.ml_output (!eval_verbose) (output p); NONE)) |
| 20896 | 1245 |
| _ => SOME) data |
1246 |
(hidden, SOME [NameSpace.pack [nsp_eval, val_name]]); |
|
| 20699 | 1247 |
val _ = serializer code'; |
| 20191 | 1248 |
val val_name_struct = NameSpace.append struct_name val_name; |
| 20600 | 1249 |
val _ = reff := NONE; |
1250 |
val _ = use_text Context.ml_output (!eval_verbose) ("val _ = (" ^ ref_name ^ " := SOME (" ^ val_name_struct ^ "))");
|
|
1251 |
in case !reff |
|
1252 |
of NONE => error ("Could not retrieve value of ML reference " ^ quote ref_name
|
|
1253 |
^ " (reference probably has been shadowed)") |
|
1254 |
| SOME value => value |
|
1255 |
end; |
|
| 20191 | 1256 |
|
| 20699 | 1257 |
structure NameMangler = NameManglerFun ( |
1258 |
type ctxt = string list; |
|
1259 |
type src = string; |
|
1260 |
val ord = string_ord; |
|
1261 |
fun mk reserved_ml (name, i) = |
|
1262 |
(Symbol.alphanum o NameSpace.base) name ^ replicate_string i "'"; |
|
1263 |
fun is_valid (reserved_ml : string list) = not o member (op =) reserved_ml; |
|
1264 |
fun maybe_unique _ _ = NONE; |
|
1265 |
fun re_mangle _ dst = error ("No such definition name: " ^ quote dst);
|
|
1266 |
); |
|
1267 |
||
| 19150 | 1268 |
fun mk_flat_ml_resolver names = |
1269 |
let |
|
1270 |
val mangler = |
|
1271 |
NameMangler.empty |
|
| 20699 | 1272 |
|> fold_map (NameMangler.declare (ThmDatabase.ml_reserved @ reserved_ml')) names |
| 19150 | 1273 |
|-> (fn _ => I) |
| 20699 | 1274 |
in NameMangler.get (ThmDatabase.ml_reserved @ reserved_ml') mangler end; |
|
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1275 |
|
| 18216 | 1276 |
end; (* local *) |
1277 |
||
| 20191 | 1278 |
|
1279 |
(** haskell serializer **) |
|
1280 |
||
| 20896 | 1281 |
fun hs_from_thingol target args = |
| 18282 | 1282 |
let |
| 20896 | 1283 |
fun hs_from_defs keyword_vars (class_syntax, tyco_syntax, const_syntax) deresolver prefix defs = |
| 18282 | 1284 |
let |
| 20896 | 1285 |
val deresolv = deresolver ""; |
1286 |
val deresolv_here = deresolver prefix; |
|
1287 |
val hs_from_def = pr_haskell class_syntax tyco_syntax const_syntax |
|
1288 |
keyword_vars deresolv_here deresolv; |
|
1289 |
in case map_filter hs_from_def defs |
|
1290 |
of [] => NONE |
|
1291 |
| ps => (SOME o Pretty.chunks o separate (str "")) ps |
|
| 18282 | 1292 |
end; |
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
1293 |
val reserved_hs = [ |
| 18702 | 1294 |
"hiding", "deriving", "where", "case", "of", "infix", "infixl", "infixr", |
1295 |
"import", "default", "forall", "let", "in", "class", "qualified", "data", |
|
1296 |
"newtype", "instance", "if", "then", "else", "type", "as", "do", "module" |
|
1297 |
] @ [ |
|
| 20401 | 1298 |
"Bool", "Integer", "Maybe", "True", "False", "Nothing", "Just", "not", "negate", |
1299 |
"String", "Char" |
|
| 18702 | 1300 |
]; |
|
19341
3414c04fbc39
added definitional code generator module: codegen_theorems.ML
haftmann
parents:
19280
diff
changeset
|
1301 |
fun hs_from_module resolv imps ((_, name), ps) = |
| 19038 | 1302 |
(Pretty.chunks) ( |
1303 |
str ("module " ^ name ^ " where")
|
|
1304 |
:: map (str o prefix "import qualified ") imps @ ( |
|
1305 |
str "" |
|
1306 |
:: separate (str "") ps |
|
1307 |
)); |
|
| 18919 | 1308 |
fun postproc (shallow, n) = |
1309 |
let |
|
1310 |
fun ch_first f = String.implode o nth_map 0 f o String.explode; |
|
| 20699 | 1311 |
in if member (op =) [nsp_module, nsp_class, nsp_tyco, nsp_dtco] shallow |
| 18919 | 1312 |
then ch_first Char.toUpper n |
1313 |
else ch_first Char.toLower n |
|
1314 |
end; |
|
| 20699 | 1315 |
val serializer = abstract_serializer (target, [[nsp_module], |
1316 |
[nsp_class], [nsp_tyco], [nsp_fun, nsp_classop], [nsp_dtco], [nsp_inst]]) |
|
| 20896 | 1317 |
"Main" (hs_from_defs (make_vars reserved_hs), hs_from_module, abstract_validator reserved_hs, postproc); |
| 18282 | 1318 |
in |
| 20896 | 1319 |
(parse_multi_file ((K o K) NONE) "hs" serializer) args |
| 18282 | 1320 |
end; |
1321 |
||
| 20896 | 1322 |
|
1323 |
(** LEGACY DIAGNOSIS **) |
|
1324 |
||
1325 |
local |
|
1326 |
||
1327 |
open CodegenThingol; |
|
1328 |
||
1329 |
in |
|
1330 |
||
1331 |
val pretty_typparms = |
|
1332 |
Pretty.list "(" ")" o Pretty.commas o map (fn (v, sort) => (Pretty.block o Pretty.breaks)
|
|
1333 |
[Pretty.str v, Pretty.str "::", Pretty.enum "&" "" "" (map Pretty.str sort)]); |
|
1334 |
||
1335 |
fun pretty_itype (tyco `%% tys) = |
|
1336 |
Pretty.enum "" "(" ")" (Pretty.str tyco :: map pretty_itype tys)
|
|
1337 |
| pretty_itype (ty1 `-> ty2) = |
|
1338 |
Pretty.enum "" "(" ")" [pretty_itype ty1, Pretty.str "->", pretty_itype ty2]
|
|
1339 |
| pretty_itype (ITyVar v) = |
|
1340 |
Pretty.str v; |
|
1341 |
||
1342 |
fun pretty_iterm (IConst (c, _)) = |
|
1343 |
Pretty.str c |
|
1344 |
| pretty_iterm (IVar v) = |
|
1345 |
Pretty.str ("?" ^ v)
|
|
1346 |
| pretty_iterm (t1 `$ t2) = |
|
1347 |
(Pretty.enclose "(" ")" o Pretty.breaks)
|
|
1348 |
[pretty_iterm t1, pretty_iterm t2] |
|
1349 |
| pretty_iterm ((v, ty) `|-> t) = |
|
1350 |
(Pretty.enclose "(" ")" o Pretty.breaks)
|
|
1351 |
[Pretty.str v, Pretty.str "::", pretty_itype ty, Pretty.str "|->", pretty_iterm t] |
|
1352 |
| pretty_iterm (INum (n, _)) = |
|
1353 |
(Pretty.str o IntInf.toString) n |
|
1354 |
| pretty_iterm (IChar (h, _)) = |
|
1355 |
(Pretty.str o quote) h |
|
1356 |
| pretty_iterm (ICase (((t, _), bs), _)) = |
|
1357 |
(Pretty.enclose "(" ")" o Pretty.breaks) [
|
|
1358 |
Pretty.str "case", |
|
1359 |
pretty_iterm t, |
|
1360 |
Pretty.enclose "(" ")" (map (fn (p, t) =>
|
|
1361 |
(Pretty.block o Pretty.breaks) [ |
|
1362 |
pretty_iterm p, |
|
1363 |
Pretty.str "=>", |
|
1364 |
pretty_iterm t |
|
1365 |
] |
|
1366 |
) bs) |
|
1367 |
]; |
|
1368 |
||
1369 |
fun pretty_def Bot = |
|
1370 |
Pretty.str "<Bot>" |
|
1371 |
| pretty_def (Fun (eqs, (vs, ty))) = |
|
1372 |
Pretty.enum " |" "" "" ( |
|
1373 |
map (fn (ps, body) => |
|
1374 |
Pretty.block [ |
|
1375 |
Pretty.enum "," "[" "]" (map pretty_iterm ps), |
|
1376 |
Pretty.str " |->", |
|
1377 |
Pretty.brk 1, |
|
1378 |
pretty_iterm body, |
|
1379 |
Pretty.str "::", |
|
1380 |
pretty_typparms vs, |
|
1381 |
Pretty.str "/", |
|
1382 |
pretty_itype ty |
|
1383 |
]) eqs |
|
1384 |
) |
|
1385 |
| pretty_def (Datatype (vs, cs)) = |
|
1386 |
Pretty.block [ |
|
1387 |
pretty_typparms vs, |
|
1388 |
Pretty.str " |=> ", |
|
1389 |
Pretty.enum " |" "" "" |
|
1390 |
(map (fn (c, tys) => (Pretty.block o Pretty.breaks) |
|
1391 |
(Pretty.str c :: map pretty_itype tys)) cs) |
|
1392 |
] |
|
1393 |
| pretty_def (Datatypecons dtname) = |
|
1394 |
Pretty.str ("cons " ^ dtname)
|
|
1395 |
| pretty_def (Class (supcls, (v, mems))) = |
|
1396 |
Pretty.block [ |
|
1397 |
Pretty.str ("class var " ^ v ^ " extending "),
|
|
1398 |
Pretty.enum "," "[" "]" (map Pretty.str supcls), |
|
1399 |
Pretty.str " with ", |
|
1400 |
Pretty.enum "," "[" "]" |
|
1401 |
(map (fn (m, ty) => Pretty.block |
|
1402 |
[Pretty.str (m ^ "::"), pretty_itype ty]) mems) |
|
1403 |
] |
|
1404 |
| pretty_def (Classmember clsname) = |
|
1405 |
Pretty.block [ |
|
1406 |
Pretty.str "class member belonging to ", |
|
1407 |
Pretty.str clsname |
|
1408 |
] |
|
1409 |
| pretty_def (Classinst ((clsname, (tyco, arity)), _)) = |
|
1410 |
Pretty.block [ |
|
1411 |
Pretty.str "class instance (",
|
|
1412 |
Pretty.str clsname, |
|
1413 |
Pretty.str ", (",
|
|
1414 |
Pretty.str tyco, |
|
1415 |
Pretty.str ", ", |
|
1416 |
Pretty.enum "," "[" "]" (map (Pretty.enum "," "{" "}" o
|
|
1417 |
map Pretty.str o snd) arity), |
|
1418 |
Pretty.str "))" |
|
1419 |
]; |
|
1420 |
||
1421 |
fun pretty_module code = |
|
1422 |
Pretty.chunks (map (fn (name, def) => Pretty.block |
|
1423 |
[Pretty.str name, Pretty.str " :=", Pretty.brk 1, pretty_def def]) |
|
1424 |
(AList.make (Graph.get_node code) (Graph.strong_conn code |> flat |> rev))); |
|
1425 |
||
1426 |
fun pretty_deps code = |
|
1427 |
let |
|
1428 |
fun one_node key = |
|
1429 |
let |
|
1430 |
val preds_ = Graph.imm_preds code key; |
|
1431 |
val succs_ = Graph.imm_succs code key; |
|
1432 |
val mutbs = gen_inter (op =) (preds_, succs_); |
|
1433 |
val preds = subtract (op =) mutbs preds_; |
|
1434 |
val succs = subtract (op =) mutbs succs_; |
|
1435 |
in |
|
1436 |
(Pretty.block o Pretty.fbreaks) ( |
|
1437 |
Pretty.str key |
|
1438 |
:: map (fn s => Pretty.str ("<-> " ^ s)) mutbs
|
|
1439 |
@ map (fn s => Pretty.str ("<-- " ^ s)) preds
|
|
1440 |
@ map (fn s => Pretty.str ("--> " ^ s)) succs
|
|
1441 |
) |
|
1442 |
end |
|
1443 |
in |
|
1444 |
code |
|
1445 |
|> Graph.strong_conn |
|
1446 |
|> flat |
|
1447 |
|> rev |
|
1448 |
|> map one_node |
|
1449 |
|> Pretty.chunks |
|
1450 |
end; |
|
1451 |
||
1452 |
end; (*local*) |
|
| 18282 | 1453 |
|
| 18702 | 1454 |
|
| 20699 | 1455 |
|
| 20896 | 1456 |
(** theory data **) |
1457 |
||
1458 |
datatype syntax_expr = SyntaxExpr of {
|
|
1459 |
class: ((string * (string -> string option)) * serial) Symtab.table, |
|
1460 |
inst: unit Symtab.table, |
|
1461 |
tyco: (itype pretty_syntax * serial) Symtab.table, |
|
1462 |
const: (iterm pretty_syntax * serial) Symtab.table |
|
1463 |
}; |
|
| 18702 | 1464 |
|
| 20896 | 1465 |
fun mk_syntax_expr ((class, inst), (tyco, const)) = |
1466 |
SyntaxExpr { class = class, inst = inst, tyco = tyco, const = const };
|
|
1467 |
fun map_syntax_expr f (SyntaxExpr { class, inst, tyco, const }) =
|
|
1468 |
mk_syntax_expr (f ((class, inst), (tyco, const))); |
|
1469 |
fun merge_syntax_expr (SyntaxExpr { class = class1, inst = inst1, tyco = tyco1, const = const1 },
|
|
1470 |
SyntaxExpr { class = class2, inst = inst2, tyco = tyco2, const = const2 }) =
|
|
1471 |
mk_syntax_expr ( |
|
1472 |
(Symtab.merge (eq_snd (op =)) (class1, class2), |
|
1473 |
Symtab.merge (op =) (inst1, inst2)), |
|
1474 |
(Symtab.merge (eq_snd (op =)) (tyco1, tyco2), |
|
1475 |
Symtab.merge (eq_snd (op =)) (const1, const2)) |
|
1476 |
); |
|
1477 |
||
1478 |
datatype syntax_modl = SyntaxModl of {
|
|
1479 |
merge: string Symtab.table, |
|
1480 |
prolog: Pretty.T Symtab.table |
|
1481 |
}; |
|
| 20699 | 1482 |
|
| 20896 | 1483 |
fun mk_syntax_modl (merge, prolog) = |
1484 |
SyntaxModl { merge = merge, prolog = prolog };
|
|
1485 |
fun map_syntax_modl f (SyntaxModl { merge, prolog }) =
|
|
1486 |
mk_syntax_modl (f (merge, prolog)); |
|
1487 |
fun merge_syntax_modl (SyntaxModl { merge = merge1, prolog = prolog1 },
|
|
1488 |
SyntaxModl { merge = merge2, prolog = prolog2 }) =
|
|
1489 |
mk_syntax_modl ( |
|
1490 |
Symtab.merge (op =) (merge1, merge2), |
|
1491 |
Symtab.merge (op =) (prolog1, prolog2) |
|
1492 |
); |
|
| 20699 | 1493 |
|
| 20896 | 1494 |
type serializer = string -> Args.T list |
1495 |
-> (string -> (string * (string -> string option)) option) |
|
1496 |
* (string |
|
1497 |
-> (int |
|
1498 |
* (fixity |
|
1499 |
-> (fixity |
|
1500 |
-> itype -> Pretty.T) |
|
1501 |
-> itype list -> Pretty.T)) |
|
1502 |
option) |
|
1503 |
* (string |
|
1504 |
-> (int |
|
1505 |
* (fixity |
|
1506 |
-> (fixity |
|
1507 |
-> iterm -> Pretty.T) |
|
1508 |
-> iterm list -> Pretty.T)) |
|
1509 |
option) |
|
1510 |
-> string list * string list option |
|
1511 |
-> CodegenThingol.code -> unit; |
|
1512 |
||
1513 |
datatype target = Target of {
|
|
1514 |
serial: serial, |
|
1515 |
serializer: serializer, |
|
1516 |
syntax_expr: syntax_expr, |
|
1517 |
syntax_modl: syntax_modl |
|
1518 |
}; |
|
| 20699 | 1519 |
|
| 20896 | 1520 |
fun mk_target (serial, (serializer, (syntax_expr, syntax_modl))) = |
1521 |
Target { serial = serial, serializer = serializer, syntax_expr = syntax_expr, syntax_modl = syntax_modl };
|
|
1522 |
fun map_target f ( Target { serial, serializer, syntax_expr, syntax_modl } ) =
|
|
1523 |
mk_target (f (serial, (serializer, (syntax_expr, syntax_modl)))); |
|
1524 |
fun merge_target target (Target { serial = serial1, serializer = serializer, syntax_expr = syntax_expr1, syntax_modl = syntax_modl1 },
|
|
1525 |
Target { serial = serial2, serializer = _, syntax_expr = syntax_expr2, syntax_modl = syntax_modl2 }) =
|
|
1526 |
if serial1 = serial2 then |
|
1527 |
mk_target (serial1, (serializer, |
|
1528 |
(merge_syntax_expr (syntax_expr1, syntax_expr2), |
|
1529 |
merge_syntax_modl (syntax_modl1, syntax_modl2)) |
|
1530 |
)) |
|
1531 |
else |
|
1532 |
error ("Incompatible serializers: " ^ quote target);
|
|
1533 |
||
1534 |
structure CodegenSerializerData = TheoryDataFun |
|
1535 |
(struct |
|
1536 |
val name = "Pure/codegen_serializer"; |
|
1537 |
type T = target Symtab.table; |
|
1538 |
val empty = Symtab.empty; |
|
1539 |
val copy = I; |
|
1540 |
val extend = I; |
|
1541 |
fun merge _ = Symtab.join merge_target; |
|
1542 |
fun print _ _ = (); |
|
1543 |
end); |
|
1544 |
||
1545 |
fun the_serializer (Target { serializer, ... }) = serializer;
|
|
1546 |
fun the_syntax_expr (Target { syntax_expr = SyntaxExpr x, ... }) = x;
|
|
1547 |
fun the_syntax_modl (Target { syntax_modl = SyntaxModl x, ... }) = x;
|
|
1548 |
||
1549 |
fun add_serializer (target, seri) thy = |
|
| 18702 | 1550 |
let |
| 20896 | 1551 |
val _ = case Symtab.lookup (CodegenSerializerData.get thy) target |
1552 |
of SOME _ => warning ("overwriting existing serializer " ^ quote target)
|
|
1553 |
| NONE => (); |
|
| 20699 | 1554 |
in |
| 20896 | 1555 |
thy |
1556 |
|> (CodegenSerializerData.map oo Symtab.map_default) |
|
1557 |
(target, mk_target (serial (), (seri, |
|
1558 |
(mk_syntax_expr ((Symtab.empty, Symtab.empty), (Symtab.empty, Symtab.empty)), |
|
1559 |
mk_syntax_modl (Symtab.empty, Symtab.empty))))) |
|
1560 |
(map_target (fn (serial, (_, syntax)) => (serial, (seri, syntax)))) |
|
| 20699 | 1561 |
end; |
1562 |
||
| 20896 | 1563 |
val _ = Context.add_setup ( |
1564 |
CodegenSerializerData.init |
|
1565 |
#> add_serializer ("SML", ml_from_thingol)
|
|
1566 |
#> add_serializer ("Haskell", hs_from_thingol)
|
|
1567 |
); |
|
1568 |
||
1569 |
fun get_serializer thy (target, args) cs = |
|
1570 |
let |
|
1571 |
val data = case Symtab.lookup (CodegenSerializerData.get thy) target |
|
1572 |
of SOME data => data |
|
1573 |
| NONE => error ("Unknown code target language: " ^ quote target);
|
|
1574 |
val seri = the_serializer data; |
|
1575 |
val { class, inst, tyco, const } = the_syntax_expr data;
|
|
1576 |
fun fun_of sys = (Option.map fst oo Symtab.lookup) sys; |
|
1577 |
in |
|
1578 |
seri target args (fun_of class, fun_of tyco, fun_of const) |
|
1579 |
(Symtab.keys class @ Symtab.keys inst @ Symtab.keys tyco @ Symtab.keys const, cs) |
|
1580 |
end; |
|
1581 |
||
1582 |
fun has_serialization f thy targets name = |
|
1583 |
forall ( |
|
1584 |
is_some o (fn tab => Symtab.lookup tab name) o f o the_syntax_expr o the |
|
1585 |
o (Symtab.lookup ((CodegenSerializerData.get) thy)) |
|
1586 |
) targets; |
|
1587 |
||
1588 |
val tyco_has_serialization = has_serialization #tyco; |
|
1589 |
val const_has_serialization = has_serialization #const; |
|
1590 |
||
1591 |
fun eval_term thy = |
|
1592 |
let |
|
1593 |
val target = "SML"; |
|
1594 |
val data = case Symtab.lookup (CodegenSerializerData.get thy) target |
|
1595 |
of SOME data => data |
|
1596 |
| NONE => error ("Unknown code target language: " ^ quote target);
|
|
1597 |
val { class, inst, tyco, const } = the_syntax_expr data;
|
|
1598 |
fun fun_of sys = (Option.map fst oo Symtab.lookup) sys; |
|
1599 |
in |
|
1600 |
eval_term_proto thy (fun_of class, fun_of tyco, fun_of const) |
|
1601 |
(Symtab.keys class @ Symtab.keys inst @ Symtab.keys tyco @ Symtab.keys const) |
|
1602 |
end; |
|
1603 |
||
1604 |
||
| 20699 | 1605 |
|
1606 |
(** target syntax interfaces **) |
|
1607 |
||
1608 |
local |
|
1609 |
||
| 20896 | 1610 |
fun map_syntax_exprs target f thy = |
1611 |
let |
|
1612 |
val _ = if is_some (Symtab.lookup (CodegenSerializerData.get thy) target) |
|
1613 |
then () |
|
1614 |
else error ("Unknown code target language: " ^ quote target);
|
|
1615 |
in |
|
1616 |
thy |
|
1617 |
|> (CodegenSerializerData.map o Symtab.map_entry target o map_target |
|
1618 |
o apsnd o apsnd o apfst o map_syntax_expr) f |
|
1619 |
end; |
|
1620 |
||
| 20699 | 1621 |
fun gen_add_syntax_class prep_class prep_const target raw_class (syntax, raw_ops) thy = |
1622 |
let |
|
1623 |
val cls = prep_class thy raw_class |
|
1624 |
val class = CodegenNames.class thy cls; |
|
| 20896 | 1625 |
fun mk_classop (const as (c, _)) = case AxClass.class_of_param thy c |
1626 |
of SOME class' => if cls = class' then CodegenNames.const thy const |
|
| 20699 | 1627 |
else error ("Not a class operation for class " ^ quote class ^ ": " ^ quote c)
|
1628 |
| NONE => error ("Not a class operation: " ^ quote c)
|
|
1629 |
val ops = (map o apfst) (mk_classop o prep_const thy) raw_ops; |
|
1630 |
val syntax_ops = AList.lookup (op =) ops; |
|
1631 |
in |
|
1632 |
thy |
|
| 20896 | 1633 |
|> (map_syntax_exprs target o apfst o apfst) |
1634 |
(Symtab.update (class, ((syntax, syntax_ops), serial ()))) |
|
| 20699 | 1635 |
end; |
1636 |
||
1637 |
fun gen_add_syntax_inst prep_class prep_tyco target (raw_tyco, raw_class) thy = |
|
1638 |
let |
|
1639 |
val inst = CodegenNames.instance thy (prep_class thy raw_class, prep_tyco thy raw_tyco); |
|
1640 |
in |
|
1641 |
thy |
|
| 20896 | 1642 |
|> (map_syntax_exprs target o apfst o apsnd) |
1643 |
(Symtab.update (inst, ())) |
|
| 20699 | 1644 |
end; |
1645 |
||
1646 |
fun gen_add_syntax_tyco prep_tyco raw_tyco target syntax thy = |
|
1647 |
let |
|
1648 |
val tyco = (CodegenNames.tyco thy o prep_tyco thy) raw_tyco; |
|
1649 |
in |
|
1650 |
thy |
|
| 20896 | 1651 |
|> (map_syntax_exprs target o apsnd o apfst) |
1652 |
(Symtab.update (tyco, (syntax, serial ()))) |
|
| 20699 | 1653 |
end; |
1654 |
||
1655 |
fun gen_add_syntax_const prep_const raw_c target syntax thy = |
|
1656 |
let |
|
1657 |
val c' = prep_const thy raw_c; |
|
1658 |
val c'' = CodegenNames.const thy c'; |
|
1659 |
in |
|
1660 |
thy |
|
| 20896 | 1661 |
|> (map_syntax_exprs target o apsnd o apsnd) |
1662 |
(Symtab.update (c'', (syntax, serial ()))) |
|
| 20699 | 1663 |
end; |
1664 |
||
1665 |
fun read_type thy raw_tyco = |
|
1666 |
let |
|
1667 |
val tyco = Sign.intern_type thy raw_tyco; |
|
1668 |
val _ = if Sign.declared_tyname thy tyco then () |
|
1669 |
else error ("No such type constructor: " ^ quote raw_tyco);
|
|
1670 |
in tyco end; |
|
1671 |
||
1672 |
fun idfs_of_const_names thy cs = |
|
1673 |
let |
|
1674 |
val cs' = AList.make (fn c => Sign.the_const_type thy c) cs; |
|
1675 |
val cs'' = map (CodegenConsts.norm_of_typ thy) cs'; |
|
1676 |
in AList.make (CodegenNames.const thy) cs'' end; |
|
1677 |
||
1678 |
fun read_quote reader consts_of target get_init gen raw_it thy = |
|
1679 |
let |
|
1680 |
val it = reader thy raw_it; |
|
1681 |
val cs = consts_of thy it; |
|
1682 |
in |
|
1683 |
gen thy cs target (get_init thy) [it] |
|
1684 |
|> (fn [it'] => (it', thy)) |
|
1685 |
end; |
|
1686 |
||
1687 |
fun parse_quote num_of reader consts_of target get_init gen adder = |
|
1688 |
parse_syntax num_of |
|
1689 |
(read_quote reader consts_of target get_init gen) |
|
1690 |
#-> (fn modifier => pair (modifier #-> adder target)); |
|
1691 |
||
1692 |
in |
|
1693 |
||
1694 |
val add_syntax_class = gen_add_syntax_class ClassPackage.read_class CodegenConsts.read_const; |
|
1695 |
val add_syntax_inst = gen_add_syntax_inst ClassPackage.read_class read_type; |
|
1696 |
||
1697 |
fun parse_syntax_tyco generate target raw_tyco = |
|
1698 |
let |
|
1699 |
fun intern thy = read_type thy raw_tyco; |
|
1700 |
fun num_of thy = Sign.arity_number thy (intern thy); |
|
1701 |
fun idf_of thy = CodegenNames.tyco thy (intern thy); |
|
1702 |
fun read_typ thy = |
|
1703 |
Sign.read_typ (thy, K NONE); |
|
1704 |
in |
|
1705 |
parse_quote num_of read_typ ((K o K) ([], [])) target idf_of generate |
|
1706 |
(gen_add_syntax_tyco read_type raw_tyco) |
|
1707 |
end; |
|
1708 |
||
1709 |
fun parse_syntax_const generate target raw_const = |
|
1710 |
let |
|
1711 |
fun intern thy = CodegenConsts.read_const thy raw_const; |
|
1712 |
fun num_of thy = (length o fst o strip_type o Sign.the_const_type thy o fst o intern) thy; |
|
1713 |
fun idf_of thy = (CodegenNames.const thy o intern) thy; |
|
1714 |
in |
|
1715 |
parse_quote num_of Sign.read_term CodegenConsts.consts_of target idf_of generate |
|
1716 |
(gen_add_syntax_const CodegenConsts.read_const raw_const) |
|
1717 |
end; |
|
1718 |
||
1719 |
fun add_pretty_list target nill cons mk_list mk_char_string target_cons thy = |
|
1720 |
let |
|
1721 |
val [(_, nil''), (cons', cons'')] = idfs_of_const_names thy [nill, cons]; |
|
1722 |
val pr = pretty_list nil'' cons'' mk_list mk_char_string target_cons; |
|
1723 |
in |
|
1724 |
thy |
|
1725 |
|> gen_add_syntax_const (K I) cons' target pr |
|
1726 |
end; |
|
1727 |
||
1728 |
fun add_pretty_ml_string target nill cons str mk_char mk_string target_implode thy = |
|
1729 |
let |
|
1730 |
val [(_, nil''), (_, cons''), (str', _)] = idfs_of_const_names thy [nill, cons, str]; |
|
1731 |
val pr = pretty_ml_string nil'' cons'' mk_char mk_string target_implode; |
|
1732 |
in |
|
1733 |
thy |
|
1734 |
|> gen_add_syntax_const (K I) str' target pr |
|
1735 |
end; |
|
1736 |
||
1737 |
fun add_undefined target undef target_undefined thy = |
|
1738 |
let |
|
1739 |
val [(undef', _)] = idfs_of_const_names thy [undef]; |
|
1740 |
fun pretty _ _ _ = str target_undefined; |
|
1741 |
in |
|
1742 |
thy |
|
1743 |
|> gen_add_syntax_const (K I) undef' target (~1, pretty) |
|
1744 |
end; |
|
1745 |
||
1746 |
end; (*local*) |
|
| 18702 | 1747 |
|
| 18216 | 1748 |
end; (* struct *) |