| author | traytel | 
| Sun, 09 Mar 2014 21:40:41 +0100 | |
| changeset 56012 | 158dc03db8be | 
| parent 55776 | 7dd1971b39c1 | 
| child 56812 | baef1c110f12 | 
| permissions | -rw-r--r-- | 
| 37745 
6315b6426200
checking generated code for various target languages
 haftmann parents: 
37449diff
changeset | 1 | (* Title: Tools/Code/code_ml.ML | 
| 28054 | 2 | Author: Florian Haftmann, TU Muenchen | 
| 3 | ||
| 4 | Serializer for SML and OCaml. | |
| 5 | *) | |
| 6 | ||
| 7 | signature CODE_ML = | |
| 8 | sig | |
| 34028 
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
 haftmann parents: 
33992diff
changeset | 9 | val target_SML: string | 
| 37745 
6315b6426200
checking generated code for various target languages
 haftmann parents: 
37449diff
changeset | 10 | val target_OCaml: string | 
| 28054 | 11 | val setup: theory -> theory | 
| 12 | end; | |
| 13 | ||
| 14 | structure Code_ML : CODE_ML = | |
| 15 | struct | |
| 16 | ||
| 55150 | 17 | open Basic_Code_Symbol; | 
| 28054 | 18 | open Basic_Code_Thingol; | 
| 19 | open Code_Printer; | |
| 20 | ||
| 21 | infixr 5 @@; | |
| 22 | infixr 5 @|; | |
| 23 | ||
| 33992 | 24 | |
| 25 | (** generic **) | |
| 26 | ||
| 28054 | 27 | val target_SML = "SML"; | 
| 28 | val target_OCaml = "OCaml"; | |
| 29 | ||
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 30 | datatype ml_binding = | 
| 35228 | 31 | ML_Function of string * (typscheme * ((iterm list * iterm) * (thm option * bool)) list) | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 32 |   | ML_Instance of (string * class) * { class: class, tyco: string, vs: (vname * sort) list,
 | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 33 | superinsts: (class * dict list list) list, | 
| 52519 
598addf65209
explicit hint for domain of class parameters in instance statements
 haftmann parents: 
52435diff
changeset | 34 | inst_params: ((string * (const * int)) * (thm * bool)) list, | 
| 
598addf65209
explicit hint for domain of class parameters in instance statements
 haftmann parents: 
52435diff
changeset | 35 | superinst_params: ((string * (const * int)) * (thm * bool)) list }; | 
| 28054 | 36 | |
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 37 | datatype ml_stmt = | 
| 33992 | 38 | ML_Exc of string * (typscheme * int) | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 39 | | ML_Val of ml_binding | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 40 | | ML_Funs of (Code_Namespace.export * ml_binding) list * Code_Symbol.T list | 
| 48003 
1d11af40b106
dropped sort constraints on datatype specifications
 haftmann parents: 
47609diff
changeset | 41 | | ML_Datas of (string * (vname list * ((string * vname list) * itype list) list)) list | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 42 | | ML_Class of string * (vname * ((class * class) list * (string * itype) list)); | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 43 | |
| 33992 | 44 | fun print_product _ [] = NONE | 
| 45 | | print_product print [x] = SOME (print x) | |
| 34178 
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
 haftmann parents: 
34071diff
changeset | 46 | | print_product print xs = (SOME o enum " *" "" "") (map print xs); | 
| 28054 | 47 | |
| 38922 | 48 | fun tuplify _ _ [] = NONE | 
| 49 | | tuplify print fxy [x] = SOME (print fxy x) | |
| 50 |   | tuplify print _ xs = SOME (enum "," "(" ")" (map (print NOBR) xs));
 | |
| 28054 | 51 | |
| 33992 | 52 | |
| 53 | (** SML serializer **) | |
| 54 | ||
| 50625 | 55 | fun print_sml_stmt tyco_syntax const_syntax reserved is_constr deresolve = | 
| 28054 | 56 | let | 
| 55150 | 57 | val deresolve_const = deresolve o Constant; | 
| 58 | val deresolve_class = deresolve o Type_Class; | |
| 59 | val deresolve_classrel = deresolve o Class_Relation; | |
| 60 | val deresolve_inst = deresolve o Class_Instance; | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 61 | fun print_tyco_expr (sym, []) = (str o deresolve) sym | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 62 | | print_tyco_expr (sym, [ty]) = | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 63 | concat [print_typ BR ty, (str o deresolve) sym] | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 64 | | print_tyco_expr (sym, tys) = | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 65 |           concat [enum "," "(" ")" (map (print_typ BR) tys), (str o deresolve) sym]
 | 
| 38923 | 66 | and print_typ fxy (tyco `%% tys) = (case tyco_syntax tyco | 
| 55150 | 67 | of NONE => print_tyco_expr (Type_Constructor tyco, tys) | 
| 47609 | 68 | | SOME (_, print) => print print_typ fxy tys) | 
| 33992 | 69 |       | print_typ fxy (ITyVar v) = str ("'" ^ v);
 | 
| 55150 | 70 | fun print_dicttyp (class, ty) = print_tyco_expr (Type_Class class, [ty]); | 
| 34178 
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
 haftmann parents: 
34071diff
changeset | 71 | fun print_typscheme_prefix (vs, p) = enum " ->" "" "" | 
| 33992 | 72 | (map_filter (fn (v, sort) => | 
| 73 | (print_product (fn class => print_dicttyp (class, ITyVar v)) sort)) vs @| p); | |
| 74 | fun print_typscheme (vs, ty) = print_typscheme_prefix (vs, print_typ NOBR ty); | |
| 75 | fun print_dicttypscheme (vs, class_ty) = print_typscheme_prefix (vs, print_dicttyp class_ty); | |
| 41100 
6c0940392fb4
dictionary constants must permit explicit weakening of classes;
 haftmann parents: 
40627diff
changeset | 76 | fun print_classrels fxy [] ps = brackify fxy ps | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 77 | | print_classrels fxy [classrel] ps = brackify fxy [(str o deresolve_classrel) classrel, brackify BR ps] | 
| 41100 
6c0940392fb4
dictionary constants must permit explicit weakening of classes;
 haftmann parents: 
40627diff
changeset | 78 | | print_classrels fxy classrels ps = | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 79 |           brackify fxy [enum " o" "(" ")" (map (str o deresolve_classrel) classrels), brackify BR ps]
 | 
| 41118 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 80 | fun print_dict is_pseudo_fun fxy (Dict (classrels, x)) = | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 81 | print_classrels fxy classrels (print_plain_dict is_pseudo_fun fxy x) | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 82 | and print_plain_dict is_pseudo_fun fxy (Dict_Const (inst, dss)) = | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 83 | ((str o deresolve_inst) inst :: | 
| 55150 | 84 | (if is_pseudo_fun (Class_Instance inst) then [str "()"] | 
| 33992 | 85 | else map_filter (print_dicts is_pseudo_fun BR) dss)) | 
| 41118 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 86 | | print_plain_dict is_pseudo_fun fxy (Dict_Var (v, (i, k))) = | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 87 | [str (if k = 1 then first_upper v ^ "_" | 
| 41100 
6c0940392fb4
dictionary constants must permit explicit weakening of classes;
 haftmann parents: 
40627diff
changeset | 88 | else first_upper v ^ string_of_int (i+1) ^ "_")] | 
| 38922 | 89 | and print_dicts is_pseudo_fun = tuplify (print_dict is_pseudo_fun); | 
| 33992 | 90 | val print_dict_args = map_filter (fn (v, sort) => print_dicts (K false) BR | 
| 41118 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 91 | (map_index (fn (i, _) => Dict ([], Dict_Var (v, (i, length sort)))) sort)); | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 92 | fun print_term is_pseudo_fun some_thm vars fxy (IConst const) = | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 93 | print_app is_pseudo_fun some_thm vars fxy (const, []) | 
| 35228 | 94 | | print_term is_pseudo_fun some_thm vars fxy (IVar NONE) = | 
| 31889 | 95 | str "_" | 
| 35228 | 96 | | print_term is_pseudo_fun some_thm vars fxy (IVar (SOME v)) = | 
| 32924 
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
 haftmann parents: 
32913diff
changeset | 97 | str (lookup_var vars v) | 
| 35228 | 98 | | print_term is_pseudo_fun some_thm vars fxy (t as t1 `$ t2) = | 
| 28054 | 99 | (case Code_Thingol.unfold_const_app t | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 100 | of SOME app => print_app is_pseudo_fun some_thm vars fxy app | 
| 35228 | 101 | | NONE => brackify fxy [print_term is_pseudo_fun some_thm vars NOBR t1, | 
| 102 | print_term is_pseudo_fun some_thm vars BR t2]) | |
| 103 | | print_term is_pseudo_fun some_thm vars fxy (t as _ `|=> _) = | |
| 28054 | 104 | let | 
| 31874 | 105 | val (binds, t') = Code_Thingol.unfold_pat_abs t; | 
| 33992 | 106 | fun print_abs (pat, ty) = | 
| 35228 | 107 | print_bind is_pseudo_fun some_thm NOBR pat | 
| 28054 | 108 | #>> (fn p => concat [str "fn", p, str "=>"]); | 
| 33992 | 109 | val (ps, vars') = fold_map print_abs binds vars; | 
| 35228 | 110 | in brackets (ps @ [print_term is_pseudo_fun some_thm vars' NOBR t']) end | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 111 | | print_term is_pseudo_fun some_thm vars fxy (ICase case_expr) = | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 112 | (case Code_Thingol.unfold_const_app (#primitive case_expr) | 
| 55150 | 113 |            of SOME (app as ({ sym = Constant const, ... }, _)) =>
 | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 114 | if is_none (const_syntax const) | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 115 | then print_case is_pseudo_fun some_thm vars fxy case_expr | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 116 | else print_app is_pseudo_fun some_thm vars fxy app | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 117 | | NONE => print_case is_pseudo_fun some_thm vars fxy case_expr) | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 118 |     and print_app_expr is_pseudo_fun some_thm vars (app as ({ sym, dicts = dss, dom = dom, ... }, ts)) =
 | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 119 | if is_constr sym then | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 120 | let val k = length dom in | 
| 33992 | 121 | if k < 2 orelse length ts = k | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 122 | then (str o deresolve) sym | 
| 38922 | 123 | :: the_list (tuplify (print_term is_pseudo_fun some_thm vars) BR ts) | 
| 35228 | 124 | else [print_term is_pseudo_fun some_thm vars BR (Code_Thingol.eta_expand k app)] | 
| 33992 | 125 | end | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 126 | else if is_pseudo_fun sym | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 127 | then (str o deresolve) sym @@ str "()" | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 128 | else (str o deresolve) sym :: map_filter (print_dicts is_pseudo_fun BR) dss | 
| 35228 | 129 | @ map (print_term is_pseudo_fun some_thm vars BR) ts | 
| 130 | and print_app is_pseudo_fun some_thm vars = gen_print_app (print_app_expr is_pseudo_fun) | |
| 38923 | 131 | (print_term is_pseudo_fun) const_syntax some_thm vars | 
| 33992 | 132 | and print_bind is_pseudo_fun = gen_print_bind (print_term is_pseudo_fun) | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 133 |     and print_case is_pseudo_fun some_thm vars fxy { clauses = [], ... } =
 | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 134 | (concat o map str) ["raise", "Fail", "\"empty case\""] | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 135 |       | print_case is_pseudo_fun some_thm vars fxy (case_expr as { clauses = [_], ... }) =
 | 
| 28054 | 136 | let | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 137 | val (binds, body) = Code_Thingol.unfold_let (ICase case_expr); | 
| 47609 | 138 | fun print_match ((pat, _), t) vars = | 
| 28054 | 139 | vars | 
| 35228 | 140 | |> print_bind is_pseudo_fun some_thm NOBR pat | 
| 33992 | 141 | |>> (fn p => semicolon [str "val", p, str "=", | 
| 35228 | 142 | print_term is_pseudo_fun some_thm vars NOBR t]) | 
| 33992 | 143 | val (ps, vars') = fold_map print_match binds vars; | 
| 28054 | 144 | in | 
| 145 | Pretty.chunks [ | |
| 34178 
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
 haftmann parents: 
34071diff
changeset | 146 | Pretty.block [str "let", Pretty.fbrk, Pretty.chunks ps], | 
| 35228 | 147 | Pretty.block [str "in", Pretty.fbrk, print_term is_pseudo_fun some_thm vars' NOBR body], | 
| 34178 
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
 haftmann parents: 
34071diff
changeset | 148 | str "end" | 
| 28054 | 149 | ] | 
| 150 | end | |
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 151 |       | print_case is_pseudo_fun some_thm vars fxy { term = t, typ = ty, clauses = clause :: clauses, ... } =
 | 
| 28054 | 152 | let | 
| 33992 | 153 | fun print_select delim (pat, body) = | 
| 28054 | 154 | let | 
| 35228 | 155 | val (p, vars') = print_bind is_pseudo_fun some_thm NOBR pat vars; | 
| 28054 | 156 | in | 
| 35228 | 157 | concat [str delim, p, str "=>", print_term is_pseudo_fun some_thm vars' NOBR body] | 
| 28054 | 158 | end; | 
| 159 | in | |
| 31665 | 160 | brackets ( | 
| 28054 | 161 | str "case" | 
| 35228 | 162 | :: print_term is_pseudo_fun some_thm vars NOBR t | 
| 33992 | 163 | :: print_select "of" clause | 
| 164 | :: map (print_select "|") clauses | |
| 28054 | 165 | ) | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 166 | end; | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 167 | fun print_val_decl print_typscheme (sym, typscheme) = concat | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 168 | [str "val", str (deresolve sym), str ":", print_typscheme typscheme]; | 
| 33992 | 169 | fun print_datatype_decl definer (tyco, (vs, cos)) = | 
| 170 | let | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 171 | fun print_co ((co, _), []) = str (deresolve_const co) | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 172 | | print_co ((co, _), tys) = concat [str (deresolve_const co), str "of", | 
| 34178 
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
 haftmann parents: 
34071diff
changeset | 173 | enum " *" "" "" (map (print_typ (INFX (2, X))) tys)]; | 
| 33992 | 174 | in | 
| 175 | concat ( | |
| 176 | str definer | |
| 55150 | 177 | :: print_tyco_expr (Type_Constructor tyco, map ITyVar vs) | 
| 33992 | 178 | :: str "=" | 
| 179 | :: separate (str "|") (map print_co cos) | |
| 180 | ) | |
| 181 | end; | |
| 182 | fun print_def is_pseudo_fun needs_typ definer | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 183 | (ML_Function (const, (vs_ty as (vs, ty), eq :: eqs))) = | 
| 29189 | 184 | let | 
| 35228 | 185 | fun print_eqn definer ((ts, t), (some_thm, _)) = | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 186 | let | 
| 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 187 | val vars = reserved | 
| 55145 
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
 haftmann parents: 
52519diff
changeset | 188 | |> intro_base_names_for (is_none o const_syntax) | 
| 
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
 haftmann parents: 
52519diff
changeset | 189 | deresolve (t :: ts) | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 190 | |> intro_vars ((fold o Code_Thingol.fold_varnames) | 
| 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 191 | (insert (op =)) ts []); | 
| 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 192 | val prolog = if needs_typ then | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 193 | concat [str definer, (str o deresolve_const) const, str ":", print_typ NOBR ty] | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 194 | else (concat o map str) [definer, deresolve_const const]; | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 195 | in | 
| 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 196 | concat ( | 
| 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 197 | prolog | 
| 55150 | 198 | :: (if is_pseudo_fun (Constant const) then [str "()"] | 
| 33992 | 199 | else print_dict_args vs | 
| 35228 | 200 | @ map (print_term is_pseudo_fun some_thm vars BR) ts) | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 201 | @ str "=" | 
| 35228 | 202 | @@ print_term is_pseudo_fun some_thm vars NOBR t | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 203 | ) | 
| 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 204 | end | 
| 33992 | 205 | val shift = if null eqs then I else | 
| 206 | map (Pretty.block o single o Pretty.block o single); | |
| 207 | in pair | |
| 55150 | 208 | (print_val_decl print_typscheme (Constant const, vs_ty)) | 
| 33992 | 209 | ((Pretty.block o Pretty.fbreaks o shift) ( | 
| 210 | print_eqn definer eq | |
| 211 | :: map (print_eqn "|") eqs | |
| 212 | )) | |
| 29189 | 213 | end | 
| 33992 | 214 | | print_def is_pseudo_fun _ definer | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 215 |           (ML_Instance (inst as (tyco, class), { vs, superinsts, inst_params, ... })) =
 | 
| 29189 | 216 | let | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 217 | fun print_super_instance (super_class, x) = | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 218 | concat [ | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 219 | (str o Long_Name.base_name o deresolve_classrel) (class, super_class), | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 220 | str "=", | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 221 | print_dict is_pseudo_fun NOBR (Dict ([], Dict_Const ((tyco, super_class), x))) | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 222 | ]; | 
| 52519 
598addf65209
explicit hint for domain of class parameters in instance statements
 haftmann parents: 
52435diff
changeset | 223 | fun print_classparam_instance ((classparam, (const, _)), (thm, _)) = | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 224 | concat [ | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 225 | (str o Long_Name.base_name o deresolve_const) classparam, | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 226 | str "=", | 
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37242diff
changeset | 227 | print_app (K false) (SOME thm) reserved NOBR (const, []) | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 228 | ]; | 
| 33992 | 229 | in pair | 
| 230 | (print_val_decl print_dicttypscheme | |
| 55150 | 231 | (Class_Instance inst, (vs, (class, tyco `%% map (ITyVar o fst) vs)))) | 
| 33992 | 232 | (concat ( | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 233 | str definer | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 234 | :: (str o deresolve_inst) inst | 
| 55150 | 235 | :: (if is_pseudo_fun (Class_Instance inst) then [str "()"] | 
| 33992 | 236 | else print_dict_args vs) | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 237 | @ str "=" | 
| 34178 
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
 haftmann parents: 
34071diff
changeset | 238 |               :: enum "," "{" "}"
 | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 239 | (map print_super_instance superinsts | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 240 | @ map print_classparam_instance inst_params) | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 241 | :: str ":" | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 242 | @@ print_dicttyp (class, tyco `%% map (ITyVar o fst) vs) | 
| 33992 | 243 | )) | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 244 | end; | 
| 55681 | 245 | fun print_stmt _ (ML_Exc (const, (vs_ty, n))) = pair | 
| 55150 | 246 | [print_val_decl print_typscheme (Constant const, vs_ty)] | 
| 33992 | 247 | ((semicolon o map str) ( | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 248 | (if n = 0 then "val" else "fun") | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 249 | :: deresolve_const const | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 250 | :: replicate n "_" | 
| 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 251 | @ "=" | 
| 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 252 | :: "raise" | 
| 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 253 | :: "Fail" | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 254 | @@ ML_Syntax.print_string const | 
| 33992 | 255 | )) | 
| 55681 | 256 | | print_stmt _ (ML_Val binding) = | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 257 | let | 
| 33992 | 258 | val (sig_p, p) = print_def (K false) true "val" binding | 
| 259 | in pair | |
| 260 | [sig_p] | |
| 261 | (semicolon [p]) | |
| 262 | end | |
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 263 | | print_stmt _ (ML_Funs ((export, binding) :: exports_bindings, pseudo_funs)) = | 
| 33992 | 264 | let | 
| 265 | val print_def' = print_def (member (op =) pseudo_funs) false; | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 266 | fun print_pseudo_fun sym = concat [ | 
| 29189 | 267 | str "val", | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 268 | (str o deresolve) sym, | 
| 29189 | 269 | str "=", | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 270 | (str o deresolve) sym, | 
| 29189 | 271 | str "();" | 
| 272 | ]; | |
| 33992 | 273 | val (sig_ps, (ps, p)) = (apsnd split_last o split_list) | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 274 | (print_def' "fun" binding :: map (print_def' "and" o snd) exports_bindings); | 
| 33992 | 275 | val pseudo_ps = map print_pseudo_fun pseudo_funs; | 
| 276 | in pair | |
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 277 | (map_filter (fn (export, p) => if Code_Namespace.not_private export then SOME p else NONE) | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 278 | ((export :: map fst exports_bindings) ~~ sig_ps)) | 
| 33992 | 279 | (Pretty.chunks (ps @ semicolon [p] :: pseudo_ps)) | 
| 280 | end | |
| 55681 | 281 | | print_stmt _ (ML_Datas [(tyco, (vs, []))]) = | 
| 33992 | 282 | let | 
| 55150 | 283 | val ty_p = print_tyco_expr (Type_Constructor tyco, map ITyVar vs); | 
| 33992 | 284 | in | 
| 285 | pair | |
| 286 | [concat [str "type", ty_p]] | |
| 55681 | 287 | (semicolon [str "datatype", ty_p, str "=", str "EMPTY__"]) | 
| 33992 | 288 | end | 
| 55681 | 289 | | print_stmt export (ML_Datas (data :: datas)) = | 
| 28054 | 290 | let | 
| 55681 | 291 | val decl_ps = print_datatype_decl "datatype" data | 
| 33992 | 292 | :: map (print_datatype_decl "and") datas; | 
| 55681 | 293 | val (ps, p) = split_last decl_ps; | 
| 33992 | 294 | in pair | 
| 55681 | 295 | (if Code_Namespace.is_public export | 
| 296 | then decl_ps | |
| 297 | else map (fn (tyco, (vs, _)) => | |
| 298 | concat [str "type", print_tyco_expr (Type_Constructor tyco, map ITyVar vs)]) | |
| 299 | (data :: datas)) | |
| 33992 | 300 | (Pretty.chunks (ps @| semicolon [p])) | 
| 301 | end | |
| 55681 | 302 | | print_stmt export (ML_Class (class, (v, (classrels, classparams)))) = | 
| 28054 | 303 | let | 
| 33992 | 304 | fun print_field s p = concat [str s, str ":", p]; | 
| 305 | fun print_proj s p = semicolon | |
| 306 | (map str ["val", s, "=", "#" ^ s, ":"] @| p); | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 307 | fun print_super_class_decl (classrel as (_, super_class)) = | 
| 33992 | 308 | print_val_decl print_dicttypscheme | 
| 55150 | 309 | (Class_Relation classrel, ([(v, [class])], (super_class, ITyVar v))); | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 310 | fun print_super_class_field (classrel as (_, super_class)) = | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 311 | print_field (deresolve_classrel classrel) (print_dicttyp (super_class, ITyVar v)); | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 312 | fun print_super_class_proj (classrel as (_, super_class)) = | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 313 | print_proj (deresolve_classrel classrel) | 
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37242diff
changeset | 314 | (print_dicttypscheme ([(v, [class])], (super_class, ITyVar v))); | 
| 33992 | 315 | fun print_classparam_decl (classparam, ty) = | 
| 316 | print_val_decl print_typscheme | |
| 55150 | 317 | (Constant classparam, ([(v, [class])], ty)); | 
| 33992 | 318 | fun print_classparam_field (classparam, ty) = | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 319 | print_field (deresolve_const classparam) (print_typ NOBR ty); | 
| 33992 | 320 | fun print_classparam_proj (classparam, ty) = | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 321 | print_proj (deresolve_const classparam) | 
| 33992 | 322 | (print_typscheme ([(v, [class])], ty)); | 
| 323 | in pair | |
| 324 | (concat [str "type", print_dicttyp (class, ITyVar v)] | |
| 55681 | 325 | :: (if Code_Namespace.is_public export | 
| 326 | then map print_super_class_decl classrels | |
| 327 | @ map print_classparam_decl classparams | |
| 328 | else [])) | |
| 33992 | 329 | (Pretty.chunks ( | 
| 28054 | 330 | concat [ | 
| 55681 | 331 | str "type", | 
| 332 | print_dicttyp (class, ITyVar v), | |
| 28054 | 333 | str "=", | 
| 34178 
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
 haftmann parents: 
34071diff
changeset | 334 |                 enum "," "{" "};" (
 | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 335 | map print_super_class_field classrels | 
| 33992 | 336 | @ map print_classparam_field classparams | 
| 28054 | 337 | ) | 
| 338 | ] | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 339 | :: map print_super_class_proj classrels | 
| 33992 | 340 | @ map print_classparam_proj classparams | 
| 341 | )) | |
| 28054 | 342 | end; | 
| 33992 | 343 | in print_stmt end; | 
| 28054 | 344 | |
| 55677 | 345 | fun print_sml_module name decls body = | 
| 38933 
bd77e092f67c
avoid strange special treatment of empty module names
 haftmann parents: 
38928diff
changeset | 346 | Pretty.chunks2 ( | 
| 55677 | 347 | Pretty.chunks [ | 
| 348 |       str ("structure " ^ name ^ " : sig"),
 | |
| 349 | (indent 2 o Pretty.chunks) decls, | |
| 350 | str "end = struct" | |
| 351 | ] | |
| 33992 | 352 | :: body | 
| 353 |     @| str ("end; (*struct " ^ name ^ "*)")
 | |
| 28054 | 354 | ); | 
| 355 | ||
| 28064 | 356 | val literals_sml = Literals {
 | 
| 357 | literal_char = prefix "#" o quote o ML_Syntax.print_char, | |
| 358 | literal_string = quote o translate_string ML_Syntax.print_char, | |
| 34944 
970e1466028d
code literals: distinguish numeral classes by different entries
 haftmann parents: 
34246diff
changeset | 359 |   literal_numeral = fn k => "(" ^ string_of_int k ^ " : IntInf.int)",
 | 
| 34178 
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
 haftmann parents: 
34071diff
changeset | 360 | literal_list = enum "," "[" "]", | 
| 28064 | 361 | infix_cons = (7, "::") | 
| 362 | }; | |
| 363 | ||
| 28054 | 364 | |
| 365 | (** OCaml serializer **) | |
| 366 | ||
| 50625 | 367 | fun print_ocaml_stmt tyco_syntax const_syntax reserved is_constr deresolve = | 
| 28054 | 368 | let | 
| 55150 | 369 | val deresolve_const = deresolve o Constant; | 
| 370 | val deresolve_class = deresolve o Type_Class; | |
| 371 | val deresolve_classrel = deresolve o Class_Relation; | |
| 372 | val deresolve_inst = deresolve o Class_Instance; | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 373 | fun print_tyco_expr (sym, []) = (str o deresolve) sym | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 374 | | print_tyco_expr (sym, [ty]) = | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 375 | concat [print_typ BR ty, (str o deresolve) sym] | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 376 | | print_tyco_expr (sym, tys) = | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 377 |           concat [enum "," "(" ")" (map (print_typ BR) tys), (str o deresolve) sym]
 | 
| 38923 | 378 | and print_typ fxy (tyco `%% tys) = (case tyco_syntax tyco | 
| 55150 | 379 | of NONE => print_tyco_expr (Type_Constructor tyco, tys) | 
| 47576 | 380 | | SOME (_, print) => print print_typ fxy tys) | 
| 33992 | 381 |       | print_typ fxy (ITyVar v) = str ("'" ^ v);
 | 
| 55150 | 382 | fun print_dicttyp (class, ty) = print_tyco_expr (Type_Class class, [ty]); | 
| 34178 
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
 haftmann parents: 
34071diff
changeset | 383 | fun print_typscheme_prefix (vs, p) = enum " ->" "" "" | 
| 33992 | 384 | (map_filter (fn (v, sort) => | 
| 385 | (print_product (fn class => print_dicttyp (class, ITyVar v)) sort)) vs @| p); | |
| 386 | fun print_typscheme (vs, ty) = print_typscheme_prefix (vs, print_typ NOBR ty); | |
| 387 | fun print_dicttypscheme (vs, class_ty) = print_typscheme_prefix (vs, print_dicttyp class_ty); | |
| 41100 
6c0940392fb4
dictionary constants must permit explicit weakening of classes;
 haftmann parents: 
40627diff
changeset | 388 | val print_classrels = | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 389 | fold_rev (fn classrel => fn p => Pretty.block [p, str ".", (str o deresolve_classrel) classrel]) | 
| 41118 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 390 | fun print_dict is_pseudo_fun fxy (Dict (classrels, x)) = | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 391 | print_plain_dict is_pseudo_fun fxy x | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 392 | |> print_classrels classrels | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 393 | and print_plain_dict is_pseudo_fun fxy (Dict_Const (inst, dss)) = | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 394 | brackify BR ((str o deresolve_inst) inst :: | 
| 55150 | 395 | (if is_pseudo_fun (Class_Instance inst) then [str "()"] | 
| 33992 | 396 | else map_filter (print_dicts is_pseudo_fun BR) dss)) | 
| 41118 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 397 | | print_plain_dict is_pseudo_fun fxy (Dict_Var (v, (i, k))) = | 
| 33992 | 398 | str (if k = 1 then "_" ^ first_upper v | 
| 399 | else "_" ^ first_upper v ^ string_of_int (i+1)) | |
| 38922 | 400 | and print_dicts is_pseudo_fun = tuplify (print_dict is_pseudo_fun); | 
| 33992 | 401 | val print_dict_args = map_filter (fn (v, sort) => print_dicts (K false) BR | 
| 41118 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 402 | (map_index (fn (i, _) => Dict ([], Dict_Var (v, (i, length sort)))) sort)); | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 403 | fun print_term is_pseudo_fun some_thm vars fxy (IConst const) = | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 404 | print_app is_pseudo_fun some_thm vars fxy (const, []) | 
| 35228 | 405 | | print_term is_pseudo_fun some_thm vars fxy (IVar NONE) = | 
| 31889 | 406 | str "_" | 
| 35228 | 407 | | print_term is_pseudo_fun some_thm vars fxy (IVar (SOME v)) = | 
| 32924 
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
 haftmann parents: 
32913diff
changeset | 408 | str (lookup_var vars v) | 
| 35228 | 409 | | print_term is_pseudo_fun some_thm vars fxy (t as t1 `$ t2) = | 
| 28054 | 410 | (case Code_Thingol.unfold_const_app t | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 411 | of SOME app => print_app is_pseudo_fun some_thm vars fxy app | 
| 35228 | 412 | | NONE => brackify fxy [print_term is_pseudo_fun some_thm vars NOBR t1, | 
| 413 | print_term is_pseudo_fun some_thm vars BR t2]) | |
| 414 | | print_term is_pseudo_fun some_thm vars fxy (t as _ `|=> _) = | |
| 28054 | 415 | let | 
| 31874 | 416 | val (binds, t') = Code_Thingol.unfold_pat_abs t; | 
| 35228 | 417 | val (ps, vars') = fold_map (print_bind is_pseudo_fun some_thm BR o fst) binds vars; | 
| 418 | in brackets (str "fun" :: ps @ str "->" @@ print_term is_pseudo_fun some_thm vars' NOBR t') end | |
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 419 | | print_term is_pseudo_fun some_thm vars fxy (ICase case_expr) = | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 420 | (case Code_Thingol.unfold_const_app (#primitive case_expr) | 
| 55150 | 421 |            of SOME (app as ({ sym = Constant const, ... }, _)) =>
 | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 422 | if is_none (const_syntax const) | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 423 | then print_case is_pseudo_fun some_thm vars fxy case_expr | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 424 | else print_app is_pseudo_fun some_thm vars fxy app | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 425 | | NONE => print_case is_pseudo_fun some_thm vars fxy case_expr) | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 426 |     and print_app_expr is_pseudo_fun some_thm vars (app as ({ sym, dicts = dss, dom = dom, ... }, ts)) =
 | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 427 | if is_constr sym then | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 428 | let val k = length dom in | 
| 33992 | 429 | if length ts = k | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 430 | then (str o deresolve) sym | 
| 38922 | 431 | :: the_list (tuplify (print_term is_pseudo_fun some_thm vars) BR ts) | 
| 35228 | 432 | else [print_term is_pseudo_fun some_thm vars BR (Code_Thingol.eta_expand k app)] | 
| 33992 | 433 | end | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 434 | else if is_pseudo_fun sym | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 435 | then (str o deresolve) sym @@ str "()" | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 436 | else (str o deresolve) sym :: map_filter (print_dicts is_pseudo_fun BR) dss | 
| 35228 | 437 | @ map (print_term is_pseudo_fun some_thm vars BR) ts | 
| 438 | and print_app is_pseudo_fun some_thm vars = gen_print_app (print_app_expr is_pseudo_fun) | |
| 38923 | 439 | (print_term is_pseudo_fun) const_syntax some_thm vars | 
| 33992 | 440 | and print_bind is_pseudo_fun = gen_print_bind (print_term is_pseudo_fun) | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 441 |     and print_case is_pseudo_fun some_thm vars fxy { clauses = [], ... } =
 | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 442 | (concat o map str) ["failwith", "\"empty case\""] | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 443 |       | print_case is_pseudo_fun some_thm vars fxy (case_expr as { clauses = [_], ... }) =
 | 
| 28054 | 444 | let | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 445 | val (binds, body) = Code_Thingol.unfold_let (ICase case_expr); | 
| 47576 | 446 | fun print_let ((pat, _), t) vars = | 
| 28054 | 447 | vars | 
| 35228 | 448 | |> print_bind is_pseudo_fun some_thm NOBR pat | 
| 28054 | 449 | |>> (fn p => concat | 
| 35228 | 450 | [str "let", p, str "=", print_term is_pseudo_fun some_thm vars NOBR t, str "in"]) | 
| 33992 | 451 | val (ps, vars') = fold_map print_let binds vars; | 
| 31665 | 452 | in | 
| 453 | brackify_block fxy (Pretty.chunks ps) [] | |
| 35228 | 454 | (print_term is_pseudo_fun some_thm vars' NOBR body) | 
| 31665 | 455 | end | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 456 |       | print_case is_pseudo_fun some_thm vars fxy { term = t, typ = ty, clauses = clause :: clauses, ... } =
 | 
| 28054 | 457 | let | 
| 33992 | 458 | fun print_select delim (pat, body) = | 
| 28054 | 459 | let | 
| 35228 | 460 | val (p, vars') = print_bind is_pseudo_fun some_thm NOBR pat vars; | 
| 461 | in concat [str delim, p, str "->", print_term is_pseudo_fun some_thm vars' NOBR body] end; | |
| 28054 | 462 | in | 
| 31665 | 463 | brackets ( | 
| 28054 | 464 | str "match" | 
| 35228 | 465 | :: print_term is_pseudo_fun some_thm vars NOBR t | 
| 33992 | 466 | :: print_select "with" clause | 
| 467 | :: map (print_select "|") clauses | |
| 28054 | 468 | ) | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 469 | end; | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 470 | fun print_val_decl print_typscheme (sym, typscheme) = concat | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 471 | [str "val", str (deresolve sym), str ":", print_typscheme typscheme]; | 
| 33992 | 472 | fun print_datatype_decl definer (tyco, (vs, cos)) = | 
| 473 | let | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 474 | fun print_co ((co, _), []) = str (deresolve_const co) | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 475 | | print_co ((co, _), tys) = concat [str (deresolve_const co), str "of", | 
| 34178 
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
 haftmann parents: 
34071diff
changeset | 476 | enum " *" "" "" (map (print_typ (INFX (2, X))) tys)]; | 
| 33992 | 477 | in | 
| 478 | concat ( | |
| 479 | str definer | |
| 55150 | 480 | :: print_tyco_expr (Type_Constructor tyco, map ITyVar vs) | 
| 33992 | 481 | :: str "=" | 
| 482 | :: separate (str "|") (map print_co cos) | |
| 483 | ) | |
| 484 | end; | |
| 485 | fun print_def is_pseudo_fun needs_typ definer | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 486 | (ML_Function (const, (vs_ty as (vs, ty), eqs))) = | 
| 28054 | 487 | let | 
| 35228 | 488 | fun print_eqn ((ts, t), (some_thm, _)) = | 
| 28054 | 489 | let | 
| 32924 
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
 haftmann parents: 
32913diff
changeset | 490 | val vars = reserved | 
| 55145 
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
 haftmann parents: 
52519diff
changeset | 491 | |> intro_base_names_for (is_none o const_syntax) | 
| 
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
 haftmann parents: 
52519diff
changeset | 492 | deresolve (t :: ts) | 
| 32924 
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
 haftmann parents: 
32913diff
changeset | 493 | |> intro_vars ((fold o Code_Thingol.fold_varnames) | 
| 28054 | 494 | (insert (op =)) ts []); | 
| 495 | in concat [ | |
| 38778 | 496 | (Pretty.block o commas) | 
| 35228 | 497 | (map (print_term is_pseudo_fun some_thm vars NOBR) ts), | 
| 28054 | 498 | str "->", | 
| 35228 | 499 | print_term is_pseudo_fun some_thm vars NOBR t | 
| 28054 | 500 | ] end; | 
| 35228 | 501 | fun print_eqns is_pseudo [((ts, t), (some_thm, _))] = | 
| 28054 | 502 | let | 
| 32924 
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
 haftmann parents: 
32913diff
changeset | 503 | val vars = reserved | 
| 55145 
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
 haftmann parents: 
52519diff
changeset | 504 | |> intro_base_names_for (is_none o const_syntax) | 
| 
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
 haftmann parents: 
52519diff
changeset | 505 | deresolve (t :: ts) | 
| 32924 
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
 haftmann parents: 
32913diff
changeset | 506 | |> intro_vars ((fold o Code_Thingol.fold_varnames) | 
| 28054 | 507 | (insert (op =)) ts []); | 
| 508 | in | |
| 509 | concat ( | |
| 29189 | 510 | (if is_pseudo then [str "()"] | 
| 35228 | 511 | else map (print_term is_pseudo_fun some_thm vars BR) ts) | 
| 28054 | 512 | @ str "=" | 
| 35228 | 513 | @@ print_term is_pseudo_fun some_thm vars NOBR t | 
| 28054 | 514 | ) | 
| 515 | end | |
| 33992 | 516 | | print_eqns _ ((eq as (([_], _), _)) :: eqs) = | 
| 28054 | 517 | Pretty.block ( | 
| 518 | str "=" | |
| 519 | :: Pretty.brk 1 | |
| 520 | :: str "function" | |
| 521 | :: Pretty.brk 1 | |
| 33992 | 522 | :: print_eqn eq | 
| 28054 | 523 | :: maps (append [Pretty.fbrk, str "|", Pretty.brk 1] | 
| 33992 | 524 | o single o print_eqn) eqs | 
| 28054 | 525 | ) | 
| 33992 | 526 | | print_eqns _ (eqs as eq :: eqs') = | 
| 28054 | 527 | let | 
| 32924 
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
 haftmann parents: 
32913diff
changeset | 528 | val vars = reserved | 
| 55145 
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
 haftmann parents: 
52519diff
changeset | 529 | |> intro_base_names_for (is_none o const_syntax) | 
| 
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
 haftmann parents: 
52519diff
changeset | 530 | deresolve (map (snd o fst) eqs) | 
| 32924 
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
 haftmann parents: 
32913diff
changeset | 531 | val dummy_parms = (map str o aux_params vars o map (fst o fst)) eqs; | 
| 28054 | 532 | in | 
| 533 | Pretty.block ( | |
| 534 | Pretty.breaks dummy_parms | |
| 535 | @ Pretty.brk 1 | |
| 536 | :: str "=" | |
| 537 | :: Pretty.brk 1 | |
| 538 | :: str "match" | |
| 539 | :: Pretty.brk 1 | |
| 38778 | 540 | :: (Pretty.block o commas) dummy_parms | 
| 28054 | 541 | :: Pretty.brk 1 | 
| 542 | :: str "with" | |
| 543 | :: Pretty.brk 1 | |
| 33992 | 544 | :: print_eqn eq | 
| 28054 | 545 | :: maps (append [Pretty.fbrk, str "|", Pretty.brk 1] | 
| 33992 | 546 | o single o print_eqn) eqs' | 
| 28054 | 547 | ) | 
| 548 | end; | |
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 549 | val prolog = if needs_typ then | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 550 | concat [str definer, (str o deresolve_const) const, str ":", print_typ NOBR ty] | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 551 | else (concat o map str) [definer, deresolve_const const]; | 
| 33992 | 552 | in pair | 
| 55150 | 553 | (print_val_decl print_typscheme (Constant const, vs_ty)) | 
| 33992 | 554 | (concat ( | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 555 | prolog | 
| 33992 | 556 | :: print_dict_args vs | 
| 55150 | 557 | @| print_eqns (is_pseudo_fun (Constant const)) eqs | 
| 33992 | 558 | )) | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 559 | end | 
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37242diff
changeset | 560 | | print_def is_pseudo_fun _ definer | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 561 |           (ML_Instance (inst as (tyco, class), { vs, superinsts, inst_params, ... })) =
 | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 562 | let | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 563 | fun print_super_instance (super_class, x) = | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 564 | concat [ | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 565 | (str o deresolve_classrel) (class, super_class), | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 566 | str "=", | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 567 | print_dict is_pseudo_fun NOBR (Dict ([], Dict_Const ((tyco, super_class), x))) | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 568 | ]; | 
| 52519 
598addf65209
explicit hint for domain of class parameters in instance statements
 haftmann parents: 
52435diff
changeset | 569 | fun print_classparam_instance ((classparam, (const, _)), (thm, _)) = | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 570 | concat [ | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 571 | (str o deresolve_const) classparam, | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 572 | str "=", | 
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37242diff
changeset | 573 | print_app (K false) (SOME thm) reserved NOBR (const, []) | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 574 | ]; | 
| 33992 | 575 | in pair | 
| 576 | (print_val_decl print_dicttypscheme | |
| 55150 | 577 | (Class_Instance inst, (vs, (class, tyco `%% map (ITyVar o fst) vs)))) | 
| 33992 | 578 | (concat ( | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 579 | str definer | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 580 | :: (str o deresolve_inst) inst | 
| 55150 | 581 | :: (if is_pseudo_fun (Class_Instance inst) then [str "()"] | 
| 43343 
e83695ea0e0a
resolving an issue with class instances that are pseudo functions in the OCaml code serializer
 bulwahn parents: 
42599diff
changeset | 582 | else print_dict_args vs) | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 583 | @ str "=" | 
| 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 584 | @@ brackets [ | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 585 |                 enum_default "()" ";" "{" "}" (map print_super_instance superinsts
 | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 586 | @ map print_classparam_instance inst_params), | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 587 | str ":", | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 588 | print_dicttyp (class, tyco `%% map (ITyVar o fst) vs) | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 589 | ] | 
| 33992 | 590 | )) | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 591 | end; | 
| 55681 | 592 | fun print_stmt _ (ML_Exc (const, (vs_ty, n))) = pair | 
| 55150 | 593 | [print_val_decl print_typscheme (Constant const, vs_ty)] | 
| 33992 | 594 | ((doublesemicolon o map str) ( | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 595 | "let" | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 596 | :: deresolve_const const | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 597 | :: replicate n "_" | 
| 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 598 | @ "=" | 
| 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 599 | :: "failwith" | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 600 | @@ ML_Syntax.print_string const | 
| 33992 | 601 | )) | 
| 55681 | 602 | | print_stmt _ (ML_Val binding) = | 
| 33636 
a9bb3f063773
accomplish mutual recursion between fun and inst
 haftmann parents: 
33519diff
changeset | 603 | let | 
| 33992 | 604 | val (sig_p, p) = print_def (K false) true "let" binding | 
| 605 | in pair | |
| 606 | [sig_p] | |
| 607 | (doublesemicolon [p]) | |
| 608 | end | |
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 609 | | print_stmt _ (ML_Funs ((export, binding) :: exports_bindings, pseudo_funs)) = | 
| 33992 | 610 | let | 
| 611 | val print_def' = print_def (member (op =) pseudo_funs) false; | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 612 | fun print_pseudo_fun sym = concat [ | 
| 29189 | 613 | str "let", | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 614 | (str o deresolve) sym, | 
| 29189 | 615 | str "=", | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 616 | (str o deresolve) sym, | 
| 29189 | 617 | str "();;" | 
| 618 | ]; | |
| 33992 | 619 | val (sig_ps, (ps, p)) = (apsnd split_last o split_list) | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 620 | (print_def' "let rec" binding :: map (print_def' "and" o snd) exports_bindings); | 
| 33992 | 621 | val pseudo_ps = map print_pseudo_fun pseudo_funs; | 
| 622 | in pair | |
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 623 | (map_filter (fn (export, p) => if Code_Namespace.not_private export then SOME p else NONE) | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 624 | ((export :: map fst exports_bindings) ~~ sig_ps)) | 
| 33992 | 625 | (Pretty.chunks (ps @ doublesemicolon [p] :: pseudo_ps)) | 
| 626 | end | |
| 55681 | 627 | | print_stmt _ (ML_Datas [(tyco, (vs, []))]) = | 
| 33992 | 628 | let | 
| 55150 | 629 | val ty_p = print_tyco_expr (Type_Constructor tyco, map ITyVar vs); | 
| 33992 | 630 | in | 
| 631 | pair | |
| 632 | [concat [str "type", ty_p]] | |
| 55681 | 633 | (doublesemicolon [str "type", ty_p, str "=", str "EMPTY__"]) | 
| 33992 | 634 | end | 
| 55681 | 635 | | print_stmt export (ML_Datas (data :: datas)) = | 
| 28054 | 636 | let | 
| 55681 | 637 | val decl_ps = print_datatype_decl "type" data | 
| 33992 | 638 | :: map (print_datatype_decl "and") datas; | 
| 55681 | 639 | val (ps, p) = split_last decl_ps; | 
| 33992 | 640 | in pair | 
| 55681 | 641 | (if Code_Namespace.is_public export | 
| 642 | then decl_ps | |
| 643 | else map (fn (tyco, (vs, _)) => | |
| 644 | concat [str "type", print_tyco_expr (Type_Constructor tyco, map ITyVar vs)]) | |
| 645 | (data :: datas)) | |
| 33992 | 646 | (Pretty.chunks (ps @| doublesemicolon [p])) | 
| 647 | end | |
| 55681 | 648 | | print_stmt export (ML_Class (class, (v, (classrels, classparams)))) = | 
| 28054 | 649 | let | 
| 33992 | 650 | fun print_field s p = concat [str s, str ":", p]; | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 651 | fun print_super_class_field (classrel as (_, super_class)) = | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 652 | print_field (deresolve_classrel classrel) (print_dicttyp (super_class, ITyVar v)); | 
| 33992 | 653 | fun print_classparam_decl (classparam, ty) = | 
| 654 | print_val_decl print_typscheme | |
| 55150 | 655 | (Constant classparam, ([(v, [class])], ty)); | 
| 33992 | 656 | fun print_classparam_field (classparam, ty) = | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 657 | print_field (deresolve_const classparam) (print_typ NOBR ty); | 
| 32924 
d2e9b2dab760
dropped Code_Printer prefix where feasible; fixed whitespace issues; more coherent terminology
 haftmann parents: 
32913diff
changeset | 658 | val w = "_" ^ first_upper v; | 
| 33992 | 659 | fun print_classparam_proj (classparam, _) = | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 660 | (concat o map str) ["let", deresolve_const classparam, w, "=", | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 661 | w ^ "." ^ deresolve_const classparam ^ ";;"]; | 
| 33992 | 662 | val type_decl_p = concat [ | 
| 55682 | 663 | str "type", | 
| 664 | print_dicttyp (class, ITyVar v), | |
| 28054 | 665 | str "=", | 
| 33992 | 666 |                 enum_default "unit" ";" "{" "}" (
 | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 667 | map print_super_class_field classrels | 
| 33992 | 668 | @ map print_classparam_field classparams | 
| 669 | ) | |
| 28054 | 670 | ]; | 
| 33992 | 671 | in pair | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 672 | (if Code_Namespace.is_public export | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 673 | then type_decl_p :: map print_classparam_decl classparams | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 674 | else [concat [str "type", print_dicttyp (class, ITyVar v)]]) | 
| 33992 | 675 | (Pretty.chunks ( | 
| 676 | doublesemicolon [type_decl_p] | |
| 677 | :: map print_classparam_proj classparams | |
| 678 | )) | |
| 679 | end; | |
| 680 | in print_stmt end; | |
| 28054 | 681 | |
| 55677 | 682 | fun print_ocaml_module name decls body = | 
| 38933 
bd77e092f67c
avoid strange special treatment of empty module names
 haftmann parents: 
38928diff
changeset | 683 | Pretty.chunks2 ( | 
| 55677 | 684 | Pretty.chunks [ | 
| 685 |       str ("module " ^ name ^ " : sig"),
 | |
| 686 | (indent 2 o Pretty.chunks) decls, | |
| 687 | str "end = struct" | |
| 688 | ] | |
| 33992 | 689 | :: body | 
| 690 |     @| str ("end;; (*struct " ^ name ^ "*)")
 | |
| 28054 | 691 | ); | 
| 692 | ||
| 28064 | 693 | val literals_ocaml = let | 
| 694 | fun chr i = | |
| 695 | let | |
| 696 | val xs = string_of_int i; | |
| 40627 
becf5d5187cc
renamed raw "explode" function to "raw_explode" to emphasize its meaning;
 wenzelm parents: 
39148diff
changeset | 697 | val ys = replicate_string (3 - length (raw_explode xs)) "0"; | 
| 28064 | 698 | in "\\" ^ ys ^ xs end; | 
| 699 | fun char_ocaml c = | |
| 700 | let | |
| 701 | val i = ord c; | |
| 702 | val s = if i < 32 orelse i = 34 orelse i = 39 orelse i = 92 orelse i > 126 | |
| 703 | then chr i else c | |
| 704 | in s end; | |
| 34944 
970e1466028d
code literals: distinguish numeral classes by different entries
 haftmann parents: 
34246diff
changeset | 705 | fun numeral_ocaml k = if k < 0 | 
| 
970e1466028d
code literals: distinguish numeral classes by different entries
 haftmann parents: 
34246diff
changeset | 706 | then "(Big_int.minus_big_int " ^ numeral_ocaml (~ k) ^ ")" | 
| 
970e1466028d
code literals: distinguish numeral classes by different entries
 haftmann parents: 
34246diff
changeset | 707 | else if k <= 1073741823 | 
| 
970e1466028d
code literals: distinguish numeral classes by different entries
 haftmann parents: 
34246diff
changeset | 708 | then "(Big_int.big_int_of_int " ^ string_of_int k ^ ")" | 
| 
970e1466028d
code literals: distinguish numeral classes by different entries
 haftmann parents: 
34246diff
changeset | 709 | else "(Big_int.big_int_of_string " ^ quote (string_of_int k) ^ ")" | 
| 28064 | 710 | in Literals {
 | 
| 34178 
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
 haftmann parents: 
34071diff
changeset | 711 | literal_char = Library.enclose "'" "'" o char_ocaml, | 
| 28064 | 712 | literal_string = quote o translate_string char_ocaml, | 
| 34944 
970e1466028d
code literals: distinguish numeral classes by different entries
 haftmann parents: 
34246diff
changeset | 713 | literal_numeral = numeral_ocaml, | 
| 34178 
a78b8d5b91cb
take care for destructive print mode properly using dedicated pretty builders
 haftmann parents: 
34071diff
changeset | 714 | literal_list = enum ";" "[" "]", | 
| 28064 | 715 | infix_cons = (6, "::") | 
| 716 | } end; | |
| 717 | ||
| 718 | ||
| 28054 | 719 | |
| 720 | (** SML/OCaml generic part **) | |
| 721 | ||
| 55681 | 722 | fun ml_program_of_program ctxt module_name reserved identifiers = | 
| 39028 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 723 | let | 
| 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 724 | fun namify_const upper base (nsp_const, nsp_type) = | 
| 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 725 | let | 
| 43326 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 wenzelm parents: 
42599diff
changeset | 726 | val (base', nsp_const') = | 
| 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 wenzelm parents: 
42599diff
changeset | 727 | Name.variant (if upper then first_upper base else base) nsp_const | 
| 39028 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 728 | in (base', (nsp_const', nsp_type)) end; | 
| 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 729 | fun namify_type base (nsp_const, nsp_type) = | 
| 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 730 | let | 
| 43326 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 wenzelm parents: 
42599diff
changeset | 731 | val (base', nsp_type') = Name.variant base nsp_type | 
| 39028 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 732 | in (base', (nsp_const, nsp_type')) end; | 
| 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 733 | fun namify_stmt (Code_Thingol.Fun _) = namify_const false | 
| 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 734 | | namify_stmt (Code_Thingol.Datatype _) = namify_type | 
| 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 735 | | namify_stmt (Code_Thingol.Datatypecons _) = namify_const true | 
| 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 736 | | namify_stmt (Code_Thingol.Class _) = namify_type | 
| 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 737 | | namify_stmt (Code_Thingol.Classrel _) = namify_const false | 
| 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 738 | | namify_stmt (Code_Thingol.Classparam _) = namify_const false | 
| 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 739 | | namify_stmt (Code_Thingol.Classinst _) = namify_const false; | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 740 | fun ml_binding_of_stmt (sym as Constant const, (export, Code_Thingol.Fun ((tysm as (vs, ty), raw_eqs), _))) = | 
| 39028 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 741 | let | 
| 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 742 | val eqs = filter (snd o snd) raw_eqs; | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 743 | val (eqs', some_sym) = if null (filter_out (null o snd) vs) then case eqs | 
| 39028 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 744 | of [(([], t), some_thm)] => if (not o null o fst o Code_Thingol.unfold_fun) ty | 
| 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 745 | then ([(([IVar (SOME "x")], t `$ IVar (SOME "x")), some_thm)], NONE) | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 746 | else (eqs, SOME (sym, member (op =) (Code_Thingol.add_constsyms t []) sym)) | 
| 39028 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 747 | | _ => (eqs, NONE) | 
| 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 748 | else (eqs, NONE) | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 749 | in ((export, ML_Function (const, (tysm, eqs'))), some_sym) end | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 750 |       | ml_binding_of_stmt (sym as Class_Instance inst, (export, Code_Thingol.Classinst (stmt as { vs, ... }))) =
 | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 751 | ((export, ML_Instance (inst, stmt)), | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 752 | if forall (null o snd) vs then SOME (sym, false) else NONE) | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 753 | | ml_binding_of_stmt (sym, _) = | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51143diff
changeset | 754 |           error ("Binding block containing illegal statement: " ^ 
 | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 755 | Code_Symbol.quote ctxt sym) | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 756 | fun modify_fun (sym, (export, stmt)) = | 
| 39028 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 757 | let | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 758 | val ((export', binding), some_value_sym) = ml_binding_of_stmt (sym, (export, stmt)); | 
| 39028 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 759 | val ml_stmt = case binding | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 760 | of ML_Function (const, ((vs, ty), [])) => | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 761 | ML_Exc (const, ((vs, ty), | 
| 39028 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 762 | (length o filter_out (null o snd)) vs + (length o fst o Code_Thingol.unfold_fun) ty)) | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 763 | | _ => case some_value_sym | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 764 | of NONE => ML_Funs ([(export', binding)], []) | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 765 | | SOME (sym, true) => ML_Funs ([(export, binding)], [sym]) | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 766 | | SOME (sym, false) => ML_Val binding | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 767 | in SOME (export, ml_stmt) end; | 
| 39031 | 768 | fun modify_funs stmts = single (SOME | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 769 | (Code_Namespace.Opaque, ML_Funs (map_split ml_binding_of_stmt stmts |> (apsnd o map_filter o Option.map) fst))) | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 770 | fun modify_datatypes stmts = | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 771 | map_filter | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 772 | (fn (Type_Constructor tyco, (export, Code_Thingol.Datatype stmt)) => SOME (export, (tyco, stmt)) | _ => NONE) stmts | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 773 | |> split_list | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 774 | |> apfst Code_Namespace.join_exports | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 775 | |> apsnd ML_Datas | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 776 | |> SOME | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 777 | |> single; | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 778 | fun modify_class stmts = | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 779 | the_single (map_filter | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 780 | (fn (Type_Class class, (export, Code_Thingol.Class stmt)) => SOME (export, (class, stmt)) | _ => NONE) stmts) | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 781 | |> apsnd ML_Class | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 782 | |> SOME | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 783 | |> single; | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 784 | fun modify_stmts ([stmt as (_, (_, stmt' as Code_Thingol.Fun _))]) = | 
| 39059 
3a11a667af75
restored and added surpression of case combinators
 haftmann parents: 
39058diff
changeset | 785 | if Code_Thingol.is_case stmt' then [] else [modify_fun stmt] | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 786 | | modify_stmts ((stmts as (_, (_, Code_Thingol.Fun _)) :: _)) = | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 787 | modify_funs (filter_out (Code_Thingol.is_case o snd o snd) stmts) | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 788 | | modify_stmts ((stmts as (_, (_, Code_Thingol.Datatypecons _)) :: _)) = | 
| 39028 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 789 | modify_datatypes stmts | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 790 | | modify_stmts ((stmts as (_, (_, Code_Thingol.Datatype _)) :: _)) = | 
| 39028 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 791 | modify_datatypes stmts | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 792 | | modify_stmts ((stmts as (_, (_, Code_Thingol.Class _)) :: _)) = | 
| 39028 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 793 | modify_class stmts | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 794 | | modify_stmts ((stmts as (_, (_, Code_Thingol.Classrel _)) :: _)) = | 
| 39028 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 795 | modify_class stmts | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 796 | | modify_stmts ((stmts as (_, (_, Code_Thingol.Classparam _)) :: _)) = | 
| 39028 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 797 | modify_class stmts | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 798 | | modify_stmts ([stmt as (_, (_, Code_Thingol.Classinst _))]) = | 
| 39028 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 799 | [modify_fun stmt] | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 800 | | modify_stmts ((stmts as (_, (_, Code_Thingol.Classinst _)) :: _)) = | 
| 39028 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 801 | modify_funs stmts | 
| 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 802 |       | modify_stmts stmts = error ("Illegal mutual dependencies: " ^
 | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 803 | (Library.commas o map (Code_Symbol.quote ctxt o fst)) stmts); | 
| 39028 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 804 | in | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 805 |     Code_Namespace.hierarchical_program ctxt {
 | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51143diff
changeset | 806 | module_name = module_name, reserved = reserved, identifiers = identifiers, | 
| 39028 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 807 | empty_nsp = (reserved, reserved), namify_module = pair, namify_stmt = namify_stmt, | 
| 55776 
7dd1971b39c1
amended some slips, rolling back currently dysfunctional export minimimalisation for Scala
 haftmann parents: 
55684diff
changeset | 808 | cyclic_modules = false, class_transitive = true, | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 809 | class_relation_public = true, empty_data = (), | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 810 | memorize_data = K I, modify_stmts = modify_stmts } | 
| 39028 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 811 | end; | 
| 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 812 | |
| 55677 | 813 | fun serialize_ml print_ml_module print_ml_stmt ctxt | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 814 |     { module_name, reserved_syms, identifiers, includes,
 | 
| 55683 | 815 | class_syntax, tyco_syntax, const_syntax } exports program = | 
| 28054 | 816 | let | 
| 39147 | 817 | |
| 818 | (* build program *) | |
| 39028 
0dd6c5a0beef
use code_namespace module for SML and OCaml code
 haftmann parents: 
38966diff
changeset | 819 |     val { deresolver, hierarchical_program = ml_program } =
 | 
| 55681 | 820 | ml_program_of_program ctxt module_name (Name.make_context reserved_syms) | 
| 55683 | 821 | identifiers exports program; | 
| 39147 | 822 | |
| 823 | (* print statements *) | |
| 55679 | 824 | fun print_stmt prefix_fragments (_, (export, stmt)) = print_ml_stmt | 
| 47576 | 825 | tyco_syntax const_syntax (make_vars reserved_syms) | 
| 55681 | 826 | (Code_Thingol.is_constr program) (deresolver prefix_fragments) export stmt | 
| 827 | |> apfst (fn decl => if Code_Namespace.not_private export then SOME decl else NONE); | |
| 39147 | 828 | |
| 829 | (* print modules *) | |
| 47576 | 830 | fun print_module _ base _ xs = | 
| 39147 | 831 | let | 
| 832 | val (raw_decls, body) = split_list xs; | |
| 55677 | 833 | val decls = maps these raw_decls | 
| 39147 | 834 | in (NONE, print_ml_module base decls body) end; | 
| 835 | ||
| 836 | (* serialization *) | |
| 837 | val p = Pretty.chunks2 (map snd includes | |
| 838 |       @ map snd (Code_Namespace.print_hierarchical {
 | |
| 839 | print_module = print_module, print_stmt = print_stmt, | |
| 840 | lift_markup = apsnd } ml_program)); | |
| 39056 | 841 | fun write width NONE = writeln o format [] width | 
| 842 | | write width (SOME p) = File.write p o format [] width; | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 843 |     fun prepare syms width p = ([("", format syms width p)], try (deresolver []));
 | 
| 28054 | 844 | in | 
| 48568 | 845 | Code_Target.serialization write prepare p | 
| 28054 | 846 | end; | 
| 847 | ||
| 38966 | 848 | val serializer_sml : Code_Target.serializer = | 
| 55677 | 849 | Code_Target.parse_args (Scan.succeed ()) #> K (serialize_ml print_sml_module print_sml_stmt); | 
| 28054 | 850 | |
| 38966 | 851 | val serializer_ocaml : Code_Target.serializer = | 
| 55677 | 852 | Code_Target.parse_args (Scan.succeed ()) #> K (serialize_ml print_ocaml_module print_ocaml_stmt); | 
| 28054 | 853 | |
| 38966 | 854 | |
| 855 | (** Isar setup **) | |
| 856 | ||
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52138diff
changeset | 857 | fun fun_syntax print_typ fxy [ty1, ty2] = | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52138diff
changeset | 858 | brackify_infix (1, R) fxy ( | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52138diff
changeset | 859 | print_typ (INFX (1, X)) ty1, | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52138diff
changeset | 860 | str "->", | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52138diff
changeset | 861 | print_typ (INFX (1, R)) ty2 | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52138diff
changeset | 862 | ); | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52138diff
changeset | 863 | |
| 28054 | 864 | val setup = | 
| 37821 | 865 | Code_Target.add_target | 
| 38966 | 866 |     (target_SML, { serializer = serializer_sml, literals = literals_sml,
 | 
| 41940 | 867 |       check = { env_var = "ISABELLE_PROCESS",
 | 
| 868 | make_destination = fn p => Path.append p (Path.explode "ROOT.ML"), | |
| 50022 
286dfcab9833
restored SML code check which got unintentionally broken: must explicitly check for error during compilation;
 haftmann parents: 
48568diff
changeset | 869 | make_command = fn _ => | 
| 51091 
c007c6bf4a35
another attempt for a uniform abort on code generation errors
 haftmann parents: 
50625diff
changeset | 870 | "\"$ISABELLE_PROCESS\" -r -q -e 'datatype ref = datatype Unsynchronized.ref; use \"ROOT.ML\" handle _ => exit 1' Pure" } }) | 
| 37821 | 871 | #> Code_Target.add_target | 
| 38966 | 872 |     (target_OCaml, { serializer = serializer_ocaml, literals = literals_ocaml,
 | 
| 41952 
c7297638599b
cleanup of former settings GHC_PATH, EXEC_GHC, EXEC_OCAML, EXEC_SWIPL, EXEC_YAP -- discontinued implicit detection;
 wenzelm parents: 
41940diff
changeset | 873 |       check = { env_var = "ISABELLE_OCAML",
 | 
| 41940 | 874 | make_destination = fn p => Path.append p (Path.explode "ROOT.ocaml"), | 
| 41952 
c7297638599b
cleanup of former settings GHC_PATH, EXEC_GHC, EXEC_OCAML, EXEC_SWIPL, EXEC_YAP -- discontinued implicit detection;
 wenzelm parents: 
41940diff
changeset | 875 | make_command = fn _ => "\"$ISABELLE_OCAML\" -w pu nums.cma ROOT.ocaml" } }) | 
| 55150 | 876 |   #> Code_Target.set_printings (Type_Constructor ("fun",
 | 
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52138diff
changeset | 877 | [(target_SML, SOME (2, fun_syntax)), (target_OCaml, SOME (2, fun_syntax))])) | 
| 28054 | 878 | #> fold (Code_Target.add_reserved target_SML) ML_Syntax.reserved_names | 
| 879 | #> fold (Code_Target.add_reserved target_SML) | |
| 38070 | 880 | ["ref" (*rebinding is illegal*), "o" (*dictionary projections use it already*), | 
| 881 | "Fail", "div", "mod" (*standard infixes*), "IntInf"] | |
| 28054 | 882 | #> fold (Code_Target.add_reserved target_OCaml) [ | 
| 883 | "and", "as", "assert", "begin", "class", | |
| 884 | "constraint", "do", "done", "downto", "else", "end", "exception", | |
| 885 | "external", "false", "for", "fun", "function", "functor", "if", | |
| 886 | "in", "include", "inherit", "initializer", "lazy", "let", "match", "method", | |
| 887 | "module", "mutable", "new", "object", "of", "open", "or", "private", "rec", | |
| 888 | "sig", "struct", "then", "to", "true", "try", "type", "val", | |
| 889 | "virtual", "when", "while", "with" | |
| 890 | ] | |
| 34944 
970e1466028d
code literals: distinguish numeral classes by different entries
 haftmann parents: 
34246diff
changeset | 891 | #> fold (Code_Target.add_reserved target_OCaml) ["failwith", "mod", "Big_int"]; | 
| 28054 | 892 | |
| 893 | end; (*struct*) |