| author | aspinall |
| Sun, 07 Jan 2007 14:05:44 +0100 | |
| changeset 22028 | c13f6b5bf2b8 |
| parent 22022 | 93f842eb51a8 |
| child 22036 | 8919dbe67c90 |
| 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; |
| 20931 | 12 |
|
| 20699 | 13 |
val add_pretty_list: string -> string -> string -> (Pretty.T list -> Pretty.T) |
14 |
-> ((string -> string) * (string -> string)) option -> int * string |
|
15 |
-> theory -> theory; |
|
16 |
val add_pretty_ml_string: string -> string -> string -> string |
|
17 |
-> (string -> string) -> (string -> string) -> string -> theory -> theory; |
|
18 |
val add_undefined: string -> string -> string -> theory -> theory; |
|
| 21837 | 19 |
val add_pretty_imperative_monad_bind: string -> string -> theory -> theory; |
| 20896 | 20 |
|
21 |
type serializer; |
|
22 |
val add_serializer : string * serializer -> theory -> theory; |
|
| 21015 | 23 |
val get_serializer: theory -> string -> Args.T list |
| 20931 | 24 |
-> string list option -> CodegenThingol.code -> unit; |
| 21082 | 25 |
val assert_serializer: theory -> string -> string; |
| 20931 | 26 |
|
27 |
val const_has_serialization: theory -> string list -> string -> bool; |
|
28 |
val tyco_has_serialization: theory -> string list -> string -> bool; |
|
29 |
||
30 |
val eval_verbose: bool ref; |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
31 |
val eval_term: theory -> CodegenThingol.code |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
32 |
-> (string * 'a option ref) * CodegenThingol.iterm -> 'a; |
| 22022 | 33 |
val code_width: int ref; |
|
18169
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
34 |
end; |
|
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
35 |
|
|
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
36 |
structure CodegenSerializer: CODEGEN_SERIALIZER = |
|
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
37 |
struct |
|
45def66f86cb
added modules for code generator generation two, not operational yet
haftmann
parents:
diff
changeset
|
38 |
|
|
19167
f237c0cb3882
refined representation of codegen intermediate language
haftmann
parents:
19150
diff
changeset
|
39 |
open BasicCodegenThingol; |
| 20845 | 40 |
val tracing = CodegenThingol.tracing; |
| 18850 | 41 |
|
| 22022 | 42 |
(** basics **) |
| 18702 | 43 |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
44 |
infixr 5 @@; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
45 |
infixr 5 @|; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
46 |
fun x @@ y = [x, y]; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
47 |
fun xs @| y = xs @ [y]; |
| 22022 | 48 |
val str = setmp print_mode [] Pretty.str; |
49 |
val concat = Pretty.block o Pretty.breaks; |
|
50 |
fun semicolon ps = Pretty.block [concat ps, str ";"]; |
|
51 |
||
52 |
(** syntax **) |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
53 |
|
| 18216 | 54 |
datatype lrx = L | R | X; |
55 |
||
| 18516 | 56 |
datatype fixity = |
| 18216 | 57 |
BR |
58 |
| NOBR |
|
59 |
| INFX of (int * lrx); |
|
60 |
||
| 22022 | 61 |
val APP = INFX (~1, L); |
62 |
||
| 20699 | 63 |
type 'a pretty_syntax = int * (fixity -> (fixity -> 'a -> Pretty.T) |
| 18865 | 64 |
-> 'a list -> Pretty.T); |
| 18516 | 65 |
|
| 18216 | 66 |
fun eval_lrx L L = false |
67 |
| eval_lrx R R = false |
|
68 |
| eval_lrx _ _ = true; |
|
69 |
||
| 21882 | 70 |
fun eval_fxy NOBR NOBR = false |
71 |
| eval_fxy BR NOBR = false |
|
72 |
| eval_fxy NOBR BR = false |
|
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
73 |
| eval_fxy (INFX (pr, lr)) (INFX (pr_ctxt, lr_ctxt)) = |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
74 |
pr < pr_ctxt |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
75 |
orelse pr = pr_ctxt |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
76 |
andalso eval_lrx lr lr_ctxt |
| 22022 | 77 |
orelse pr_ctxt = ~1 |
| 21882 | 78 |
| eval_fxy _ (INFX _) = false |
| 22022 | 79 |
| eval_fxy (INFX _) NOBR = false |
| 21882 | 80 |
| eval_fxy _ _ = true; |
| 18216 | 81 |
|
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
82 |
fun gen_brackify _ [p] = p |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
83 |
| gen_brackify true (ps as _::_) = Pretty.enclose "(" ")" ps
|
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
84 |
| gen_brackify false (ps as _::_) = Pretty.block ps; |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
85 |
|
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
86 |
fun brackify fxy_ctxt ps = |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
87 |
gen_brackify (eval_fxy BR fxy_ctxt) (Pretty.breaks ps); |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
88 |
|
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
89 |
fun brackify_infix infx fxy_ctxt ps = |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
90 |
gen_brackify (eval_fxy (INFX infx) fxy_ctxt) (Pretty.breaks ps); |
| 18216 | 91 |
|
| 20976 | 92 |
fun mk_app mk_app' from_expr const_syntax fxy (app as ((c, (_, ty)), ts)) = |
| 20191 | 93 |
case const_syntax c |
| 20976 | 94 |
of NONE => brackify fxy (mk_app' app) |
| 20699 | 95 |
| SOME (i, pr) => |
96 |
let |
|
97 |
val k = if i < 0 then (length o fst o CodegenThingol.unfold_fun) ty else i |
|
| 21882 | 98 |
in if k = length ts |
99 |
then |
|
100 |
pr fxy from_expr ts |
|
101 |
else if k < length ts |
|
102 |
then case chop k ts of (ts1, ts2) => |
|
| 22022 | 103 |
brackify fxy (pr APP from_expr ts1 :: map (from_expr BR) ts2) |
| 20976 | 104 |
else from_expr fxy (CodegenThingol.eta_expand app k) |
| 20699 | 105 |
end; |
106 |
||
| 20896 | 107 |
val first_upper = implode o nth_map 0 Symbol.to_ascii_upper o explode; |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
108 |
val first_lower = implode o nth_map 0 Symbol.to_ascii_lower o explode; |
| 20699 | 109 |
|
110 |
||
| 20896 | 111 |
(* user-defined syntax *) |
| 20699 | 112 |
|
113 |
datatype 'a mixfix = |
|
114 |
Arg of fixity |
|
| 20931 | 115 |
| Pretty of Pretty.T; |
| 18865 | 116 |
|
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
117 |
fun mk_mixfix (fixity_this, mfx) = |
| 18702 | 118 |
let |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
119 |
fun is_arg (Arg _) = true |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
120 |
| is_arg _ = false; |
| 21882 | 121 |
val i = (length o filter is_arg) mfx; |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
122 |
fun fillin _ [] [] = |
| 19008 | 123 |
[] |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
124 |
| fillin pr (Arg fxy :: mfx) (a :: args) = |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
125 |
pr fxy a :: fillin pr mfx args |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
126 |
| fillin pr (Pretty p :: mfx) args = |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
127 |
p :: fillin pr mfx args |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
128 |
| fillin _ [] _ = |
| 20389 | 129 |
error ("Inconsistent mixfix: too many arguments")
|
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
130 |
| fillin _ _ [] = |
| 20389 | 131 |
error ("Inconsistent mixfix: too less arguments");
|
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
132 |
in |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
133 |
(i, fn fixity_ctxt => fn pr => fn args => |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
134 |
gen_brackify (eval_fxy fixity_this fixity_ctxt) (fillin pr mfx args)) |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
135 |
end; |
| 18702 | 136 |
|
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
137 |
fun parse_infix (x, i) s = |
| 18702 | 138 |
let |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
139 |
val l = case x of L => INFX (i, L) | _ => INFX (i, X); |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
140 |
val r = case x of R => INFX (i, R) | _ => INFX (i, X); |
| 18702 | 141 |
in |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
142 |
mk_mixfix (INFX (i, x), [Arg l, (Pretty o Pretty.brk) 1, (Pretty o str) s, (Pretty o Pretty.brk) 1, Arg r]) |
| 18702 | 143 |
end; |
144 |
||
| 20931 | 145 |
fun parse_mixfix s = |
| 18335 | 146 |
let |
| 20931 | 147 |
val sym_any = Scan.one Symbol.not_eof; |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
148 |
val parse = Scan.optional ($$ "!" >> K true) false -- Scan.repeat ( |
| 21130 | 149 |
($$ "(" -- $$ "_" -- $$ ")" >> K (Arg NOBR))
|
| 20931 | 150 |
|| ($$ "_" >> K (Arg BR)) |
151 |
|| ($$ "/" |-- Scan.repeat ($$ " ") >> (Pretty o Pretty.brk o length)) |
|
| 18702 | 152 |
|| (Scan.repeat1 |
| 20931 | 153 |
( $$ "'" |-- sym_any |
| 21130 | 154 |
|| Scan.unless ($$ "_" || $$ "/" || $$ "(" |-- $$ "_" |-- $$ ")")
|
| 18702 | 155 |
sym_any) >> (Pretty o str o implode))); |
| 20931 | 156 |
in case Scan.finite Symbol.stopper parse (Symbol.explode s) |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
157 |
of ((_, p as [_]), []) => mk_mixfix (NOBR, p) |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
158 |
| ((b, p as _ :: _ :: _), []) => mk_mixfix (if b then NOBR else BR, p) |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
159 |
| _ => Scan.!! (the_default ("malformed mixfix annotation: " ^ quote s) o snd) Scan.fail ()
|
| 18702 | 160 |
end; |
161 |
||
| 21082 | 162 |
fun parse_args f args = |
| 21895 | 163 |
case Scan.read Args.stopper f args |
164 |
of SOME x => x |
|
165 |
| NONE => error "bad serializer arguments"; |
|
| 21082 | 166 |
|
| 18702 | 167 |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
168 |
(* module names *) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
169 |
|
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
170 |
val dest_name = |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
171 |
apfst NameSpace.implode o split_last o fst o split_last o NameSpace.explode; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
172 |
|
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
173 |
fun mk_modl_name_tab empty_names prefix module_alias code = |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
174 |
let |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
175 |
fun nsp_map f = NameSpace.explode #> map f #> NameSpace.implode; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
176 |
fun mk_alias name = |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
177 |
case module_alias name |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
178 |
of SOME name' => name' |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
179 |
| NONE => nsp_map (fn name => (the_single o fst) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
180 |
(Name.variants [name] empty_names)) name; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
181 |
fun mk_prefix name = |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
182 |
case prefix |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
183 |
of SOME prefix => NameSpace.append prefix name |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
184 |
| NONE => name; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
185 |
val tab = |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
186 |
Symtab.empty |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
187 |
|> Graph.fold ((fn name => Symtab.default (name, (mk_alias #> mk_prefix) name)) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
188 |
o fst o dest_name o fst) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
189 |
code |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
190 |
in (fn name => (the o Symtab.lookup tab) name) end; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
191 |
|
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
192 |
|
| 20896 | 193 |
(* list and string serializer *) |
| 20401 | 194 |
|
195 |
fun implode_list c_nil c_cons e = |
|
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
196 |
let |
| 19202 | 197 |
fun dest_cons (IConst (c, _) `$ e1 `$ e2) = |
| 20401 | 198 |
if c = c_cons |
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
199 |
then SOME (e1, e2) |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
200 |
else NONE |
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
201 |
| dest_cons _ = NONE; |
| 20401 | 202 |
val (es, e') = CodegenThingol.unfoldr dest_cons e; |
203 |
in case e' |
|
204 |
of IConst (c, _) => if c = c_nil then SOME es else NONE |
|
205 |
| _ => NONE |
|
206 |
end; |
|
207 |
||
208 |
fun implode_string mk_char mk_string es = |
|
209 |
if forall (fn IChar _ => true | _ => false) es |
|
| 21015 | 210 |
then (SOME o str o mk_string o implode o map (fn IChar c => mk_char c)) es |
| 20401 | 211 |
else NONE; |
212 |
||
213 |
fun pretty_ml_string c_nil c_cons mk_char mk_string target_implode = |
|
214 |
let |
|
| 21837 | 215 |
fun pretty fxy pr [t] = |
216 |
case implode_list c_nil c_cons t |
|
217 |
of SOME ts => (case implode_string mk_char mk_string ts |
|
| 20401 | 218 |
of SOME p => p |
| 21837 | 219 |
| NONE => Pretty.block [str target_implode, Pretty.brk 1, pr BR t]) |
220 |
| NONE => Pretty.block [str target_implode, Pretty.brk 1, pr BR t] |
|
| 20699 | 221 |
in (1, pretty) end; |
| 20401 | 222 |
|
223 |
fun pretty_list c_nil c_cons mk_list mk_char_string (target_fxy, target_cons) = |
|
224 |
let |
|
| 21837 | 225 |
fun default fxy pr t1 t2 = |
| 20401 | 226 |
brackify_infix (target_fxy, R) fxy [ |
| 21837 | 227 |
pr (INFX (target_fxy, X)) t1, |
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
228 |
str target_cons, |
| 21837 | 229 |
pr (INFX (target_fxy, R)) t2 |
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
230 |
]; |
| 21837 | 231 |
fun pretty fxy pr [t1, t2] = |
232 |
case Option.map (cons t1) (implode_list c_nil c_cons t2) |
|
233 |
of SOME ts => |
|
| 20401 | 234 |
(case mk_char_string |
235 |
of SOME (mk_char, mk_string) => |
|
| 21837 | 236 |
(case implode_string mk_char mk_string ts |
| 20401 | 237 |
of SOME p => p |
| 21837 | 238 |
| NONE => mk_list (map (pr NOBR) ts)) |
239 |
| NONE => mk_list (map (pr NOBR) ts)) |
|
240 |
| NONE => default fxy pr t1 t2; |
|
241 |
in (2, pretty) end; |
|
242 |
||
243 |
fun pretty_imperative_monad_bind c_bind = |
|
244 |
let |
|
245 |
fun pretty fxy pr [t1, (v, ty) `|-> t2] = |
|
246 |
pr fxy (ICase ((t1, ty), ([(IVar v, t2)]))) |
|
247 |
| pretty _ _ _ = error "bad arguments for imperative monad bind"; |
|
| 20699 | 248 |
in (2, pretty) end; |
|
18704
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
249 |
|
|
2c86ced392a8
substantial improvement in serialization handling
haftmann
parents:
18702
diff
changeset
|
250 |
|
| 21015 | 251 |
|
| 22022 | 252 |
(** SML/OCaml serializer **) |
| 20896 | 253 |
|
254 |
datatype ml_def = |
|
255 |
MLFuns of (string * ((iterm list * iterm) list * typscheme)) list |
|
256 |
| MLDatas of (string * ((vname * sort) list * (string * itype list) list)) list |
|
257 |
| MLClass of string * (class list * (vname * (string * itype) list)) |
|
258 |
| MLClassinst of string * ((class * (string * (vname * sort) list)) |
|
259 |
* ((class * (string * inst list list)) list |
|
260 |
* (string * iterm) list)); |
|
261 |
||
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
262 |
fun pr_sml tyco_syntax const_syntax keyword_vars deresolv is_cons ml_def = |
| 20896 | 263 |
let |
264 |
fun dictvar v = |
|
| 20699 | 265 |
first_upper v ^ "_"; |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
266 |
val label = translate_string (fn "." => "__" | c => c) o NameSpace.qualifier; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
267 |
val mk_classop_name = suffix "_" o snd o dest_name; |
| 20896 | 268 |
fun pr_tyvar (v, []) = |
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
269 |
str "()" |
| 20896 | 270 |
| pr_tyvar (v, sort) = |
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
271 |
let |
| 20896 | 272 |
fun pr_class class = |
273 |
str ("'" ^ v ^ " " ^ deresolv class);
|
|
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
274 |
in |
|
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
275 |
Pretty.block [ |
|
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
276 |
str "(",
|
| 20896 | 277 |
(str o dictvar) v, |
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
278 |
str ":", |
|
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
279 |
case sort |
| 20896 | 280 |
of [class] => pr_class class |
281 |
| _ => Pretty.enum " *" "" "" (map pr_class sort), |
|
282 |
str ")" |
|
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
283 |
] |
|
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19936
diff
changeset
|
284 |
end; |
| 20896 | 285 |
fun pr_insts fxy iys = |
| 18885 | 286 |
let |
| 20896 | 287 |
fun pr_proj s = str ("#" ^ s);
|
288 |
fun pr_lookup [] p = |
|
| 20401 | 289 |
p |
| 20896 | 290 |
| pr_lookup [p'] p = |
291 |
brackify BR [p', p] |
|
292 |
| pr_lookup (ps as _ :: _) p = |
|
293 |
brackify BR [Pretty.enum " o" "(" ")" ps, p];
|
|
294 |
fun pr_inst fxy (Instance (inst, iss)) = |
|
| 18885 | 295 |
brackify fxy ( |
| 20896 | 296 |
(str o deresolv) inst |
297 |
:: map (pr_insts BR) iss |
|
| 18885 | 298 |
) |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
299 |
| pr_inst fxy (Context ((classes, i), (v, k))) = |
| 20896 | 300 |
pr_lookup (map (pr_proj o label) classes |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
301 |
@ (if k = 1 then [] else [(pr_proj o string_of_int) (i+1)]) |
| 20896 | 302 |
) ((str o dictvar) v); |
303 |
in case iys |
|
| 19253 | 304 |
of [] => str "()" |
| 20896 | 305 |
| [iy] => pr_inst fxy iy |
306 |
| _ :: _ => (Pretty.list "(" ")" o map (pr_inst NOBR)) iys
|
|
| 18885 | 307 |
end; |
| 20896 | 308 |
fun pr_tycoexpr fxy (tyco, tys) = |
| 18963 | 309 |
let |
| 20896 | 310 |
val tyco' = (str o deresolv) tyco |
311 |
in case map (pr_typ BR) tys |
|
| 18963 | 312 |
of [] => tyco' |
313 |
| [p] => Pretty.block [p, Pretty.brk 1, tyco'] |
|
314 |
| (ps as _::_) => Pretty.block [Pretty.list "(" ")" ps, Pretty.brk 1, tyco']
|
|
315 |
end |
|
| 20896 | 316 |
and pr_typ fxy (tyco `%% tys) = |
| 18702 | 317 |
(case tyco_syntax tyco |
| 20896 | 318 |
of NONE => pr_tycoexpr fxy (tyco, tys) |
| 20699 | 319 |
| SOME (i, pr) => |
320 |
if not (i = length tys) |
|
| 20389 | 321 |
then error ("Number of argument mismatch in customary serialization: "
|
| 18865 | 322 |
^ (string_of_int o length) tys ^ " given, " |
| 20699 | 323 |
^ string_of_int i ^ " expected") |
| 20896 | 324 |
else pr fxy pr_typ tys) |
325 |
| pr_typ fxy (ITyVar v) = |
|
| 18885 | 326 |
str ("'" ^ v);
|
| 20896 | 327 |
fun pr_term vars fxy (IConst c) = |
328 |
pr_app vars fxy (c, []) |
|
329 |
| pr_term vars fxy (IVar v) = |
|
| 21082 | 330 |
str (CodegenThingol.lookup_var vars v) |
| 20896 | 331 |
| pr_term vars fxy (t as t1 `$ t2) = |
332 |
(case CodegenThingol.unfold_const_app t |
|
333 |
of SOME c_ts => pr_app vars fxy c_ts |
|
| 18865 | 334 |
| NONE => |
| 20976 | 335 |
brackify fxy [pr_term vars NOBR t1, pr_term vars BR t2]) |
| 20896 | 336 |
| pr_term vars fxy (t as _ `|-> _) = |
| 20105 | 337 |
let |
| 21015 | 338 |
val (ps, t') = CodegenThingol.unfold_abs t; |
339 |
fun pr ((v, NONE), _) vars = |
|
340 |
let |
|
| 21082 | 341 |
val vars' = CodegenThingol.intro_vars [v] vars; |
| 21015 | 342 |
in |
| 22022 | 343 |
(concat [str "fn", str (CodegenThingol.lookup_var vars' v), str "=>"], vars') |
| 21015 | 344 |
end |
345 |
| pr ((v, SOME p), _) vars = |
|
346 |
let |
|
| 21093 | 347 |
val vars' = CodegenThingol.intro_vars [v] vars; |
348 |
val vs = CodegenThingol.fold_varnames (insert (op =)) p []; |
|
349 |
val vars'' = CodegenThingol.intro_vars vs vars'; |
|
| 21015 | 350 |
in |
| 22022 | 351 |
(concat [str "fn", str (CodegenThingol.lookup_var vars' v), str "as", |
| 21093 | 352 |
pr_term vars'' NOBR p, str "=>"], vars'') |
| 21015 | 353 |
end; |
354 |
val (ps', vars') = fold_map pr ps vars; |
|
355 |
in brackify BR (ps' @ [pr_term vars' NOBR t']) end |
|
356 |
| pr_term vars fxy (INum n) = |
|
| 20896 | 357 |
brackify BR [(str o IntInf.toString) n, str ":", str "IntInf.int"] |
| 21015 | 358 |
| 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
|
359 |
(str o prefix "#" o quote) |
| 20203 | 360 |
(let val i = ord c |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
361 |
in if i < 32 orelse i = 34 orelse i = 92 |
| 20105 | 362 |
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
|
363 |
else c |
|
07eeb832f28d
introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents:
19597
diff
changeset
|
364 |
end) |
| 21015 | 365 |
| pr_term vars fxy (t as ICase (_, [_])) = |
| 18216 | 366 |
let |
| 20976 | 367 |
val (ts, t') = CodegenThingol.unfold_let t; |
368 |
fun mk ((p, _), t'') vars = |
|
| 20699 | 369 |
let |
| 20896 | 370 |
val vs = CodegenThingol.fold_varnames (insert (op =)) p []; |
| 21082 | 371 |
val vars' = CodegenThingol.intro_vars vs vars; |
| 20699 | 372 |
in |
373 |
(Pretty.block [ |
|
| 22022 | 374 |
concat [ |
| 20699 | 375 |
str "val", |
| 20896 | 376 |
pr_term vars' NOBR p, |
| 20699 | 377 |
str "=", |
| 20976 | 378 |
pr_term vars NOBR t'' |
| 20699 | 379 |
], |
380 |
str ";" |
|
| 20896 | 381 |
], vars') |
| 20699 | 382 |
end |
| 20896 | 383 |
val (binds, vars') = fold_map mk ts vars; |
384 |
in |
|
385 |
Pretty.chunks [ |
|
386 |
[str ("let"), Pretty.fbrk, binds |> Pretty.chunks] |> Pretty.block,
|
|
| 20976 | 387 |
[str ("in"), Pretty.fbrk, pr_term vars' NOBR t'] |> Pretty.block,
|
| 20896 | 388 |
str ("end")
|
389 |
] end |
|
| 21015 | 390 |
| pr_term vars fxy (ICase ((td, ty), b::bs)) = |
| 18216 | 391 |
let |
| 20896 | 392 |
fun pr definer (p, t) = |
| 20699 | 393 |
let |
| 20896 | 394 |
val vs = CodegenThingol.fold_varnames (insert (op =)) p []; |
| 21082 | 395 |
val vars' = CodegenThingol.intro_vars vs vars; |
| 20699 | 396 |
in |
| 22022 | 397 |
concat [ |
| 20699 | 398 |
str definer, |
| 20896 | 399 |
pr_term vars' NOBR p, |
| 20699 | 400 |
str "=>", |
| 20896 | 401 |
pr_term vars' NOBR t |
| 20699 | 402 |
] |
| 20896 | 403 |
end; |
404 |
in |
|
405 |
(Pretty.enclose "(" ")" o single o brackify fxy) (
|
|
406 |
str "case" |
|
407 |
:: pr_term vars NOBR td |
|
408 |
:: pr "of" b |
|
409 |
:: map (pr "|") bs |
|
410 |
) |
|
411 |
end |
|
| 20976 | 412 |
and pr_app' vars (app as ((c, (iss, ty)), ts)) = |
413 |
if is_cons c then let |
|
414 |
val k = (length o fst o CodegenThingol.unfold_fun) ty |
|
415 |
in if k < 2 then |
|
416 |
(str o deresolv) c :: map (pr_term vars BR) ts |
|
417 |
else if k = length ts then |
|
418 |
[(str o deresolv) c, Pretty.enum "," "(" ")" (map (pr_term vars NOBR) ts)]
|
|
419 |
else [pr_term vars BR (CodegenThingol.eta_expand app k)] end else |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
420 |
(str o deresolv) c |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
421 |
:: ((map (pr_insts BR) o filter_out null) iss @ map (pr_term vars BR) ts) |
| 20896 | 422 |
and pr_app vars fxy (app as ((c, (iss, ty)), ts)) = |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
423 |
mk_app (pr_app' vars) (pr_term vars) const_syntax fxy app; |
| 21882 | 424 |
fun pr_def (MLFuns (funns as (funn :: funns'))) = |
| 20896 | 425 |
let |
426 |
val definer = |
|
427 |
let |
|
428 |
fun mk [] [] = "val" |
|
429 |
| mk (_::_) _ = "fun" |
|
430 |
| mk [] vs = if (null o filter_out (null o snd)) vs then "val" else "fun"; |
|
431 |
fun chk (_, ((ts, _) :: _, (vs, _))) NONE = SOME (mk ts vs) |
|
432 |
| chk (_, ((ts, _) :: _, (vs, _))) (SOME defi) = |
|
433 |
if defi = mk ts vs then SOME defi |
|
434 |
else error ("Mixing simultaneous vals and funs not implemented");
|
|
435 |
in the (fold chk funns NONE) end; |
|
436 |
fun pr_funn definer (name, (eqs as eq::eqs', (raw_vs, ty))) = |
|
437 |
let |
|
438 |
val vs = filter_out (null o snd) raw_vs; |
|
439 |
val shift = if null eqs' then I else |
|
440 |
map (Pretty.block o single o Pretty.block o single); |
|
441 |
fun pr_eq definer (ts, t) = |
|
442 |
let |
|
443 |
val consts = map_filter |
|
444 |
(fn c => if (is_some o const_syntax) c |
|
445 |
then NONE else (SOME o NameSpace.base o deresolv) c) |
|
446 |
((fold o CodegenThingol.fold_constnames) (insert (op =)) (t :: ts) []); |
|
447 |
val vars = keyword_vars |
|
| 21082 | 448 |
|> CodegenThingol.intro_vars consts |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
449 |
|> CodegenThingol.intro_vars ((fold o CodegenThingol.fold_unbound_varnames) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
450 |
(insert (op =)) ts []); |
| 20896 | 451 |
in |
| 22022 | 452 |
concat ( |
| 20896 | 453 |
[str definer, (str o deresolv) name] |
454 |
@ (if null ts andalso null vs |
|
455 |
andalso not (ty = ITyVar "_")(*for evaluation*) |
|
456 |
then [str ":", pr_typ NOBR ty] |
|
457 |
else |
|
458 |
map pr_tyvar vs |
|
459 |
@ map (pr_term vars BR) ts) |
|
460 |
@ [str "=", pr_term vars NOBR t] |
|
461 |
) |
|
462 |
end |
|
463 |
in |
|
464 |
(Pretty.block o Pretty.fbreaks o shift) ( |
|
465 |
pr_eq definer eq |
|
466 |
:: map (pr_eq "|") eqs' |
|
467 |
) |
|
468 |
end; |
|
469 |
val (ps, p) = split_last (pr_funn definer funn :: map (pr_funn "and") funns'); |
|
470 |
in Pretty.chunks (ps @ [Pretty.block ([p, str ";"])]) end |
|
471 |
| pr_def (MLDatas (datas as (data :: datas'))) = |
|
472 |
let |
|
473 |
fun pr_co (co, []) = |
|
474 |
str (deresolv co) |
|
475 |
| pr_co (co, tys) = |
|
| 22022 | 476 |
concat [ |
| 20896 | 477 |
str (deresolv co), |
478 |
str "of", |
|
| 21882 | 479 |
Pretty.enum " *" "" "" (map (pr_typ (INFX (2, X))) tys) |
| 20896 | 480 |
]; |
481 |
fun pr_data definer (tyco, (vs, cos)) = |
|
| 22022 | 482 |
concat ( |
| 20896 | 483 |
str definer |
484 |
:: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs) |
|
485 |
:: str "=" |
|
486 |
:: separate (str "|") (map pr_co cos) |
|
487 |
); |
|
488 |
val (ps, p) = split_last (pr_data "datatype" data :: map (pr_data "and") datas'); |
|
489 |
in Pretty.chunks (ps @ [Pretty.block ([p, str ";"])]) end |
|
490 |
| pr_def (MLClass (class, (superclasses, (v, classops)))) = |
|
491 |
let |
|
492 |
val w = dictvar v; |
|
493 |
fun pr_superclass class = |
|
| 22022 | 494 |
(concat o map str) [ |
| 20896 | 495 |
label class, ":", "'" ^ v, deresolv class |
496 |
]; |
|
497 |
fun pr_classop (classop, ty) = |
|
| 22022 | 498 |
concat [ |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
499 |
(*FIXME?*) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
500 |
(str o mk_classop_name) classop, str ":", pr_typ NOBR ty |
| 20896 | 501 |
]; |
502 |
fun pr_classop_fun (classop, _) = |
|
| 22022 | 503 |
concat [ |
| 20896 | 504 |
str "fun", |
505 |
(str o deresolv) classop, |
|
506 |
Pretty.enclose "(" ")" [str (w ^ ":'" ^ v ^ " " ^ deresolv class)],
|
|
507 |
str "=", |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
508 |
str ("#" ^ mk_classop_name classop),
|
| 20896 | 509 |
str (w ^ ";") |
510 |
]; |
|
511 |
in |
|
512 |
Pretty.chunks ( |
|
| 22022 | 513 |
concat [ |
| 20896 | 514 |
str ("type '" ^ v),
|
515 |
(str o deresolv) class, |
|
516 |
str "=", |
|
517 |
Pretty.enum "," "{" "};" (
|
|
518 |
map pr_superclass superclasses @ map pr_classop classops |
|
519 |
) |
|
520 |
] |
|
521 |
:: map pr_classop_fun classops |
|
522 |
) |
|
523 |
end |
|
524 |
| pr_def (MLClassinst (inst, ((class, (tyco, arity)), (superarities, classop_defs)))) = |
|
525 |
let |
|
526 |
fun pr_superclass (superclass, superinst_iss) = |
|
| 22022 | 527 |
concat [ |
| 20896 | 528 |
(str o label) superclass, |
529 |
str "=", |
|
530 |
pr_insts NOBR [Instance superinst_iss] |
|
531 |
]; |
|
532 |
fun pr_classop_def (classop, t) = |
|
533 |
let |
|
534 |
val consts = map_filter |
|
535 |
(fn c => if (is_some o const_syntax) c |
|
536 |
then NONE else (SOME o NameSpace.base o deresolv) c) |
|
537 |
(CodegenThingol.fold_constnames (insert (op =)) t []); |
|
538 |
val vars = keyword_vars |
|
| 21082 | 539 |
|> CodegenThingol.intro_vars consts; |
| 20896 | 540 |
in |
| 22022 | 541 |
concat [ |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
542 |
(str o mk_classop_name) classop, |
| 20896 | 543 |
str "=", |
544 |
pr_term vars NOBR t |
|
545 |
] |
|
546 |
end; |
|
547 |
in |
|
| 22022 | 548 |
concat ([ |
| 20896 | 549 |
str (if null arity then "val" else "fun"), |
550 |
(str o deresolv) inst ] @ |
|
551 |
map pr_tyvar arity @ [ |
|
552 |
str "=", |
|
553 |
Pretty.enum "," "{" "}" (map pr_superclass superarities @ map pr_classop_def classop_defs),
|
|
554 |
str ":", |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
555 |
pr_tycoexpr NOBR (class, [tyco `%% map (ITyVar o fst) arity]), |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
556 |
str ";;" |
| 20896 | 557 |
]) |
558 |
end; |
|
559 |
in pr_def ml_def end; |
|
|
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
560 |
|
| 21837 | 561 |
fun pr_sml_modl name content = |
562 |
Pretty.chunks ([ |
|
563 |
str ("structure " ^ name ^ " = "),
|
|
564 |
str "struct", |
|
565 |
str "" |
|
566 |
] @ content @ [ |
|
567 |
str "", |
|
568 |
str ("end; (*struct " ^ name ^ "*)")
|
|
569 |
]); |
|
570 |
||
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
571 |
fun pr_ocaml tyco_syntax const_syntax keyword_vars deresolv is_cons ml_def = |
| 21837 | 572 |
let |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
573 |
fun dictvar v = "_" ^ first_upper v; |
| 21837 | 574 |
fun pr_tyvar (v, []) = |
575 |
str "()" |
|
576 |
| pr_tyvar (v, sort) = |
|
577 |
let |
|
578 |
fun pr_class class = |
|
579 |
str ("'" ^ v ^ " " ^ deresolv class);
|
|
580 |
in |
|
581 |
Pretty.block [ |
|
582 |
str "(",
|
|
583 |
(str o dictvar) v, |
|
584 |
str ":", |
|
585 |
case sort |
|
586 |
of [class] => pr_class class |
|
587 |
| _ => Pretty.enum " *" "" "" (map pr_class sort), |
|
588 |
str ")" |
|
589 |
] |
|
590 |
end; |
|
591 |
fun pr_insts fxy iys = |
|
592 |
let |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
593 |
fun dot p2 p1 = Pretty.block [p1, str ".", str p2]; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
594 |
fun proj k i p = (brackify BR o map str) [ |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
595 |
"match", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
596 |
p, |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
597 |
"with", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
598 |
replicate i "_" |> nth_map k (K "d") |> separate (", ") |> implode,
|
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
599 |
"-> d" |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
600 |
] |
| 21837 | 601 |
fun pr_lookup [] p = |
602 |
p |
|
603 |
| pr_lookup [p'] p = |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
604 |
dot p' p |
| 21837 | 605 |
| pr_lookup (ps as _ :: _) p = |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
606 |
fold_rev dot ps p; |
| 21837 | 607 |
fun pr_inst fxy (Instance (inst, iss)) = |
608 |
brackify fxy ( |
|
609 |
(str o deresolv) inst |
|
610 |
:: map (pr_insts BR) iss |
|
611 |
) |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
612 |
| pr_inst fxy (Context ((classes, k), (v, i))) = |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
613 |
if i = 1 then pr_lookup (map deresolv classes) ((str o dictvar) v) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
614 |
else pr_lookup (map deresolv classes) (proj k i (dictvar v)); |
| 21837 | 615 |
in case iys |
616 |
of [] => str "()" |
|
617 |
| [iy] => pr_inst fxy iy |
|
618 |
| _ :: _ => (Pretty.list "(" ")" o map (pr_inst NOBR)) iys
|
|
619 |
end; |
|
620 |
fun pr_tycoexpr fxy (tyco, tys) = |
|
621 |
let |
|
622 |
val tyco' = (str o deresolv) tyco |
|
623 |
in case map (pr_typ BR) tys |
|
624 |
of [] => tyco' |
|
625 |
| [p] => Pretty.block [p, Pretty.brk 1, tyco'] |
|
626 |
| (ps as _::_) => Pretty.block [Pretty.list "(" ")" ps, Pretty.brk 1, tyco']
|
|
627 |
end |
|
628 |
and pr_typ fxy (tyco `%% tys) = |
|
629 |
(case tyco_syntax tyco |
|
630 |
of NONE => pr_tycoexpr fxy (tyco, tys) |
|
631 |
| SOME (i, pr) => |
|
632 |
if not (i = length tys) |
|
633 |
then error ("Number of argument mismatch in customary serialization: "
|
|
634 |
^ (string_of_int o length) tys ^ " given, " |
|
635 |
^ string_of_int i ^ " expected") |
|
636 |
else pr fxy pr_typ tys) |
|
637 |
| pr_typ fxy (ITyVar v) = |
|
638 |
str ("'" ^ v);
|
|
639 |
fun pr_term vars fxy (IConst c) = |
|
640 |
pr_app vars fxy (c, []) |
|
641 |
| pr_term vars fxy (IVar v) = |
|
642 |
str (CodegenThingol.lookup_var vars v) |
|
643 |
| pr_term vars fxy (t as t1 `$ t2) = |
|
644 |
(case CodegenThingol.unfold_const_app t |
|
645 |
of SOME c_ts => pr_app vars fxy c_ts |
|
646 |
| NONE => |
|
647 |
brackify fxy [pr_term vars NOBR t1, pr_term vars BR t2]) |
|
648 |
| pr_term vars fxy (t as _ `|-> _) = |
|
649 |
let |
|
650 |
val (ps, t') = CodegenThingol.unfold_abs t; |
|
651 |
fun pr ((v, NONE), _) vars = |
|
652 |
let |
|
653 |
val vars' = CodegenThingol.intro_vars [v] vars; |
|
654 |
in |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
655 |
(str (CodegenThingol.lookup_var vars' v), vars') |
| 21837 | 656 |
end |
657 |
| pr ((v, SOME p), _) vars = |
|
658 |
let |
|
659 |
val vars' = CodegenThingol.intro_vars [v] vars; |
|
660 |
val vs = CodegenThingol.fold_varnames (insert (op =)) p []; |
|
661 |
val vars'' = CodegenThingol.intro_vars vs vars'; |
|
662 |
in |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
663 |
(brackify BR [ |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
664 |
pr_term vars'' NOBR p, |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
665 |
str "as", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
666 |
str (CodegenThingol.lookup_var vars' v) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
667 |
], vars'') |
| 21837 | 668 |
end; |
669 |
val (ps', vars') = fold_map pr ps vars; |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
670 |
in brackify BR ( |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
671 |
str "fun" |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
672 |
:: ps' |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
673 |
@ str "->" |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
674 |
@@ pr_term vars' NOBR t' |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
675 |
) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
676 |
end |
| 21837 | 677 |
| pr_term vars fxy (INum n) = |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
678 |
brackify BR [str "Big_int.big_int_of_int", (str o IntInf.toString) n] |
| 21837 | 679 |
| pr_term vars _ (IChar c) = |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
680 |
(str o enclose "'" "'") |
| 21837 | 681 |
(let val i = ord c |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
682 |
in if i < 32 orelse i = 39 orelse i = 92 |
| 21837 | 683 |
then prefix "\\" (string_of_int i) |
684 |
else c |
|
685 |
end) |
|
686 |
| pr_term vars fxy (t as ICase (_, [_])) = |
|
687 |
let |
|
688 |
val (ts, t') = CodegenThingol.unfold_let t; |
|
689 |
fun mk ((p, _), t'') vars = |
|
690 |
let |
|
691 |
val vs = CodegenThingol.fold_varnames (insert (op =)) p []; |
|
692 |
val vars' = CodegenThingol.intro_vars vs vars; |
|
693 |
in |
|
| 22022 | 694 |
(concat [ |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
695 |
str "let", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
696 |
pr_term vars' NOBR p, |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
697 |
str "=", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
698 |
pr_term vars NOBR t'', |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
699 |
str "in" |
| 21837 | 700 |
], vars') |
701 |
end |
|
702 |
val (binds, vars') = fold_map mk ts vars; |
|
703 |
in |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
704 |
Pretty.chunks (binds @ [pr_term vars' NOBR t']) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
705 |
end |
| 21837 | 706 |
| pr_term vars fxy (ICase ((td, ty), b::bs)) = |
707 |
let |
|
708 |
fun pr definer (p, t) = |
|
709 |
let |
|
710 |
val vs = CodegenThingol.fold_varnames (insert (op =)) p []; |
|
711 |
val vars' = CodegenThingol.intro_vars vs vars; |
|
712 |
in |
|
| 22022 | 713 |
concat [ |
| 21837 | 714 |
str definer, |
715 |
pr_term vars' NOBR p, |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
716 |
str "->", |
| 21837 | 717 |
pr_term vars' NOBR t |
718 |
] |
|
719 |
end; |
|
720 |
in |
|
721 |
(Pretty.enclose "(" ")" o single o brackify fxy) (
|
|
722 |
str "match" |
|
723 |
:: pr_term vars NOBR td |
|
724 |
:: pr "with" b |
|
725 |
:: map (pr "|") bs |
|
726 |
) |
|
727 |
end |
|
728 |
and pr_app' vars (app as ((c, (iss, ty)), ts)) = |
|
| 21952 | 729 |
if is_cons c then |
730 |
if (length o fst o CodegenThingol.unfold_fun) ty = length ts |
|
731 |
then case ts |
|
732 |
of [] => [(str o deresolv) c] |
|
733 |
| [t] => [(str o deresolv) c, pr_term vars BR t] |
|
734 |
| _ => [(str o deresolv) c, Pretty.enum "," "(" ")" (map (pr_term vars NOBR) ts)]
|
|
735 |
else [pr_term vars BR (CodegenThingol.eta_expand app ((length o fst o CodegenThingol.unfold_fun) ty))] |
|
736 |
else (str o deresolv) c |
|
737 |
:: ((map (pr_insts BR) o filter_out null) iss @ map (pr_term vars BR) ts) |
|
| 21837 | 738 |
and pr_app vars fxy (app as ((c, (iss, ty)), ts)) = |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
739 |
mk_app (pr_app' vars) (pr_term vars) const_syntax fxy app; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
740 |
fun pr_def (MLFuns (funns as funn :: funns')) = |
| 21837 | 741 |
let |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
742 |
fun fish_parm _ (w as SOME _) = w |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
743 |
| fish_parm (IVar v) NONE = SOME v |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
744 |
| fish_parm _ NONE = NONE; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
745 |
fun fillup_parm _ (_, SOME v) = v |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
746 |
| fillup_parm x (i, NONE) = x ^ string_of_int i; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
747 |
fun fish_parms vars eqs = |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
748 |
let |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
749 |
val raw_fished = fold (map2 fish_parm) eqs (replicate (length (hd eqs)) NONE); |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
750 |
val x = Name.variant (map_filter I raw_fished) "x"; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
751 |
val fished = map_index (fillup_parm x) raw_fished; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
752 |
val vars' = CodegenThingol.intro_vars fished vars; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
753 |
in map (CodegenThingol.lookup_var vars') fished end; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
754 |
fun pr_eq (ts, t) = |
| 21837 | 755 |
let |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
756 |
val consts = map_filter |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
757 |
(fn c => if (is_some o const_syntax) c |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
758 |
then NONE else (SOME o NameSpace.base o deresolv) c) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
759 |
((fold o CodegenThingol.fold_constnames) (insert (op =)) (t :: ts) []); |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
760 |
val vars = keyword_vars |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
761 |
|> CodegenThingol.intro_vars consts |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
762 |
|> CodegenThingol.intro_vars ((fold o CodegenThingol.fold_unbound_varnames) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
763 |
(insert (op =)) ts []); |
| 22022 | 764 |
in concat [ |
| 21952 | 765 |
(Pretty.block o Pretty.commas) (map (pr_term vars NOBR) ts), |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
766 |
str "->", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
767 |
pr_term vars NOBR t |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
768 |
] end; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
769 |
fun pr_eqs [(ts, t)] = |
| 21837 | 770 |
let |
771 |
val consts = map_filter |
|
772 |
(fn c => if (is_some o const_syntax) c |
|
773 |
then NONE else (SOME o NameSpace.base o deresolv) c) |
|
774 |
((fold o CodegenThingol.fold_constnames) (insert (op =)) (t :: ts) []); |
|
775 |
val vars = keyword_vars |
|
776 |
|> CodegenThingol.intro_vars consts |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
777 |
|> CodegenThingol.intro_vars ((fold o CodegenThingol.fold_unbound_varnames) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
778 |
(insert (op =)) ts []); |
| 21837 | 779 |
in |
| 22022 | 780 |
concat ( |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
781 |
map (pr_term vars BR) ts |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
782 |
@ str "=" |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
783 |
@@ pr_term vars NOBR t |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
784 |
) |
| 21837 | 785 |
end |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
786 |
| pr_eqs (eqs as (eq as ([_], _)) :: eqs') = |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
787 |
Pretty.block ( |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
788 |
str "=" |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
789 |
:: Pretty.brk 1 |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
790 |
:: str "function" |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
791 |
:: Pretty.brk 1 |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
792 |
:: pr_eq eq |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
793 |
:: maps (append [Pretty.fbrk, str "|", Pretty.brk 1] o single o pr_eq) eqs' |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
794 |
) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
795 |
| pr_eqs (eqs as eq :: eqs') = |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
796 |
let |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
797 |
val consts = map_filter |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
798 |
(fn c => if (is_some o const_syntax) c |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
799 |
then NONE else (SOME o NameSpace.base o deresolv) c) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
800 |
((fold o CodegenThingol.fold_constnames) (insert (op =)) (map snd eqs) []); |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
801 |
val vars = keyword_vars |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
802 |
|> CodegenThingol.intro_vars consts; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
803 |
val dummy_parms = (map str o fish_parms vars o map fst) eqs; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
804 |
in |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
805 |
Pretty.block ( |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
806 |
Pretty.breaks dummy_parms |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
807 |
@ Pretty.brk 1 |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
808 |
:: str "=" |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
809 |
:: Pretty.brk 1 |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
810 |
:: str "match" |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
811 |
:: Pretty.brk 1 |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
812 |
:: (Pretty.block o Pretty.commas) dummy_parms |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
813 |
:: Pretty.brk 1 |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
814 |
:: str "with" |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
815 |
:: Pretty.brk 1 |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
816 |
:: pr_eq eq |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
817 |
:: maps (append [Pretty.fbrk, str "|", Pretty.brk 1] o single o pr_eq) eqs' |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
818 |
) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
819 |
end; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
820 |
fun pr_funn definer (name, (eqs, (vs, ty))) = |
| 22022 | 821 |
concat ( |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
822 |
str definer |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
823 |
:: (str o deresolv) name |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
824 |
:: map_filter (fn (_, []) => NONE | v => SOME (pr_tyvar v)) vs |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
825 |
@| pr_eqs eqs |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
826 |
); |
| 21837 | 827 |
val (ps, p) = split_last (pr_funn "let rec" funn :: map (pr_funn "and") funns'); |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
828 |
in Pretty.chunks (ps @ [Pretty.block ([p, str ";;"])]) end |
| 21837 | 829 |
| pr_def (MLDatas (datas as (data :: datas'))) = |
830 |
let |
|
831 |
fun pr_co (co, []) = |
|
832 |
str (deresolv co) |
|
833 |
| pr_co (co, tys) = |
|
| 22022 | 834 |
concat [ |
| 21837 | 835 |
str (deresolv co), |
836 |
str "of", |
|
| 21882 | 837 |
Pretty.enum " *" "" "" (map (pr_typ (INFX (2, X))) tys) |
| 21837 | 838 |
]; |
839 |
fun pr_data definer (tyco, (vs, cos)) = |
|
| 22022 | 840 |
concat ( |
| 21837 | 841 |
str definer |
842 |
:: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs) |
|
843 |
:: str "=" |
|
844 |
:: separate (str "|") (map pr_co cos) |
|
845 |
); |
|
846 |
val (ps, p) = split_last (pr_data "type" data :: map (pr_data "and") datas'); |
|
847 |
in Pretty.chunks (ps @ [Pretty.block ([p, str ";;"])]) end |
|
848 |
| pr_def (MLClass (class, (superclasses, (v, classops)))) = |
|
849 |
let |
|
850 |
val w = dictvar v; |
|
851 |
fun pr_superclass class = |
|
| 22022 | 852 |
(concat o map str) [ |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
853 |
deresolv class, ":", "'" ^ v, deresolv class |
| 21837 | 854 |
]; |
855 |
fun pr_classop (classop, ty) = |
|
| 22022 | 856 |
concat [ |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
857 |
(str o deresolv) classop, str ":", pr_typ NOBR ty |
| 21837 | 858 |
]; |
859 |
fun pr_classop_fun (classop, _) = |
|
| 22022 | 860 |
concat [ |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
861 |
str "let", |
| 21837 | 862 |
(str o deresolv) classop, |
863 |
Pretty.enclose "(" ")" [str (w ^ ":'" ^ v ^ " " ^ deresolv class)],
|
|
864 |
str "=", |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
865 |
str (w ^ "." ^ deresolv classop ^ ";;") |
| 21837 | 866 |
]; |
867 |
in |
|
868 |
Pretty.chunks ( |
|
| 22022 | 869 |
concat [ |
| 21837 | 870 |
str ("type '" ^ v),
|
871 |
(str o deresolv) class, |
|
872 |
str "=", |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
873 |
Pretty.enum ";" "{" "};;" (
|
| 21837 | 874 |
map pr_superclass superclasses @ map pr_classop classops |
875 |
) |
|
876 |
] |
|
877 |
:: map pr_classop_fun classops |
|
878 |
) |
|
879 |
end |
|
880 |
| pr_def (MLClassinst (inst, ((class, (tyco, arity)), (superarities, classop_defs)))) = |
|
881 |
let |
|
882 |
fun pr_superclass (superclass, superinst_iss) = |
|
| 22022 | 883 |
concat [ |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
884 |
(str o deresolv) superclass, |
| 21837 | 885 |
str "=", |
886 |
pr_insts NOBR [Instance superinst_iss] |
|
887 |
]; |
|
888 |
fun pr_classop_def (classop, t) = |
|
889 |
let |
|
890 |
val consts = map_filter |
|
891 |
(fn c => if (is_some o const_syntax) c |
|
892 |
then NONE else (SOME o NameSpace.base o deresolv) c) |
|
893 |
(CodegenThingol.fold_constnames (insert (op =)) t []); |
|
894 |
val vars = keyword_vars |
|
895 |
|> CodegenThingol.intro_vars consts; |
|
896 |
in |
|
| 22022 | 897 |
concat [ |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
898 |
(str o deresolv) classop, |
| 21837 | 899 |
str "=", |
900 |
pr_term vars NOBR t |
|
901 |
] |
|
902 |
end; |
|
903 |
in |
|
| 22022 | 904 |
concat ( |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
905 |
str "let" |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
906 |
:: (str o deresolv) inst |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
907 |
:: map pr_tyvar arity |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
908 |
@ str "=" |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
909 |
@@ (Pretty.enclose "(" ");;" o Pretty.breaks) [
|
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
910 |
Pretty.enum ";" "{" "}" (map pr_superclass superarities @ map pr_classop_def classop_defs),
|
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
911 |
str ":", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
912 |
pr_tycoexpr NOBR (class, [tyco `%% map (ITyVar o fst) arity]) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
913 |
] |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
914 |
) |
| 21837 | 915 |
end; |
916 |
in pr_def ml_def end; |
|
917 |
||
918 |
fun pr_ocaml_modl name content = |
|
919 |
Pretty.chunks ([ |
|
920 |
str ("module " ^ name ^ " = "),
|
|
921 |
str "struct", |
|
922 |
str "" |
|
923 |
] @ content @ [ |
|
924 |
str "", |
|
925 |
str ("end;; (*struct " ^ name ^ "*)")
|
|
926 |
]); |
|
927 |
||
| 22022 | 928 |
val code_width = ref 80; |
929 |
fun code_output p = Pretty.setmp_margin (!code_width) Pretty.output p ^ "\n"; |
|
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
930 |
|
| 21837 | 931 |
fun seri_ml pr_def pr_modl output reserved_user module_alias module_prolog |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
932 |
(_ : string -> (string * (string -> string option)) option) tyco_syntax const_syntax code = |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
933 |
let |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
934 |
val is_cons = fn node => case CodegenThingol.get_def code node |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
935 |
of CodegenThingol.Datatypecons _ => true |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
936 |
| _ => false; |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
937 |
datatype node = |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
938 |
Def of string * ml_def option |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
939 |
| Module of string * ((Name.context * Name.context) * node Graph.T); |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
940 |
val empty_names = ML_Syntax.reserved |> fold Name.declare reserved_user; |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
941 |
val empty_module = ((empty_names, empty_names), Graph.empty); |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
942 |
fun map_node [] f = f |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
943 |
| map_node (m::ms) f = |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
944 |
Graph.default_node (m, Module (m, empty_module)) |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
945 |
#> Graph.map_node m (fn (Module (dmodlname, (nsp, nodes))) => Module (dmodlname, (nsp, map_node ms f nodes))); |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
946 |
fun map_nsp_yield [] f (nsp, nodes) = |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
947 |
let |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
948 |
val (x, nsp') = f nsp |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
949 |
in (x, (nsp', nodes)) end |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
950 |
| map_nsp_yield (m::ms) f (nsp, nodes) = |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
951 |
let |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
952 |
val (x, nodes') = |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
953 |
nodes |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
954 |
|> Graph.default_node (m, Module (m, empty_module)) |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
955 |
|> Graph.map_node_yield m (fn Module (dmodlname, nsp_nodes) => |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
956 |
let |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
957 |
val (x, nsp_nodes') = map_nsp_yield ms f nsp_nodes |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
958 |
in (x, Module (dmodlname, nsp_nodes')) end) |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
959 |
in (x, (nsp, nodes')) end; |
| 21719 | 960 |
val init_vars = CodegenThingol.make_vars (ML_Syntax.reserved_names @ reserved_user); |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
961 |
val name_modl = mk_modl_name_tab empty_names NONE module_alias code; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
962 |
fun name_def upper name nsp = |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
963 |
let |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
964 |
val (_, base) = dest_name name; |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
965 |
val base' = if upper then first_upper base else base; |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
966 |
val ([base''], nsp') = Name.variants [base'] nsp; |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
967 |
in (base'', nsp') end; |
| 21130 | 968 |
fun map_nsp_fun f (nsp_fun, nsp_typ) = |
969 |
let |
|
970 |
val (x, nsp_fun') = f nsp_fun |
|
971 |
in (x, (nsp_fun', nsp_typ)) end; |
|
972 |
fun map_nsp_typ f (nsp_fun, nsp_typ) = |
|
973 |
let |
|
974 |
val (x, nsp_typ') = f nsp_typ |
|
975 |
in (x, (nsp_fun, nsp_typ')) end; |
|
976 |
fun mk_funs defs = |
|
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
977 |
fold_map |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
978 |
(fn (name, CodegenThingol.Fun info) => |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
979 |
map_nsp_fun (name_def false name) >> (fn base => (base, (name, info))) |
| 21130 | 980 |
| (name, def) => error ("Function block containing illegal def: " ^ quote name)
|
981 |
) defs |
|
982 |
>> (split_list #> apsnd MLFuns); |
|
983 |
fun mk_datatype defs = |
|
984 |
fold_map |
|
985 |
(fn (name, CodegenThingol.Datatype info) => |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
986 |
map_nsp_typ (name_def false name) >> (fn base => (base, SOME (name, info))) |
| 21130 | 987 |
| (name, CodegenThingol.Datatypecons _) => |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
988 |
map_nsp_fun (name_def true name) >> (fn base => (base, NONE)) |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
989 |
| (name, def) => error ("Datatype block containing illegal def: " ^ quote name)
|
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
990 |
) defs |
| 21130 | 991 |
>> (split_list #> apsnd (map_filter I |
992 |
#> (fn [] => error ("Datatype block without data: " ^ (commas o map (quote o fst)) defs)
|
|
993 |
| infos => MLDatas infos))); |
|
994 |
fun mk_class defs = |
|
995 |
fold_map |
|
996 |
(fn (name, CodegenThingol.Class info) => |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
997 |
map_nsp_typ (name_def false name) >> (fn base => (base, SOME (name, info))) |
| 21130 | 998 |
| (name, CodegenThingol.Classop _) => |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
999 |
map_nsp_fun (name_def false name) >> (fn base => (base, NONE)) |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1000 |
| (name, def) => error ("Class block containing illegal def: " ^ quote name)
|
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1001 |
) defs |
| 21130 | 1002 |
>> (split_list #> apsnd (map_filter I |
1003 |
#> (fn [] => error ("Class block without class: " ^ (commas o map (quote o fst)) defs)
|
|
1004 |
| [info] => MLClass info))); |
|
1005 |
fun mk_inst [(name, CodegenThingol.Classinst info)] = |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1006 |
map_nsp_fun (name_def false name) |
| 21130 | 1007 |
>> (fn base => ([base], MLClassinst (name, info))); |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1008 |
fun add_group mk defs nsp_nodes = |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1009 |
let |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1010 |
val names as (name :: names') = map fst defs; |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1011 |
val deps = |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1012 |
[] |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1013 |
|> fold (fold (insert (op =)) o Graph.imm_succs code) names |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1014 |
|> subtract (op =) names; |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1015 |
val (modls, _) = (split_list o map dest_name) names; |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1016 |
val modl = (the_single o distinct (op =)) modls |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1017 |
handle Empty => |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1018 |
error ("Illegal mutual dependencies: " ^ commas names);
|
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1019 |
val modl' = name_modl modl; |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1020 |
val modl_explode = NameSpace.explode modl'; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1021 |
fun add_dep name name'' = |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1022 |
let |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1023 |
val modl'' = (name_modl o fst o dest_name) name''; |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1024 |
in if modl' = modl'' then |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1025 |
map_node modl_explode |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1026 |
(Graph.add_edge (name, name'')) |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1027 |
else let |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1028 |
val (common, (diff1::_, diff2::_)) = chop_prefix (op =) (modl_explode, NameSpace.explode modl''); |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1029 |
in |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1030 |
map_node common |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1031 |
(fn gr => Graph.add_edge_acyclic (diff1, diff2) gr |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1032 |
handle Graph.CYCLES _ => error ("Dependency "
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1033 |
^ quote name |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1034 |
^ " -> " ^ quote name'' ^ " would result in module dependency cycle")) |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1035 |
end end; |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1036 |
in |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1037 |
nsp_nodes |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1038 |
|> map_nsp_yield modl_explode (mk defs) |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1039 |
|-> (fn (base' :: bases', def') => |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1040 |
apsnd (map_node modl_explode (Graph.new_node (name, (Def (base', SOME def'))) |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1041 |
#> fold2 (fn name' => fn base' => Graph.new_node (name', (Def (base', NONE)))) names' bases'))) |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1042 |
|> apsnd (fold (fn name => fold (add_dep name) deps) names) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1043 |
|> apsnd (fold (map_node modl_explode o Graph.add_edge) (product names names)) |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1044 |
end; |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1045 |
fun group_defs [(_, CodegenThingol.Bot)] = |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1046 |
I |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1047 |
| group_defs ((defs as (_, CodegenThingol.Fun _)::_)) = |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1048 |
add_group mk_funs defs |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1049 |
| group_defs ((defs as (_, CodegenThingol.Datatypecons _)::_)) = |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1050 |
add_group mk_datatype defs |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1051 |
| group_defs ((defs as (_, CodegenThingol.Datatype _)::_)) = |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1052 |
add_group mk_datatype defs |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1053 |
| group_defs ((defs as (_, CodegenThingol.Class _)::_)) = |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1054 |
add_group mk_class defs |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1055 |
| group_defs ((defs as (_, CodegenThingol.Classop _)::_)) = |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1056 |
add_group mk_class defs |
| 21130 | 1057 |
| group_defs ((defs as [(_, CodegenThingol.Classinst _)])) = |
1058 |
add_group mk_inst defs |
|
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1059 |
| group_defs defs = error ("Illegal mutual dependencies: " ^ (commas o map fst) defs)
|
| 21130 | 1060 |
val (_, nodes) = |
1061 |
empty_module |
|
1062 |
|> fold group_defs (map (AList.make (Graph.get_node code)) |
|
1063 |
(rev (Graph.strong_conn code))) |
|
1064 |
fun deresolver prefix name = |
|
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1065 |
let |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1066 |
val modl = (fst o dest_name) name; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1067 |
val modl' = (NameSpace.explode o name_modl) modl; |
| 21130 | 1068 |
val (_, (_, remainder)) = chop_prefix (op =) (prefix, modl'); |
1069 |
val defname' = |
|
1070 |
nodes |
|
1071 |
|> fold (fn m => fn g => case Graph.get_node g m |
|
1072 |
of Module (_, (_, g)) => g) modl' |
|
1073 |
|> (fn g => case Graph.get_node g name of Def (defname, _) => defname); |
|
|
21858
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents:
21837
diff
changeset
|
1074 |
in NameSpace.implode (remainder @ [defname']) end handle Graph.UNDEF _ => |
| 21285 | 1075 |
"(raise Fail \"undefined name " ^ name ^ "\")"; |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1076 |
fun the_prolog modlname = case module_prolog modlname |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1077 |
of NONE => [] |
| 21130 | 1078 |
| SOME p => [p, str ""]; |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1079 |
fun pr_node prefix (Def (_, NONE)) = |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1080 |
NONE |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1081 |
| pr_node prefix (Def (_, SOME def)) = |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1082 |
SOME (pr_def tyco_syntax const_syntax init_vars (deresolver prefix) is_cons def) |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1083 |
| pr_node prefix (Module (dmodlname, (_, nodes))) = |
|
21858
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents:
21837
diff
changeset
|
1084 |
SOME (pr_modl dmodlname (the_prolog (NameSpace.implode (prefix @ [dmodlname])) |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1085 |
@ separate (str "") ((map_filter (pr_node (prefix @ [dmodlname]) o Graph.get_node nodes) |
| 21837 | 1086 |
o rev o flat o Graph.strong_conn) nodes))); |
1087 |
val p = pr_modl "ROOT" (the_prolog "" @ separate (str "") ((map_filter |
|
1088 |
(pr_node [] o Graph.get_node nodes) o rev o flat o Graph.strong_conn) nodes)) |
|
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1089 |
in output p end; |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1090 |
|
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1091 |
val isar_seri_sml = |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1092 |
let |
| 22022 | 1093 |
fun output_file file = File.write (Path.explode file) o code_output; |
1094 |
val output_diag = writeln o code_output; |
|
1095 |
val output_internal = use_text Output.ml_output false o code_output; |
|
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1096 |
in |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1097 |
parse_args ((Args.$$$ "-" >> K output_diag |
| 21548 | 1098 |
|| Args.$$$ "#" >> K output_internal |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1099 |
|| Args.name >> output_file) |
| 21837 | 1100 |
>> (fn output => seri_ml pr_sml pr_sml_modl output)) |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1101 |
end; |
| 20896 | 1102 |
|
| 21837 | 1103 |
val isar_seri_ocaml = |
1104 |
let |
|
| 22022 | 1105 |
fun output_file file = File.write (Path.explode file) o code_output; |
1106 |
val output_diag = writeln o code_output; |
|
| 21837 | 1107 |
in |
1108 |
parse_args ((Args.$$$ "-" >> K output_diag |
|
1109 |
|| Args.name >> output_file) |
|
1110 |
>> (fn output => seri_ml pr_ocaml pr_ocaml_modl output)) |
|
1111 |
end; |
|
| 20896 | 1112 |
|
| 21162 | 1113 |
|
| 20896 | 1114 |
(** Haskell serializer **) |
1115 |
||
| 21082 | 1116 |
fun pr_haskell class_syntax tyco_syntax const_syntax keyword_vars deresolv_here deresolv deriving_show def = |
|
19042
630b8dd0b31a
exported some interfaces useful for other code generator approaches
haftmann
parents:
19038
diff
changeset
|
1117 |
let |
| 20896 | 1118 |
fun class_name class = case class_syntax class |
1119 |
of NONE => deresolv class |
|
1120 |
| SOME (class, _) => class; |
|
1121 |
fun classop_name class classop = case class_syntax class |
|
| 22022 | 1122 |
of NONE => deresolv_here classop |
| 20896 | 1123 |
| SOME (_, classop_syntax) => case classop_syntax classop |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1124 |
of NONE => (snd o dest_name) classop |
| 20896 | 1125 |
| SOME classop => classop |
1126 |
fun pr_typparms tyvars vs = |
|
1127 |
case maps (fn (v, sort) => map (pair v) sort) vs |
|
1128 |
of [] => str "" |
|
1129 |
| xs => Pretty.block [ |
|
1130 |
Pretty.enum "," "(" ")" (
|
|
1131 |
map (fn (v, class) => str |
|
| 21082 | 1132 |
(class_name class ^ " " ^ CodegenThingol.lookup_var tyvars v)) xs |
| 20896 | 1133 |
), |
1134 |
str " => " |
|
1135 |
]; |
|
1136 |
fun pr_tycoexpr tyvars fxy (tyco, tys) = |
|
1137 |
brackify fxy (str tyco :: map (pr_typ tyvars BR) tys) |
|
1138 |
and pr_typ tyvars fxy (tycoexpr as tyco `%% tys) = |
|
1139 |
(case tyco_syntax tyco |
|
1140 |
of NONE => |
|
1141 |
pr_tycoexpr tyvars fxy (deresolv tyco, tys) |
|
1142 |
| SOME (i, pr) => |
|
1143 |
if not (i = length tys) |
|
1144 |
then error ("Number of argument mismatch in customary serialization: "
|
|
1145 |
^ (string_of_int o length) tys ^ " given, " |
|
1146 |
^ string_of_int i ^ " expected") |
|
1147 |
else pr fxy (pr_typ tyvars) tys) |
|
1148 |
| pr_typ tyvars fxy (ITyVar v) = |
|
| 21082 | 1149 |
(str o CodegenThingol.lookup_var tyvars) v; |
| 20896 | 1150 |
fun pr_typscheme_expr tyvars (vs, tycoexpr) = |
1151 |
Pretty.block [pr_typparms tyvars vs, pr_tycoexpr tyvars NOBR tycoexpr]; |
|
1152 |
fun pr_typscheme tyvars (vs, ty) = |
|
1153 |
Pretty.block [pr_typparms tyvars vs, pr_typ tyvars NOBR ty]; |
|
1154 |
fun pr_term vars fxy (IConst c) = |
|
1155 |
pr_app vars fxy (c, []) |
|
1156 |
| pr_term vars fxy (t as (t1 `$ t2)) = |
|
1157 |
(case CodegenThingol.unfold_const_app t |
|
1158 |
of SOME app => pr_app vars fxy app |
|
1159 |
| _ => |
|
1160 |
brackify fxy [ |
|
1161 |
pr_term vars NOBR t1, |
|
1162 |
pr_term vars BR t2 |
|
1163 |
]) |
|
1164 |
| pr_term vars fxy (IVar v) = |
|
| 21082 | 1165 |
(str o CodegenThingol.lookup_var vars) v |
| 20896 | 1166 |
| 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
|
1167 |
let |
| 21015 | 1168 |
val (ps, t') = CodegenThingol.unfold_abs t; |
1169 |
fun pr ((v, SOME p), _) vars = |
|
1170 |
let |
|
| 21093 | 1171 |
val vars' = CodegenThingol.intro_vars [v] vars; |
1172 |
val vs = CodegenThingol.fold_varnames (insert (op =)) p []; |
|
1173 |
val vars'' = CodegenThingol.intro_vars vs vars'; |
|
1174 |
in |
|
| 22022 | 1175 |
(concat [str (CodegenThingol.lookup_var vars' v), |
| 21093 | 1176 |
str "@", pr_term vars'' BR p], vars'') |
1177 |
end |
|
| 21015 | 1178 |
| pr ((v, NONE), _) vars = |
1179 |
let |
|
| 21082 | 1180 |
val vars' = CodegenThingol.intro_vars [v] vars; |
1181 |
in (str (CodegenThingol.lookup_var vars' v), vars') end; |
|
| 21015 | 1182 |
val (ps', vars') = fold_map pr ps vars; |
| 20896 | 1183 |
in |
1184 |
brackify BR ( |
|
1185 |
str "\\" |
|
| 21015 | 1186 |
:: ps' @ [ |
| 20896 | 1187 |
str "->", |
1188 |
pr_term vars' NOBR t' |
|
1189 |
]) |
|
1190 |
end |
|
| 21015 | 1191 |
| pr_term vars fxy (INum n) = |
| 20896 | 1192 |
if n > 0 then |
1193 |
(str o IntInf.toString) n |
|
1194 |
else |
|
1195 |
brackify BR [(str o Library.prefix "-" o IntInf.toString o IntInf.~) n] |
|
| 21015 | 1196 |
| pr_term vars fxy (IChar c) = |
| 20896 | 1197 |
(str o enclose "'" "'") |
1198 |
(let val i = (Char.ord o the o Char.fromString) c |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1199 |
in if i < 32 orelse i = 39 orelse i = 92 |
| 20896 | 1200 |
then Library.prefix "\\" (string_of_int i) |
1201 |
else c |
|
1202 |
end) |
|
| 21015 | 1203 |
| pr_term vars fxy (t as ICase (_, [_])) = |
| 18216 | 1204 |
let |
| 20896 | 1205 |
val (ts, t) = CodegenThingol.unfold_let t; |
1206 |
fun pr ((p, _), t) vars = |
|
1207 |
let |
|
1208 |
val vs = CodegenThingol.fold_varnames (insert (op =)) p []; |
|
| 21082 | 1209 |
val vars' = CodegenThingol.intro_vars vs vars; |
| 20896 | 1210 |
in |
| 22022 | 1211 |
(semicolon [ |
| 20896 | 1212 |
pr_term vars' BR p, |
1213 |
str "=", |
|
1214 |
pr_term vars NOBR t |
|
1215 |
], vars') |
|
1216 |
end; |
|
1217 |
val (binds, vars') = fold_map pr ts vars; |
|
| 22022 | 1218 |
in |
1219 |
Pretty.block_enclose ( |
|
1220 |
str "let {",
|
|
1221 |
Pretty.block [str "} in ", pr_term vars' NOBR t] |
|
1222 |
) binds |
|
1223 |
end |
|
| 21015 | 1224 |
| pr_term vars fxy (ICase ((td, _), bs)) = |
| 20896 | 1225 |
let |
1226 |
fun pr (p, t) = |
|
1227 |
let |
|
1228 |
val vs = CodegenThingol.fold_varnames (insert (op =)) p []; |
|
| 21082 | 1229 |
val vars' = CodegenThingol.intro_vars vs vars; |
| 20896 | 1230 |
in |
| 22022 | 1231 |
semicolon [ |
| 20896 | 1232 |
pr_term vars' NOBR p, |
1233 |
str "->", |
|
1234 |
pr_term vars' NOBR t |
|
1235 |
] |
|
1236 |
end |
|
1237 |
in |
|
| 22022 | 1238 |
Pretty.block_enclose ( |
1239 |
concat [str "(case", pr_term vars NOBR td, str "of", str "{"],
|
|
1240 |
str "})" |
|
1241 |
) (map pr bs) |
|
| 20896 | 1242 |
end |
| 20976 | 1243 |
and pr_app' vars ((c, _), ts) = |
| 20896 | 1244 |
(str o deresolv) c :: map (pr_term vars BR) ts |
1245 |
and pr_app vars fxy = |
|
1246 |
mk_app (pr_app' vars) (pr_term vars) const_syntax fxy; |
|
| 21882 | 1247 |
fun pr_def (name, CodegenThingol.Fun (eqs, (vs, ty))) = |
| 20896 | 1248 |
let |
| 21082 | 1249 |
val tyvars = CodegenThingol.intro_vars (map fst vs) keyword_vars; |
| 20896 | 1250 |
fun pr_eq (ts, t) = |
| 20699 | 1251 |
let |
1252 |
val consts = map_filter |
|
1253 |
(fn c => if (is_some o const_syntax) c |
|
| 20896 | 1254 |
then NONE else (SOME o NameSpace.base o deresolv) c) |
1255 |
((fold o CodegenThingol.fold_constnames) (insert (op =)) (t :: ts) []); |
|
1256 |
val vars = keyword_vars |
|
| 21082 | 1257 |
|> CodegenThingol.intro_vars consts |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1258 |
|> CodegenThingol.intro_vars ((fold o CodegenThingol.fold_unbound_varnames) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1259 |
(insert (op =)) ts []); |
| 20699 | 1260 |
in |
| 22022 | 1261 |
semicolon ( |
| 20896 | 1262 |
(str o deresolv_here) name |
1263 |
:: map (pr_term vars BR) ts |
|
| 22022 | 1264 |
@ str "=" |
1265 |
@@ pr_term vars NOBR t |
|
| 20699 | 1266 |
) |
| 20896 | 1267 |
end; |
1268 |
in |
|
1269 |
Pretty.chunks ( |
|
1270 |
Pretty.block [ |
|
1271 |
(str o suffix " ::" o deresolv_here) name, |
|
1272 |
Pretty.brk 1, |
|
| 22022 | 1273 |
pr_typscheme tyvars (vs, ty), |
1274 |
str ";" |
|
| 20896 | 1275 |
] |
| 21882 | 1276 |
:: map pr_eq eqs |
| 20896 | 1277 |
) |
| 21082 | 1278 |
end |
| 20896 | 1279 |
| pr_def (name, CodegenThingol.Datatype (vs, [(co, [ty])])) = |
1280 |
let |
|
| 21082 | 1281 |
val tyvars = CodegenThingol.intro_vars (map fst vs) keyword_vars; |
| 18216 | 1282 |
in |
| 22022 | 1283 |
semicolon ( |
1284 |
str "newtype" |
|
1285 |
:: pr_typscheme_expr tyvars (vs, (deresolv_here name, map (ITyVar o fst) vs)) |
|
1286 |
:: str "=" |
|
1287 |
:: (str o deresolv_here) co |
|
1288 |
:: pr_typ tyvars BR ty |
|
1289 |
:: (if deriving_show name then [str "deriving (Read, Show)"] else []) |
|
1290 |
) |
|
| 21082 | 1291 |
end |
| 20896 | 1292 |
| pr_def (name, CodegenThingol.Datatype (vs, co :: cos)) = |
1293 |
let |
|
| 21082 | 1294 |
val tyvars = CodegenThingol.intro_vars (map fst vs) keyword_vars; |
| 20896 | 1295 |
fun pr_co (co, tys) = |
| 22022 | 1296 |
concat ( |
| 20896 | 1297 |
(str o deresolv_here) co |
1298 |
:: map (pr_typ tyvars BR) tys |
|
1299 |
) |
|
1300 |
in |
|
| 22022 | 1301 |
semicolon ( |
| 20896 | 1302 |
str "data" |
1303 |
:: pr_typscheme_expr tyvars (vs, (deresolv_here name, map (ITyVar o fst) vs)) |
|
1304 |
:: str "=" |
|
1305 |
:: pr_co co |
|
1306 |
:: map ((fn p => Pretty.block [str "| ", p]) o pr_co) cos |
|
| 22022 | 1307 |
@ (if deriving_show name then [str "deriving (Read, Show)"] else []) |
1308 |
) |
|
| 21082 | 1309 |
end |
| 20896 | 1310 |
| pr_def (name, CodegenThingol.Class (superclasss, (v, classops))) = |
1311 |
let |
|
| 21082 | 1312 |
val tyvars = CodegenThingol.intro_vars [v] keyword_vars; |
| 20896 | 1313 |
fun pr_classop (classop, ty) = |
| 22022 | 1314 |
semicolon [ |
1315 |
(str o classop_name name) classop, |
|
1316 |
str "::", |
|
| 20896 | 1317 |
pr_typ tyvars NOBR ty |
| 19136 | 1318 |
] |
| 20896 | 1319 |
in |
| 22022 | 1320 |
Pretty.block_enclose ( |
1321 |
Pretty.block [ |
|
1322 |
str "class ", |
|
1323 |
pr_typparms tyvars [(v, superclasss)], |
|
1324 |
str (deresolv_here name ^ " " ^ CodegenThingol.lookup_var tyvars v), |
|
1325 |
str " where {"
|
|
1326 |
], |
|
1327 |
str "};" |
|
1328 |
) (map pr_classop classops) |
|
| 21082 | 1329 |
end |
| 20896 | 1330 |
| pr_def (_, CodegenThingol.Classinst ((class, (tyco, vs)), (_, classop_defs))) = |
1331 |
let |
|
| 21082 | 1332 |
val tyvars = CodegenThingol.intro_vars (map fst vs) keyword_vars; |
| 22022 | 1333 |
fun pr_instdef (classop, t) = |
| 20896 | 1334 |
let |
1335 |
val consts = map_filter |
|
1336 |
(fn c => if (is_some o const_syntax) c |
|
1337 |
then NONE else (SOME o NameSpace.base o deresolv) c) |
|
1338 |
(CodegenThingol.fold_constnames (insert (op =)) t []); |
|
1339 |
val vars = keyword_vars |
|
| 21082 | 1340 |
|> CodegenThingol.intro_vars consts; |
| 20896 | 1341 |
in |
| 22022 | 1342 |
semicolon [ |
| 20896 | 1343 |
(str o classop_name class) classop, |
1344 |
str "=", |
|
1345 |
pr_term vars NOBR t |
|
1346 |
] |
|
| 22022 | 1347 |
end; |
1348 |
in |
|
1349 |
Pretty.block_enclose ( |
|
1350 |
Pretty.block [ |
|
1351 |
str "instance ", |
|
1352 |
pr_typparms tyvars vs, |
|
1353 |
str (class_name class ^ " "), |
|
1354 |
pr_typ tyvars BR (tyco `%% map (ITyVar o fst) vs), |
|
1355 |
str " where {"
|
|
1356 |
], |
|
1357 |
str "};" |
|
1358 |
) (map pr_instdef classop_defs) |
|
| 21082 | 1359 |
end; |
| 20940 | 1360 |
in pr_def def end; |
| 20896 | 1361 |
|
| 21015 | 1362 |
val reserved_haskell = [ |
1363 |
"hiding", "deriving", "where", "case", "of", "infix", "infixl", "infixr", |
|
1364 |
"import", "default", "forall", "let", "in", "class", "qualified", "data", |
|
1365 |
"newtype", "instance", "if", "then", "else", "type", "as", "do", "module" |
|
1366 |
]; |
|
1367 |
||
| 21082 | 1368 |
fun seri_haskell module_prefix destination string_classes reserved_user module_alias module_prolog |
| 21015 | 1369 |
class_syntax tyco_syntax const_syntax code = |
1370 |
let |
|
|
21858
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents:
21837
diff
changeset
|
1371 |
val _ = Option.map File.check destination; |
| 21082 | 1372 |
val empty_names = Name.make_context (reserved_haskell @ reserved_user); |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1373 |
val name_modl = mk_modl_name_tab empty_names module_prefix module_alias code |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1374 |
fun add_def (name, (def, deps : string list)) = |
| 21082 | 1375 |
let |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1376 |
val (modl, base) = dest_name name; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1377 |
fun name_def base = Name.variants [base] #>> the_single; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1378 |
fun add_fun upper (nsp_fun, nsp_typ) = |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1379 |
let |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1380 |
val (base', nsp_fun') = name_def (if upper then first_upper base else base) nsp_fun |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1381 |
in (base', (nsp_fun', nsp_typ)) end; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1382 |
fun add_typ (nsp_fun, nsp_typ) = |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1383 |
let |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1384 |
val (base', nsp_typ') = name_def (first_upper base) nsp_typ |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1385 |
in (base', (nsp_fun, nsp_typ')) end; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1386 |
val add_name = |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1387 |
case def |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1388 |
of CodegenThingol.Bot => pair base |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1389 |
| CodegenThingol.Fun _ => add_fun false |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1390 |
| CodegenThingol.Datatype _ => add_typ |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1391 |
| CodegenThingol.Datatypecons _ => add_fun true |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1392 |
| CodegenThingol.Class _ => add_typ |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1393 |
| CodegenThingol.Classop _ => add_fun false |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1394 |
| CodegenThingol.Classinst _ => pair base; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1395 |
val modlname' = name_modl modl; |
| 21082 | 1396 |
fun add_def base' = |
1397 |
case def |
|
| 21093 | 1398 |
of CodegenThingol.Bot => I |
1399 |
| CodegenThingol.Datatypecons _ => I |
|
| 21082 | 1400 |
cons (name, ((NameSpace.append modlname' base', base'), NONE)) |
1401 |
| CodegenThingol.Classop _ => |
|
1402 |
cons (name, ((NameSpace.append modlname' base', base'), NONE)) |
|
1403 |
| _ => cons (name, ((NameSpace.append modlname' base', base'), SOME def)); |
|
1404 |
in |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1405 |
Symtab.map_default (modlname', ([], ([], (empty_names, empty_names)))) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1406 |
(apfst (fold (insert (op =)) deps)) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1407 |
#> `(fn code => add_name ((snd o snd o the o Symtab.lookup code) modlname')) |
| 21082 | 1408 |
#-> (fn (base', names) => |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1409 |
(Symtab.map_entry modlname' o apsnd) (fn (defs, _) => |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1410 |
(add_def base' defs, names))) |
| 21082 | 1411 |
end; |
1412 |
val code' = |
|
1413 |
fold add_def (AList.make (fn name => (Graph.get_node code name, Graph.imm_succs code name)) |
|
1414 |
(Graph.strong_conn code |> flat)) Symtab.empty; |
|
1415 |
val init_vars = CodegenThingol.make_vars (reserved_haskell @ reserved_user); |
|
1416 |
fun deresolv name = |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1417 |
(fst o fst o the o AList.lookup (op =) ((fst o snd o the |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1418 |
o Symtab.lookup code') ((name_modl o fst o dest_name) name))) name |
| 21285 | 1419 |
handle Option => "(error \"undefined name " ^ name ^ "\")"; |
| 21082 | 1420 |
fun deresolv_here name = |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1421 |
(snd o fst o the o AList.lookup (op =) ((fst o snd o the |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1422 |
o Symtab.lookup code') ((name_modl o fst o dest_name) name))) name |
| 21285 | 1423 |
handle Option => "(error \"undefined name " ^ name ^ "\")"; |
| 21082 | 1424 |
fun deriving_show tyco = |
1425 |
let |
|
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1426 |
fun deriv _ "fun" = false |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1427 |
| deriv tycos tyco = member (op =) tycos tyco orelse |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1428 |
case the_default CodegenThingol.Bot (try (Graph.get_node code) tyco) |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1429 |
of CodegenThingol.Bot => true |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1430 |
| CodegenThingol.Datatype (_, cs) => forall (deriv' (tyco :: tycos)) |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1431 |
(maps snd cs) |
| 21082 | 1432 |
and deriv' tycos (tyco `%% tys) = deriv tycos tyco |
1433 |
andalso forall (deriv' tycos) tys |
|
1434 |
| deriv' _ (ITyVar _) = true |
|
1435 |
in deriv [] tyco end; |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1436 |
fun seri_def qualified = pr_haskell class_syntax tyco_syntax const_syntax init_vars |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1437 |
deresolv_here (if qualified then deresolv else deresolv_here) (if string_classes then deriving_show else K false); |
| 22022 | 1438 |
fun write_module (SOME destination) modlname = |
| 21082 | 1439 |
let |
1440 |
val filename = case modlname |
|
|
21858
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents:
21837
diff
changeset
|
1441 |
of "" => Path.explode "Main.hs" |
|
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents:
21837
diff
changeset
|
1442 |
| _ => (Path.ext "hs" o Path.explode o implode o separate "/" o NameSpace.explode) modlname; |
| 21082 | 1443 |
val pathname = Path.append destination filename; |
1444 |
val _ = File.mkdir (Path.dir pathname); |
|
| 22022 | 1445 |
in File.write pathname end |
1446 |
| write_module NONE _ = writeln; |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1447 |
fun seri_module (modlname', (imports, (defs, _))) = |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1448 |
let |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1449 |
val imports' = |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1450 |
imports |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1451 |
|> map (name_modl o fst o dest_name) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1452 |
|> distinct (op =) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1453 |
|> remove (op =) modlname'; |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1454 |
val qualified = |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1455 |
imports |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1456 |
|> map_filter (try deresolv) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1457 |
|> map NameSpace.base |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1458 |
|> has_duplicates (op =); |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1459 |
val mk_import = str o (if qualified |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1460 |
then prefix "import qualified " |
| 22022 | 1461 |
else prefix "import ") o suffix ";"; |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1462 |
in |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1463 |
Pretty.chunks ( |
| 22022 | 1464 |
str ("module " ^ modlname' ^ " where {")
|
1465 |
:: str "" |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1466 |
:: map mk_import imports' |
| 22022 | 1467 |
@ str "" |
1468 |
:: separate (str "") ((case module_prolog modlname' |
|
1469 |
of SOME prolog => [prolog] |
|
1470 |
| NONE => []) |
|
1471 |
@ map_filter |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1472 |
(fn (name, (_, SOME def)) => SOME (seri_def qualified (name, def)) |
| 22022 | 1473 |
| (_, (_, NONE)) => NONE) defs) |
1474 |
@ str "" |
|
1475 |
@@ str "}" |
|
1476 |
) |
|
1477 |
|> code_output |
|
1478 |
|> write_module destination modlname' |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1479 |
end; |
| 21082 | 1480 |
in Symtab.fold (fn modl => fn () => seri_module modl) code' () end; |
| 21015 | 1481 |
|
| 21082 | 1482 |
val isar_seri_haskell = |
1483 |
parse_args (Scan.option (Args.$$$ "root" -- Args.colon |-- Args.name) |
|
1484 |
-- Scan.optional (Args.$$$ "string_classes" >> K true) false |
|
1485 |
-- (Args.$$$ "-" >> K NONE || Args.name >> SOME) |
|
1486 |
>> (fn ((module_prefix, string_classes), destination) => |
|
|
21858
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents:
21837
diff
changeset
|
1487 |
seri_haskell module_prefix (Option.map Path.explode destination) string_classes)); |
| 21015 | 1488 |
|
1489 |
||
1490 |
(** diagnosis serializer **) |
|
1491 |
||
1492 |
fun seri_diagnosis _ _ _ _ _ code = |
|
1493 |
let |
|
| 21082 | 1494 |
val init_vars = CodegenThingol.make_vars reserved_haskell; |
1495 |
val pr = pr_haskell (K NONE) (K NONE) (K NONE) init_vars I I (K false); |
|
| 21015 | 1496 |
in |
1497 |
[] |
|
| 21093 | 1498 |
|> Graph.fold (fn (name, (def, _)) => case try pr (name, def) of SOME p => cons p | NONE => I) code |
| 22022 | 1499 |
|> Pretty.chunks2 |
1500 |
|> code_output |
|
1501 |
|> writeln |
|
| 21015 | 1502 |
end; |
1503 |
||
1504 |
||
| 20896 | 1505 |
|
1506 |
(** theory data **) |
|
1507 |
||
1508 |
datatype syntax_expr = SyntaxExpr of {
|
|
1509 |
class: ((string * (string -> string option)) * serial) Symtab.table, |
|
1510 |
inst: unit Symtab.table, |
|
1511 |
tyco: (itype pretty_syntax * serial) Symtab.table, |
|
1512 |
const: (iterm pretty_syntax * serial) Symtab.table |
|
1513 |
}; |
|
| 18702 | 1514 |
|
| 20896 | 1515 |
fun mk_syntax_expr ((class, inst), (tyco, const)) = |
1516 |
SyntaxExpr { class = class, inst = inst, tyco = tyco, const = const };
|
|
1517 |
fun map_syntax_expr f (SyntaxExpr { class, inst, tyco, const }) =
|
|
1518 |
mk_syntax_expr (f ((class, inst), (tyco, const))); |
|
1519 |
fun merge_syntax_expr (SyntaxExpr { class = class1, inst = inst1, tyco = tyco1, const = const1 },
|
|
1520 |
SyntaxExpr { class = class2, inst = inst2, tyco = tyco2, const = const2 }) =
|
|
1521 |
mk_syntax_expr ( |
|
1522 |
(Symtab.merge (eq_snd (op =)) (class1, class2), |
|
1523 |
Symtab.merge (op =) (inst1, inst2)), |
|
1524 |
(Symtab.merge (eq_snd (op =)) (tyco1, tyco2), |
|
1525 |
Symtab.merge (eq_snd (op =)) (const1, const2)) |
|
1526 |
); |
|
1527 |
||
1528 |
datatype syntax_modl = SyntaxModl of {
|
|
| 21015 | 1529 |
alias: string Symtab.table, |
| 20896 | 1530 |
prolog: Pretty.T Symtab.table |
1531 |
}; |
|
| 20699 | 1532 |
|
| 21015 | 1533 |
fun mk_syntax_modl (alias, prolog) = |
1534 |
SyntaxModl { alias = alias, prolog = prolog };
|
|
1535 |
fun map_syntax_modl f (SyntaxModl { alias, prolog }) =
|
|
1536 |
mk_syntax_modl (f (alias, prolog)); |
|
1537 |
fun merge_syntax_modl (SyntaxModl { alias = alias1, prolog = prolog1 },
|
|
1538 |
SyntaxModl { alias = alias2, prolog = prolog2 }) =
|
|
| 20896 | 1539 |
mk_syntax_modl ( |
| 21015 | 1540 |
Symtab.merge (op =) (alias1, alias2), |
| 20896 | 1541 |
Symtab.merge (op =) (prolog1, prolog2) |
1542 |
); |
|
| 20699 | 1543 |
|
| 21015 | 1544 |
type serializer = Args.T list |
1545 |
-> string list |
|
1546 |
-> (string -> string option) |
|
1547 |
-> (string -> Pretty.T option) |
|
1548 |
-> (string -> (string * (string -> string option)) option) |
|
1549 |
-> (string -> (int * (fixity -> (fixity -> itype -> Pretty.T) -> itype list -> Pretty.T)) option) |
|
1550 |
-> (string -> (int * (fixity -> (fixity -> iterm -> Pretty.T) -> iterm list -> Pretty.T)) option) |
|
1551 |
-> CodegenThingol.code -> unit; |
|
| 20896 | 1552 |
|
1553 |
datatype target = Target of {
|
|
1554 |
serial: serial, |
|
1555 |
serializer: serializer, |
|
1556 |
syntax_expr: syntax_expr, |
|
| 21015 | 1557 |
syntax_modl: syntax_modl, |
1558 |
reserved: string list |
|
| 20896 | 1559 |
}; |
| 20699 | 1560 |
|
| 21015 | 1561 |
fun mk_target (serial, ((serializer, reserved), (syntax_expr, syntax_modl))) = |
1562 |
Target { serial = serial, reserved = reserved, serializer = serializer, syntax_expr = syntax_expr, syntax_modl = syntax_modl };
|
|
1563 |
fun map_target f ( Target { serial, serializer, reserved, syntax_expr, syntax_modl } ) =
|
|
1564 |
mk_target (f (serial, ((serializer, reserved), (syntax_expr, syntax_modl)))); |
|
1565 |
fun merge_target target (Target { serial = serial1, serializer = serializer, reserved = reserved1,
|
|
1566 |
syntax_expr = syntax_expr1, syntax_modl = syntax_modl1 }, |
|
1567 |
Target { serial = serial2, serializer = _, reserved = reserved2,
|
|
1568 |
syntax_expr = syntax_expr2, syntax_modl = syntax_modl2 }) = |
|
| 20896 | 1569 |
if serial1 = serial2 then |
| 21015 | 1570 |
mk_target (serial1, ((serializer, merge (op =) (reserved1, reserved2)), |
| 20896 | 1571 |
(merge_syntax_expr (syntax_expr1, syntax_expr2), |
1572 |
merge_syntax_modl (syntax_modl1, syntax_modl2)) |
|
1573 |
)) |
|
1574 |
else |
|
1575 |
error ("Incompatible serializers: " ^ quote target);
|
|
1576 |
||
1577 |
structure CodegenSerializerData = TheoryDataFun |
|
1578 |
(struct |
|
1579 |
val name = "Pure/codegen_serializer"; |
|
1580 |
type T = target Symtab.table; |
|
1581 |
val empty = Symtab.empty; |
|
1582 |
val copy = I; |
|
1583 |
val extend = I; |
|
1584 |
fun merge _ = Symtab.join merge_target; |
|
1585 |
fun print _ _ = (); |
|
1586 |
end); |
|
1587 |
||
1588 |
fun the_serializer (Target { serializer, ... }) = serializer;
|
|
| 21015 | 1589 |
fun the_reserved (Target { reserved, ... }) = reserved;
|
| 20896 | 1590 |
fun the_syntax_expr (Target { syntax_expr = SyntaxExpr x, ... }) = x;
|
1591 |
fun the_syntax_modl (Target { syntax_modl = SyntaxModl x, ... }) = x;
|
|
1592 |
||
1593 |
fun add_serializer (target, seri) thy = |
|
| 18702 | 1594 |
let |
| 20896 | 1595 |
val _ = case Symtab.lookup (CodegenSerializerData.get thy) target |
1596 |
of SOME _ => warning ("overwriting existing serializer " ^ quote target)
|
|
1597 |
| NONE => (); |
|
| 20699 | 1598 |
in |
| 20896 | 1599 |
thy |
1600 |
|> (CodegenSerializerData.map oo Symtab.map_default) |
|
| 21015 | 1601 |
(target, mk_target (serial (), ((seri, []), |
| 20896 | 1602 |
(mk_syntax_expr ((Symtab.empty, Symtab.empty), (Symtab.empty, Symtab.empty)), |
1603 |
mk_syntax_modl (Symtab.empty, Symtab.empty))))) |
|
| 21015 | 1604 |
(map_target (fn (serial, ((_, keywords), syntax)) => (serial, ((seri, keywords), syntax)))) |
| 20699 | 1605 |
end; |
1606 |
||
| 21015 | 1607 |
fun map_seri_data target f thy = |
1608 |
let |
|
1609 |
val _ = if is_some (Symtab.lookup (CodegenSerializerData.get thy) target) |
|
1610 |
then () |
|
1611 |
else error ("Unknown code target language: " ^ quote target);
|
|
1612 |
in |
|
1613 |
thy |
|
1614 |
|> (CodegenSerializerData.map o Symtab.map_entry target o map_target) f |
|
1615 |
end; |
|
1616 |
||
1617 |
val target_diag = "diag"; |
|
1618 |
||
| 20896 | 1619 |
val _ = Context.add_setup ( |
1620 |
CodegenSerializerData.init |
|
| 21130 | 1621 |
#> add_serializer ("SML", isar_seri_sml)
|
| 21837 | 1622 |
#> add_serializer ("OCaml", isar_seri_ocaml)
|
| 21082 | 1623 |
#> add_serializer ("Haskell", isar_seri_haskell)
|
| 21015 | 1624 |
#> add_serializer (target_diag, (fn _ => fn _ => seri_diagnosis)) |
| 20896 | 1625 |
); |
1626 |
||
| 22007 | 1627 |
fun get_serializer thy target args = fn cs => |
| 20896 | 1628 |
let |
1629 |
val data = case Symtab.lookup (CodegenSerializerData.get thy) target |
|
1630 |
of SOME data => data |
|
1631 |
| NONE => error ("Unknown code target language: " ^ quote target);
|
|
1632 |
val seri = the_serializer data; |
|
| 21015 | 1633 |
val reserved = the_reserved data; |
1634 |
val { alias, prolog } = the_syntax_modl data;
|
|
| 20896 | 1635 |
val { class, inst, tyco, const } = the_syntax_expr data;
|
1636 |
fun fun_of sys = (Option.map fst oo Symtab.lookup) sys; |
|
| 21015 | 1637 |
val project = if target = target_diag then I |
1638 |
else CodegenThingol.project_code |
|
1639 |
(Symtab.keys class @ Symtab.keys inst @ Symtab.keys tyco @ Symtab.keys const) cs; |
|
| 20896 | 1640 |
in |
| 21015 | 1641 |
project #> seri args reserved (Symtab.lookup alias) (Symtab.lookup prolog) |
1642 |
(fun_of class) (fun_of tyco) (fun_of const) |
|
| 20896 | 1643 |
end; |
1644 |
||
| 21162 | 1645 |
val eval_verbose = ref false; |
1646 |
||
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1647 |
fun eval_term thy code ((ref_name, reff), t) = |
| 21162 | 1648 |
let |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1649 |
val val_name = "eval.EVAL.EVAL"; |
| 21162 | 1650 |
val val_name' = "ROOT.eval.EVAL"; |
1651 |
val data = (the o Symtab.lookup (CodegenSerializerData.get thy)) "SML" |
|
1652 |
val reserved = the_reserved data; |
|
1653 |
val { alias, prolog } = the_syntax_modl data;
|
|
1654 |
val { class, inst, tyco, const } = the_syntax_expr data;
|
|
1655 |
fun fun_of sys = (Option.map fst oo Symtab.lookup) sys; |
|
1656 |
fun eval p = ( |
|
1657 |
reff := NONE; |
|
1658 |
if !eval_verbose then Pretty.writeln p else (); |
|
1659 |
use_text Output.ml_output (!eval_verbose) |
|
1660 |
((Pretty.output o Pretty.chunks) [p, |
|
1661 |
str ("val _ = (" ^ ref_name ^ " := SOME " ^ val_name' ^ ")")
|
|
1662 |
]); |
|
1663 |
case !reff |
|
1664 |
of NONE => error ("Could not retrieve value of ML reference " ^ quote ref_name
|
|
1665 |
^ " (reference probably has been shadowed)") |
|
1666 |
| SOME value => value |
|
1667 |
); |
|
1668 |
in |
|
1669 |
code |
|
1670 |
|> CodegenThingol.add_eval_def (val_name, t) |
|
1671 |
|> CodegenThingol.project_code |
|
1672 |
(Symtab.keys class @ Symtab.keys inst @ Symtab.keys tyco @ Symtab.keys const) |
|
1673 |
(SOME [val_name]) |
|
| 21837 | 1674 |
|> seri_ml pr_sml pr_sml_modl I reserved (Symtab.lookup alias) (Symtab.lookup prolog) |
| 21162 | 1675 |
(fun_of class) (fun_of tyco) (fun_of const) |
1676 |
|> eval |
|
1677 |
end; |
|
1678 |
||
| 21082 | 1679 |
fun assert_serializer thy target = |
1680 |
case Symtab.lookup (CodegenSerializerData.get thy) target |
|
1681 |
of SOME data => target |
|
1682 |
| NONE => error ("Unknown code target language: " ^ quote target);
|
|
1683 |
||
| 20896 | 1684 |
fun has_serialization f thy targets name = |
1685 |
forall ( |
|
1686 |
is_some o (fn tab => Symtab.lookup tab name) o f o the_syntax_expr o the |
|
| 22007 | 1687 |
o Symtab.lookup (CodegenSerializerData.get thy) |
| 20896 | 1688 |
) targets; |
1689 |
||
1690 |
val tyco_has_serialization = has_serialization #tyco; |
|
1691 |
val const_has_serialization = has_serialization #const; |
|
1692 |
||
1693 |
||
| 20699 | 1694 |
|
| 20931 | 1695 |
(** ML and toplevel interface **) |
| 20699 | 1696 |
|
1697 |
local |
|
1698 |
||
| 21015 | 1699 |
fun map_syntax_exprs target = |
1700 |
map_seri_data target o (apsnd o apsnd o apfst o map_syntax_expr); |
|
1701 |
fun map_syntax_modls target = |
|
1702 |
map_seri_data target o (apsnd o apsnd o apsnd o map_syntax_modl); |
|
1703 |
fun map_reserveds target = |
|
1704 |
map_seri_data target o (apsnd o apfst o apsnd); |
|
| 20896 | 1705 |
|
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1706 |
fun gen_add_syntax_class prep_class prep_const target raw_class raw_syn thy = |
| 20699 | 1707 |
let |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1708 |
val cls = prep_class thy raw_class; |
| 20699 | 1709 |
val class = CodegenNames.class thy cls; |
| 20896 | 1710 |
fun mk_classop (const as (c, _)) = case AxClass.class_of_param thy c |
1711 |
of SOME class' => if cls = class' then CodegenNames.const thy const |
|
| 20699 | 1712 |
else error ("Not a class operation for class " ^ quote class ^ ": " ^ quote c)
|
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1713 |
| NONE => error ("Not a class operation: " ^ quote c);
|
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1714 |
fun mk_syntax_ops raw_ops = AList.lookup (op =) |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1715 |
((map o apfst) (mk_classop o prep_const thy) raw_ops); |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1716 |
in case raw_syn |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1717 |
of SOME (syntax, raw_ops) => |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1718 |
thy |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1719 |
|> (map_syntax_exprs target o apfst o apfst) |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1720 |
(Symtab.update (class, ((syntax, mk_syntax_ops raw_ops), serial ()))) |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1721 |
| NONE => |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1722 |
thy |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1723 |
|> (map_syntax_exprs target o apfst o apfst) |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1724 |
(Symtab.delete_safe class) |
| 20699 | 1725 |
end; |
1726 |
||
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1727 |
fun gen_add_syntax_inst prep_class prep_tyco target (raw_tyco, raw_class) add_del thy = |
| 20699 | 1728 |
let |
1729 |
val inst = CodegenNames.instance thy (prep_class thy raw_class, prep_tyco thy raw_tyco); |
|
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1730 |
in if add_del then |
| 20699 | 1731 |
thy |
| 20896 | 1732 |
|> (map_syntax_exprs target o apfst o apsnd) |
1733 |
(Symtab.update (inst, ())) |
|
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1734 |
else |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1735 |
thy |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1736 |
|> (map_syntax_exprs target o apfst o apsnd) |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1737 |
(Symtab.delete_safe inst) |
| 20699 | 1738 |
end; |
1739 |
||
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1740 |
fun gen_add_syntax_tyco prep_tyco target raw_tyco raw_syn thy = |
| 20699 | 1741 |
let |
| 21015 | 1742 |
val tyco = prep_tyco thy raw_tyco; |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1743 |
val tyco' = if tyco = "fun" then "fun" else CodegenNames.tyco thy tyco; |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1744 |
fun check_args (syntax as (n, _)) = if n <> Sign.arity_number thy tyco |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1745 |
then error ("Number of arguments mismatch in syntax for type constructor " ^ quote tyco)
|
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1746 |
else syntax |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1747 |
in case raw_syn |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1748 |
of SOME syntax => |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1749 |
thy |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1750 |
|> (map_syntax_exprs target o apsnd o apfst) |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1751 |
(Symtab.update (tyco', (check_args syntax, serial ()))) |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1752 |
| NONE => |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1753 |
thy |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1754 |
|> (map_syntax_exprs target o apsnd o apfst) |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1755 |
(Symtab.delete_safe tyco') |
| 20699 | 1756 |
end; |
1757 |
||
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1758 |
fun gen_add_syntax_const prep_const target raw_c raw_syn thy = |
| 20699 | 1759 |
let |
| 21015 | 1760 |
val c = prep_const thy raw_c; |
1761 |
val c' = CodegenNames.const thy c; |
|
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1762 |
fun check_args (syntax as (n, _)) = if n > (length o fst o strip_type o Sign.the_const_type thy o fst) c |
| 21015 | 1763 |
then error ("Too many arguments in syntax for constant " ^ (quote o fst) c)
|
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1764 |
else syntax; |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1765 |
in case raw_syn |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1766 |
of SOME syntax => |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1767 |
thy |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1768 |
|> (map_syntax_exprs target o apsnd o apsnd) |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1769 |
(Symtab.update (c', (check_args syntax, serial ()))) |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1770 |
| NONE => |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1771 |
thy |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1772 |
|> (map_syntax_exprs target o apsnd o apsnd) |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1773 |
(Symtab.delete_safe c') |
| 20699 | 1774 |
end; |
1775 |
||
| 21285 | 1776 |
(*fun gen_add_syntax_monad prep_tyco target raw_tyco monad_tyco thy = |
1777 |
let |
|
1778 |
val _ = if |
|
1779 |
in |
|
1780 |
thy |
|
1781 |
end;*) |
|
1782 |
||
|
21463
42dd50268c8b
completed class parameter handling in axclass.ML
haftmann
parents:
21389
diff
changeset
|
1783 |
fun read_class thy raw_class = |
|
42dd50268c8b
completed class parameter handling in axclass.ML
haftmann
parents:
21389
diff
changeset
|
1784 |
let |
|
42dd50268c8b
completed class parameter handling in axclass.ML
haftmann
parents:
21389
diff
changeset
|
1785 |
val class = Sign.intern_class thy raw_class; |
|
42dd50268c8b
completed class parameter handling in axclass.ML
haftmann
parents:
21389
diff
changeset
|
1786 |
val _ = AxClass.get_definition thy class; |
|
42dd50268c8b
completed class parameter handling in axclass.ML
haftmann
parents:
21389
diff
changeset
|
1787 |
in class end; |
|
42dd50268c8b
completed class parameter handling in axclass.ML
haftmann
parents:
21389
diff
changeset
|
1788 |
|
| 20699 | 1789 |
fun read_type thy raw_tyco = |
1790 |
let |
|
1791 |
val tyco = Sign.intern_type thy raw_tyco; |
|
1792 |
val _ = if Sign.declared_tyname thy tyco then () |
|
1793 |
else error ("No such type constructor: " ^ quote raw_tyco);
|
|
1794 |
in tyco end; |
|
1795 |
||
| 21837 | 1796 |
fun idfs_of_const_names thy c = |
| 20699 | 1797 |
let |
| 21837 | 1798 |
val c' = (c, Sign.the_const_type thy c); |
1799 |
val c'' = CodegenConsts.norm_of_typ thy c'; |
|
1800 |
in (c'', CodegenNames.const thy c'') end; |
|
| 20699 | 1801 |
|
|
21463
42dd50268c8b
completed class parameter handling in axclass.ML
haftmann
parents:
21389
diff
changeset
|
1802 |
val add_syntax_class = gen_add_syntax_class read_class CodegenConsts.read_const; |
|
42dd50268c8b
completed class parameter handling in axclass.ML
haftmann
parents:
21389
diff
changeset
|
1803 |
val add_syntax_inst = gen_add_syntax_inst read_class read_type; |
| 21015 | 1804 |
val add_syntax_tyco = gen_add_syntax_tyco read_type; |
1805 |
val add_syntax_const = gen_add_syntax_const CodegenConsts.read_const; |
|
1806 |
||
1807 |
fun add_reserved target = |
|
1808 |
map_reserveds target o insert (op =); |
|
1809 |
||
1810 |
fun add_modl_alias target = |
|
1811 |
map_syntax_modls target o apfst o Symtab.update o apsnd CodegenNames.check_modulename; |
|
1812 |
||
1813 |
fun add_modl_prolog target = |
|
1814 |
map_syntax_modls target o apsnd o |
|
1815 |
(fn (modl, NONE) => Symtab.delete modl | (modl, SOME prolog) => |
|
1816 |
Symtab.update (modl, Pretty.str prolog)); |
|
| 20931 | 1817 |
|
1818 |
fun zip_list (x::xs) f g = |
|
| 21015 | 1819 |
f |
1820 |
#-> (fn y => |
|
1821 |
fold_map (fn x => g |-- f >> pair x) xs |
|
| 20931 | 1822 |
#-> (fn xys => pair ((x, y) :: xys))); |
| 20699 | 1823 |
|
| 20931 | 1824 |
structure P = OuterParse |
1825 |
and K = OuterKeyword |
|
| 20699 | 1826 |
|
| 20931 | 1827 |
fun parse_multi_syntax parse_thing parse_syntax = |
1828 |
P.and_list1 parse_thing |
|
| 21015 | 1829 |
#-> (fn things => Scan.repeat1 (P.$$$ "(" |-- P.name --
|
1830 |
(zip_list things parse_syntax (P.$$$ "and")) --| P.$$$ ")")); |
|
| 20699 | 1831 |
|
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1832 |
val (infixK, infixlK, infixrK) = ("infix", "infixl", "infixr");
|
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1833 |
|
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1834 |
fun parse_syntax xs = |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1835 |
Scan.option (( |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1836 |
((P.$$$ infixK >> K X) |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1837 |
|| (P.$$$ infixlK >> K L) |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1838 |
|| (P.$$$ infixrK >> K R)) |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1839 |
-- P.nat >> parse_infix |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1840 |
|| Scan.succeed parse_mixfix) |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1841 |
-- P.string |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1842 |
>> (fn (parse, s) => parse s)) xs; |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1843 |
|
| 21285 | 1844 |
val (code_classK, code_instanceK, code_typeK, code_constK, code_monadK, |
| 21015 | 1845 |
code_reservedK, code_modulenameK, code_moduleprologK) = |
| 21285 | 1846 |
("code_class", "code_instance", "code_type", "code_const", "code_monad",
|
| 21015 | 1847 |
"code_reserved", "code_modulename", "code_moduleprolog"); |
| 20699 | 1848 |
|
| 21015 | 1849 |
in |
1850 |
||
1851 |
fun add_pretty_list target nill cons mk_list mk_char_string target_cons thy = |
|
| 20699 | 1852 |
let |
| 21837 | 1853 |
val (_, nil'') = idfs_of_const_names thy nill; |
1854 |
val (cons', cons'') = idfs_of_const_names thy cons; |
|
| 21015 | 1855 |
val pr = pretty_list nil'' cons'' mk_list mk_char_string target_cons; |
| 20699 | 1856 |
in |
| 21015 | 1857 |
thy |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1858 |
|> gen_add_syntax_const (K I) target cons' (SOME pr) |
| 20699 | 1859 |
end; |
1860 |
||
| 21015 | 1861 |
fun add_pretty_ml_string target nill cons str mk_char mk_string target_implode thy = |
| 20699 | 1862 |
let |
| 21837 | 1863 |
val (_, nil'') = idfs_of_const_names thy nill; |
1864 |
val (_, cons'') = idfs_of_const_names thy cons; |
|
1865 |
val (str', _) = idfs_of_const_names thy str; |
|
| 21015 | 1866 |
val pr = pretty_ml_string nil'' cons'' mk_char mk_string target_implode; |
| 20699 | 1867 |
in |
| 21015 | 1868 |
thy |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1869 |
|> gen_add_syntax_const (K I) target str' (SOME pr) |
| 20699 | 1870 |
end; |
1871 |
||
| 21015 | 1872 |
fun add_undefined target undef target_undefined thy = |
1873 |
let |
|
| 21837 | 1874 |
val (undef', _) = idfs_of_const_names thy undef; |
1875 |
fun pr _ _ _ = str target_undefined; |
|
| 21015 | 1876 |
in |
1877 |
thy |
|
| 21837 | 1878 |
|> gen_add_syntax_const (K I) target undef' (SOME (~1, pr)) |
1879 |
end; |
|
1880 |
||
1881 |
fun add_pretty_imperative_monad_bind target bind thy = |
|
1882 |
let |
|
1883 |
val (bind', bind'') = idfs_of_const_names thy bind; |
|
1884 |
val pr = pretty_imperative_monad_bind bind'' |
|
1885 |
in |
|
1886 |
thy |
|
1887 |
|> gen_add_syntax_const (K I) target bind' (SOME pr) |
|
| 21015 | 1888 |
end; |
| 20931 | 1889 |
|
1890 |
val code_classP = |
|
1891 |
OuterSyntax.command code_classK "define code syntax for class" K.thy_decl ( |
|
1892 |
parse_multi_syntax P.xname |
|
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1893 |
(Scan.option (P.string -- Scan.optional (P.$$$ "where" |-- Scan.repeat1 |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1894 |
(P.term --| (P.$$$ "\\<equiv>" || P.$$$ "==") -- P.string)) [])) |
| 20931 | 1895 |
>> (Toplevel.theory oo fold) (fn (target, syns) => |
1896 |
fold (fn (raw_class, syn) => add_syntax_class target raw_class syn) syns) |
|
1897 |
); |
|
1898 |
||
1899 |
val code_instanceP = |
|
1900 |
OuterSyntax.command code_instanceK "define code syntax for instance" K.thy_decl ( |
|
1901 |
parse_multi_syntax (P.xname --| P.$$$ "::" -- P.xname) |
|
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1902 |
((P.minus >> K true) || Scan.succeed false) |
| 20931 | 1903 |
>> (Toplevel.theory oo fold) (fn (target, syns) => |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1904 |
fold (fn (raw_inst, add_del) => add_syntax_inst target raw_inst add_del) syns) |
| 20931 | 1905 |
); |
1906 |
||
1907 |
val code_typeP = |
|
1908 |
OuterSyntax.command code_typeK "define code syntax for type constructor" K.thy_decl ( |
|
| 21015 | 1909 |
parse_multi_syntax P.xname parse_syntax |
1910 |
>> (Toplevel.theory oo fold) (fn (target, syns) => |
|
1911 |
fold (fn (raw_tyco, syn) => add_syntax_tyco target raw_tyco syn) syns) |
|
| 20931 | 1912 |
); |
1913 |
||
1914 |
val code_constP = |
|
1915 |
OuterSyntax.command code_constK "define code syntax for constant" K.thy_decl ( |
|
| 21015 | 1916 |
parse_multi_syntax P.term parse_syntax |
1917 |
>> (Toplevel.theory oo fold) (fn (target, syns) => |
|
1918 |
fold (fn (raw_const, syn) => add_syntax_const target raw_const syn) syns) |
|
| 20931 | 1919 |
); |
1920 |
||
| 21285 | 1921 |
(*val code_monadP = |
1922 |
OuterSyntax.command code_typeK "define code syntax for open state monads" K.thy_decl ( |
|
1923 |
parse_multi_syntax P.xname parse_syntax |
|
1924 |
>> (Toplevel.theory oo fold) (fn (target, syns) => |
|
1925 |
fold (fn (raw_tyco, syn) => add_syntax_monad target raw_tyco syn) syns) |
|
1926 |
);*) |
|
1927 |
||
| 21015 | 1928 |
val code_reservedP = |
1929 |
OuterSyntax.command code_reservedK "declare words as reserved for target language" K.thy_decl ( |
|
1930 |
P.name -- Scan.repeat1 P.name |
|
1931 |
>> (fn (target, reserveds) => (Toplevel.theory o fold (add_reserved target)) reserveds) |
|
1932 |
) |
|
| 20699 | 1933 |
|
| 21015 | 1934 |
val code_modulenameP = |
| 21082 | 1935 |
OuterSyntax.command code_modulenameK "alias module to other name" K.thy_decl ( |
| 21015 | 1936 |
P.name -- Scan.repeat1 (P.name -- P.name) |
1937 |
>> (fn (target, modlnames) => (Toplevel.theory o fold (add_modl_alias target)) modlnames) |
|
1938 |
) |
|
| 20699 | 1939 |
|
| 21015 | 1940 |
val code_moduleprologP = |
| 21082 | 1941 |
OuterSyntax.command code_moduleprologK "add prolog to module" K.thy_decl ( |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1942 |
P.name -- Scan.repeat1 (P.name -- (P.text >> (fn "-" => NONE | s => SOME s))) |
| 21015 | 1943 |
>> (fn (target, prologs) => (Toplevel.theory o fold (add_modl_prolog target)) prologs) |
1944 |
) |
|
1945 |
||
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1946 |
val _ = OuterSyntax.add_keywords [infixK, infixlK, infixrK]; |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1947 |
|
| 21015 | 1948 |
val _ = OuterSyntax.add_parsers [code_classP, code_instanceP, code_typeP, code_constP, |
1949 |
code_reservedP, code_modulenameP, code_moduleprologP]; |
|
| 20699 | 1950 |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1951 |
(*including serializer defaults*) |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1952 |
val _ = Context.add_setup ( |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1953 |
gen_add_syntax_tyco (K I) "SML" "fun" (SOME (2, fn fxy => fn pr_typ => fn [ty1, ty2] => |
|
21463
42dd50268c8b
completed class parameter handling in axclass.ML
haftmann
parents:
21389
diff
changeset
|
1954 |
(gen_brackify (case fxy of NOBR => false | _ => eval_fxy (INFX (1, R)) fxy) o Pretty.breaks) [ |
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1955 |
pr_typ (INFX (1, X)) ty1, |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1956 |
str "->", |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1957 |
pr_typ (INFX (1, R)) ty2 |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1958 |
])) |
| 21837 | 1959 |
#> gen_add_syntax_tyco (K I) "OCaml" "fun" (SOME (2, fn fxy => fn pr_typ => fn [ty1, ty2] => |
1960 |
(gen_brackify (case fxy of NOBR => false | _ => eval_fxy (INFX (1, R)) fxy) o Pretty.breaks) [ |
|
1961 |
pr_typ (INFX (1, X)) ty1, |
|
1962 |
str "->", |
|
1963 |
pr_typ (INFX (1, R)) ty2 |
|
1964 |
])) |
|
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1965 |
#> gen_add_syntax_tyco (K I) "Haskell" "fun" (SOME (2, fn fxy => fn pr_typ => fn [ty1, ty2] => |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1966 |
brackify_infix (1, R) fxy [ |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1967 |
pr_typ (INFX (1, X)) ty1, |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1968 |
str "->", |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1969 |
pr_typ (INFX (1, R)) ty2 |
|
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
1970 |
])) |
| 21952 | 1971 |
(*IntInt resp. Big_int are added later when code extraction for numerals is set up*) |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1972 |
#> add_reserved "SML" "o" (*dictionary projections use it already*) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1973 |
#> fold (add_reserved "Haskell") [ |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1974 |
"Prelude", "Main", "Bool", "Maybe", "Either", "Ordering", "Char", "String", "Int", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1975 |
"Integer", "Float", "Double", "Rational", "IO", "Eq", "Ord", "Enum", "Bounded", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1976 |
"Num", "Real", "Integral", "Fractional", "Floating", "RealFloat", "Monad", "Functor", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1977 |
"AlreadyExists", "ArithException", "ArrayException", "AssertionFailed", "AsyncException", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1978 |
"BlockedOnDeadMVar", "Deadlock", "Denormal", "DivideByZero", "DotNetException", "DynException", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1979 |
"Dynamic", "EOF", "EQ", "EmptyRec", "ErrorCall", "ExitException", "ExitFailure", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1980 |
"ExitSuccess", "False", "GT", "HeapOverflow", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1981 |
"IO", "IOError", "IOException", "IllegalOperation", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1982 |
"IndexOutOfBounds", "Just", "Key", "LT", "Left", "LossOfPrecision", "NoMethodError", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1983 |
"NoSuchThing", "NonTermination", "Nothing", "Obj", "OtherError", "Overflow", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1984 |
"PatternMatchFail", "PermissionDenied", "ProtocolError", "RecConError", "RecSelError", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1985 |
"RecUpdError", "ResourceBusy", "ResourceExhausted", "Right", "StackOverflow", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1986 |
"ThreadKilled", "True", "TyCon", "TypeRep", "UndefinedElement", "Underflow", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1987 |
"UnsupportedOperation", "UserError", "abs", "absReal", "acos", "acosh", "all", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1988 |
"and", "any", "appendFile", "asTypeOf", "asciiTab", "asin", "asinh", "atan", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1989 |
"atan2", "atanh", "basicIORun", "blockIO", "boundedEnumFrom", "boundedEnumFromThen", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1990 |
"boundedEnumFromThenTo", "boundedEnumFromTo", "boundedPred", "boundedSucc", "break", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1991 |
"catch", "catchException", "ceiling", "compare", "concat", "concatMap", "const", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1992 |
"cos", "cosh", "curry", "cycle", "decodeFloat", "denominator", "div", "divMod", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1993 |
"doubleToRatio", "doubleToRational", "drop", "dropWhile", "either", "elem", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1994 |
"emptyRec", "encodeFloat", "enumFrom", "enumFromThen", "enumFromThenTo", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1995 |
"enumFromTo", "error", "even", "exp", "exponent", "fail", "filter", "flip", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1996 |
"floatDigits", "floatProperFraction", "floatRadix", "floatRange", "floatToRational", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1997 |
"floor", "fmap", "foldl", "foldl'", "foldl1", "foldr", "foldr1", "fromDouble", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1998 |
"fromEnum", "fromEnum_0", "fromInt", "fromInteger", "fromIntegral", "fromObj", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
1999 |
"fromRational", "fst", "gcd", "getChar", "getContents", "getLine", "head", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
2000 |
"id", "inRange", "index", "init", "intToRatio", "interact", "ioError", "isAlpha", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
2001 |
"isAlphaNum", "isDenormalized", "isDigit", "isHexDigit", "isIEEE", "isInfinite", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
2002 |
"isLower", "isNaN", "isNegativeZero", "isOctDigit", "isSpace", "isUpper", "iterate", "iterate'", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
2003 |
"last", "lcm", "length", "lex", "lexDigits", "lexLitChar", "lexmatch", "lines", "log", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
2004 |
"logBase", "lookup", "loop", "map", "mapM", "mapM_", "max", "maxBound", "maximum", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
2005 |
"maybe", "min", "minBound", "minimum", "mod", "negate", "nonnull", "not", "notElem", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
2006 |
"null", "numerator", "numericEnumFrom", "numericEnumFromThen", "numericEnumFromThenTo", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
2007 |
"numericEnumFromTo", "odd", "or", "otherwise", "pi", "pred", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
2008 |
"print", "product", "properFraction", "protectEsc", "putChar", "putStr", "putStrLn", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
2009 |
"quot", "quotRem", "range", "rangeSize", "rationalToDouble", "rationalToFloat", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
2010 |
"rationalToRealFloat", "read", "readDec", "readField", "readFieldName", "readFile", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
2011 |
"readFloat", "readHex", "readIO", "readInt", "readList", "readLitChar", "readLn", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
2012 |
"readOct", "readParen", "readSigned", "reads", "readsPrec", "realFloatToRational", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
2013 |
"realToFrac", "recip", "reduce", "rem", "repeat", "replicate", "return", "reverse", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
2014 |
"round", "scaleFloat", "scanl", "scanl1", "scanr", "scanr1", "seq", "sequence", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
2015 |
"sequence_", "show", "showChar", "showException", "showField", "showList", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
2016 |
"showLitChar", "showParen", "showString", "shows", "showsPrec", "significand", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
2017 |
"signum", "signumReal", "sin", "sinh", "snd", "span", "splitAt", "sqrt", "subtract", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
2018 |
"succ", "sum", "tail", "take", "takeWhile", "takeWhile1", "tan", "tanh", "threadToIOResult", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
2019 |
"throw", "toEnum", "toInt", "toInteger", "toObj", "toRational", "truncate", "uncurry", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
2020 |
"undefined", "unlines", "unsafeCoerce", "unsafeIndex", "unsafeRangeSize", "until", "unwords", |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
2021 |
"unzip", "unzip3", "userError", "words", "writeFile", "zip", "zip3", "zipWith", "zipWith3" |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
2022 |
] (*due to weird handling of ':', we can't do anything else than to import *all* prelude symbols*) |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21895
diff
changeset
|
2023 |
|
|
21122
b1fdd08e0ea3
new serialization syntax; experimental extensions
haftmann
parents:
21094
diff
changeset
|
2024 |
) |
| 21082 | 2025 |
|
| 20699 | 2026 |
end; (*local*) |
| 18702 | 2027 |
|
| 21093 | 2028 |
end; (*struct*) |