| author | wenzelm | 
| Sat, 29 Apr 2017 11:06:39 +0200 | |
| changeset 65633 | 826311fca263 | 
| parent 65531 | 24544e3f183d | 
| child 66326 | 9eb8a2d07852 | 
| permissions | -rw-r--r-- | 
| 37745 
6315b6426200
checking generated code for various target languages
 haftmann parents: 
37669diff
changeset | 1 | (* Title: Tools/Code/code_scala.ML | 
| 
6315b6426200
checking generated code for various target languages
 haftmann parents: 
37669diff
changeset | 2 | Author: Florian Haftmann, TU Muenchen | 
| 34294 | 3 | |
| 4 | Serializer for Scala. | |
| 5 | *) | |
| 6 | ||
| 7 | signature CODE_SCALA = | |
| 8 | sig | |
| 37745 
6315b6426200
checking generated code for various target languages
 haftmann parents: 
37669diff
changeset | 9 | val target: string | 
| 58520 
a4d1f8041af0
accomplish potentially case-insenstive file systems for Scala
 haftmann parents: 
58398diff
changeset | 10 | val case_insensitive: bool Config.T; | 
| 34294 | 11 | end; | 
| 12 | ||
| 13 | structure Code_Scala : CODE_SCALA = | |
| 14 | struct | |
| 15 | ||
| 55150 | 16 | open Basic_Code_Symbol; | 
| 34294 | 17 | open Basic_Code_Thingol; | 
| 18 | open Code_Printer; | |
| 19 | ||
| 20 | infixr 5 @@; | |
| 21 | infixr 5 @|; | |
| 22 | ||
| 58520 
a4d1f8041af0
accomplish potentially case-insenstive file systems for Scala
 haftmann parents: 
58398diff
changeset | 23 | (** preliminary: option to circumvent clashes on case-insensitive file systems **) | 
| 
a4d1f8041af0
accomplish potentially case-insenstive file systems for Scala
 haftmann parents: 
58398diff
changeset | 24 | |
| 
a4d1f8041af0
accomplish potentially case-insenstive file systems for Scala
 haftmann parents: 
58398diff
changeset | 25 | val case_insensitive = Attrib.setup_config_bool @{binding "scala_case_insensitive"} (K false);
 | 
| 
a4d1f8041af0
accomplish potentially case-insenstive file systems for Scala
 haftmann parents: 
58398diff
changeset | 26 | |
| 34294 | 27 | |
| 28 | (** Scala serializer **) | |
| 29 | ||
| 63350 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 30 | val target = "Scala"; | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 31 | |
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 32 | datatype scala_stmt = Fun of typscheme * ((iterm list * iterm) * (thm option * bool)) list | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 33 | | Datatype of vname list * ((string * vname list) * itype list) list | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 34 | | Class of (vname * ((class * class) list * (string * itype) list)) | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 35 |       * (string * { vs: (vname * sort) list,
 | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 36 | inst_params: ((string * (const * int)) * (thm * bool)) list, | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 37 | superinst_params: ((string * (const * int)) * (thm * bool)) list }) list; | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 38 | |
| 47609 | 39 | fun print_scala_stmt tyco_syntax const_syntax reserved | 
| 50626 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 haftmann parents: 
48568diff
changeset | 40 | args_num is_constr (deresolve, deresolve_full) = | 
| 34294 | 41 | let | 
| 55150 | 42 | val deresolve_const = deresolve o Constant; | 
| 43 | val deresolve_tyco = deresolve o Type_Constructor; | |
| 44 | val deresolve_class = deresolve o Type_Class; | |
| 56812 | 45 | fun lookup_tyvar tyvars = lookup_var tyvars o Name.enforce_case true; | 
| 46 | fun intro_tyvars vs = intro_vars (map (Name.enforce_case true o fst) vs); | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 47 | fun print_tyco_expr tyvars fxy (sym, tys) = applify "[" "]" | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 48 | (print_typ tyvars NOBR) fxy ((str o deresolve) sym) tys | 
| 38923 | 49 | and print_typ tyvars fxy (tyco `%% tys) = (case tyco_syntax tyco | 
| 55150 | 50 | of NONE => print_tyco_expr tyvars fxy (Type_Constructor tyco, tys) | 
| 47609 | 51 | | SOME (_, print) => print (print_typ tyvars) fxy tys) | 
| 37243 
6e2ac5358d6e
capitalized type variables; added yield as keyword
 haftmann parents: 
37224diff
changeset | 52 | | print_typ tyvars fxy (ITyVar v) = (str o lookup_tyvar tyvars) v; | 
| 55150 | 53 | fun print_dicttyp tyvars (class, ty) = print_tyco_expr tyvars NOBR (Type_Class class, [ty]); | 
| 37639 | 54 | fun print_tupled_typ tyvars ([], ty) = | 
| 55 | print_typ tyvars NOBR ty | |
| 56 | | print_tupled_typ tyvars ([ty1], ty2) = | |
| 57 | concat [print_typ tyvars BR ty1, str "=>", print_typ tyvars NOBR ty2] | |
| 58 | | print_tupled_typ tyvars (tys, ty) = | |
| 59 |           concat [enum "," "(" ")" (map (print_typ tyvars NOBR) tys),
 | |
| 60 | str "=>", print_typ tyvars NOBR ty]; | |
| 61 | fun constraint p1 p2 = Pretty.block [p1, str ":", Pretty.brk 1, p2]; | |
| 34294 | 62 | fun print_var vars NONE = str "_" | 
| 63303 | 63 | | print_var vars (SOME v) = (str o lookup_var vars) v; | 
| 64 | fun applify_dict tyvars (Dict (_, d)) = applify_plain_dict tyvars d | |
| 65 | and applify_plain_dict tyvars (Dict_Const (inst, dss)) = | |
| 66 | applify_dictss tyvars ((str o deresolve o Class_Instance) inst) dss | |
| 67 |       | applify_plain_dict tyvars (Dict_Var { var, class, ... }) =
 | |
| 68 | Pretty.block [str "implicitly", | |
| 69 | enclose "[" "]" [Pretty.block [(str o deresolve_class) class, | |
| 70 | enclose "[" "]" [(str o lookup_tyvar tyvars) var]]]] | |
| 71 | and applify_dictss tyvars p dss = | |
| 72 |       applify "(" ")" (applify_dict tyvars) NOBR p (flat dss)
 | |
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 73 | fun print_term tyvars is_pat some_thm vars fxy (IConst const) = | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 74 | print_app tyvars is_pat some_thm vars fxy (const, []) | 
| 35228 | 75 | | print_term tyvars is_pat some_thm vars fxy (t as (t1 `$ t2)) = | 
| 34294 | 76 | (case Code_Thingol.unfold_const_app t | 
| 35228 | 77 | of SOME app => print_app tyvars is_pat some_thm vars fxy app | 
| 37639 | 78 |             | _ => applify "(" ")" (print_term tyvars is_pat some_thm vars NOBR) fxy
 | 
| 79 | (print_term tyvars is_pat some_thm vars BR t1) [t2]) | |
| 35228 | 80 | | print_term tyvars is_pat some_thm vars fxy (IVar v) = | 
| 34294 | 81 | print_var vars v | 
| 61130 
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
 haftmann parents: 
59323diff
changeset | 82 | | print_term tyvars is_pat some_thm vars fxy (t as _ `|=> _) = | 
| 34294 | 83 | let | 
| 61130 
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
 haftmann parents: 
59323diff
changeset | 84 | val (vs_tys, body) = Code_Thingol.unfold_abs t; | 
| 
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
 haftmann parents: 
59323diff
changeset | 85 | val (ps, vars') = fold_map (print_abs_head tyvars) vs_tys vars; | 
| 
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
 haftmann parents: 
59323diff
changeset | 86 | val vars' = intro_vars (map_filter fst vs_tys) vars; | 
| 34294 | 87 | in | 
| 61130 
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
 haftmann parents: 
59323diff
changeset | 88 | brackets (ps @| print_term tyvars false some_thm vars' NOBR body) | 
| 41939 | 89 | end | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 90 | | print_term tyvars is_pat some_thm vars fxy (ICase case_expr) = | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 91 | (case Code_Thingol.unfold_const_app (#primitive case_expr) | 
| 55150 | 92 |            of SOME (app as ({ sym = Constant const, ... }, _)) =>
 | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 93 | if is_none (const_syntax const) | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 94 | then print_case tyvars some_thm vars fxy case_expr | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 95 | else print_app tyvars is_pat some_thm vars fxy app | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 96 | | NONE => print_case tyvars some_thm vars fxy case_expr) | 
| 61130 
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
 haftmann parents: 
59323diff
changeset | 97 | and print_abs_head tyvars (some_v, ty) vars = | 
| 
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
 haftmann parents: 
59323diff
changeset | 98 | let | 
| 
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
 haftmann parents: 
59323diff
changeset | 99 | val vars' = intro_vars (the_list some_v) vars; | 
| 
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
 haftmann parents: 
59323diff
changeset | 100 | in | 
| 
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
 haftmann parents: 
59323diff
changeset | 101 | (concat [ | 
| 
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
 haftmann parents: 
59323diff
changeset | 102 |                enclose "(" ")" [constraint (print_var vars' some_v) (print_typ tyvars NOBR ty)],
 | 
| 
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
 haftmann parents: 
59323diff
changeset | 103 | str "=>" | 
| 
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
 haftmann parents: 
59323diff
changeset | 104 | ], vars') | 
| 
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
 haftmann parents: 
59323diff
changeset | 105 | end | 
| 37464 | 106 | and print_app tyvars is_pat some_thm vars fxy | 
| 63303 | 107 |         (app as ({ sym, typargs, dom, dicts, ... }, ts)) =
 | 
| 34294 | 108 | let | 
| 109 | val k = length ts; | |
| 50626 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 haftmann parents: 
48568diff
changeset | 110 | val typargs' = if is_pat then [] else typargs; | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 111 | val syntax = case sym of | 
| 55150 | 112 | Constant const => const_syntax const | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 113 | | _ => NONE; | 
| 63303 | 114 | val applify_dicts = | 
| 115 | if is_pat orelse is_some syntax orelse is_constr sym | |
| 116 | orelse Code_Thingol.unambiguous_dictss dicts | |
| 117 | then fn p => K p | |
| 118 | else applify_dictss tyvars; | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 119 | val (l, print') = case syntax of | 
| 63303 | 120 | NONE => (args_num sym, fn fxy => fn ts => applify_dicts | 
| 121 |               (gen_applify (is_constr sym) "(" ")"
 | |
| 37639 | 122 | (print_term tyvars is_pat some_thm vars NOBR) fxy | 
| 123 | (applify "[" "]" (print_typ tyvars NOBR) | |
| 63303 | 124 | NOBR ((str o deresolve) sym) typargs') ts) dicts) | 
| 125 | | SOME (k, Plain_printer s) => (k, fn fxy => fn ts => applify_dicts | |
| 126 |               (applify "(" ")"
 | |
| 37881 | 127 | (print_term tyvars is_pat some_thm vars NOBR) fxy | 
| 128 | (applify "[" "]" (print_typ tyvars NOBR) | |
| 63303 | 129 | NOBR (str s) typargs') ts) dicts) | 
| 55153 
eedd549de3ef
more suitable names, without any notion of "activating"
 haftmann parents: 
55150diff
changeset | 130 | | SOME (k, Complex_printer print) => | 
| 38059 | 131 | (k, fn fxy => fn ts => print (print_term tyvars is_pat some_thm) some_thm vars fxy | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 132 | (ts ~~ take k dom)) | 
| 38059 | 133 | in if k = l then print' fxy ts | 
| 34294 | 134 | else if k < l then | 
| 35228 | 135 | print_term tyvars is_pat some_thm vars fxy (Code_Thingol.eta_expand l app) | 
| 34294 | 136 | else let | 
| 137 | val (ts1, ts23) = chop l ts; | |
| 138 | in | |
| 38059 | 139 | Pretty.block (print' BR ts1 :: map (fn t => Pretty.block | 
| 35228 | 140 |           [str ".apply(", print_term tyvars is_pat some_thm vars NOBR t, str ")"]) ts23)
 | 
| 34294 | 141 | end end | 
| 37464 | 142 | and print_bind tyvars some_thm fxy p = | 
| 143 | gen_print_bind (print_term tyvars true) some_thm fxy p | |
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 144 |     and print_case tyvars some_thm vars fxy { clauses = [], ... } =
 | 
| 48073 
1b609a7837ef
prefer sys.error over plain error in Scala to avoid deprecation warning
 haftmann parents: 
48072diff
changeset | 145 | (brackify fxy o Pretty.breaks o map str) ["sys.error(\"empty case\")"] | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 146 |       | print_case tyvars some_thm vars fxy (case_expr as { clauses = [_], ... }) =
 | 
| 34294 | 147 | let | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 148 | val (bind :: binds, body) = Code_Thingol.unfold_let (ICase case_expr); | 
| 41781 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 149 | fun print_match_val ((pat, ty), t) vars = | 
| 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 150 | vars | 
| 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 151 | |> print_bind tyvars some_thm BR pat | 
| 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 152 | |>> (fn p => (false, concat [str "val", constraint p (print_typ tyvars NOBR ty), | 
| 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 153 | str "=", print_term tyvars false some_thm vars NOBR t])); | 
| 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 154 | fun print_match_seq t vars = | 
| 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 155 | ((true, print_term tyvars false some_thm vars NOBR t), vars); | 
| 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 156 | fun print_match is_first ((IVar NONE, ty), t) = | 
| 41784 | 157 | if Code_Thingol.is_IAbs t andalso is_first | 
| 41781 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 158 | then print_match_val ((IVar NONE, ty), t) | 
| 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 159 | else print_match_seq t | 
| 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 160 | | print_match _ ((pat, ty), t) = | 
| 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 161 | print_match_val ((pat, ty), t); | 
| 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 162 | val (seps_ps, vars') = | 
| 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 163 | vars |> print_match true bind ||>> fold_map (print_match false) binds |>> uncurry cons; | 
| 38059 | 164 | val all_seps_ps = seps_ps @ [(true, print_term tyvars false some_thm vars' NOBR body)]; | 
| 165 | fun insert_seps [(_, p)] = [p] | |
| 166 | | insert_seps ((_, p) :: (seps_ps as (sep, _) :: _)) = | |
| 167 | (if sep then Pretty.block [p, str ";"] else p) :: insert_seps seps_ps | |
| 41781 
32a7726d2136
more idiomatic printing of let cascades and type variable constraints
 haftmann parents: 
41687diff
changeset | 168 |           in brackify_block fxy (str "{") (insert_seps all_seps_ps) (str "}") end
 | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 169 |       | print_case tyvars some_thm vars fxy { term = t, typ = ty, clauses = clauses as _ :: _, ... } =
 | 
| 34294 | 170 | let | 
| 171 | fun print_select (pat, body) = | |
| 172 | let | |
| 37464 | 173 | val (p_pat, vars') = print_bind tyvars some_thm NOBR pat vars; | 
| 174 | val p_body = print_term tyvars false some_thm vars' NOBR body | |
| 175 | in concat [str "case", p_pat, str "=>", p_body] end; | |
| 46850 | 176 | in | 
| 177 | map print_select clauses | |
| 178 |             |> Pretty.block_enclose (concat [print_term tyvars false some_thm vars NOBR t, str "match", str "{"], str "}")
 | |
| 179 | |> single | |
| 180 |             |> enclose "(" ")"
 | |
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 181 | end; | 
| 63350 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 182 | fun print_context tyvars vs s = applify "[" "]" | 
| 37639 | 183 | (fn (v, sort) => (Pretty.block o map str) | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 184 | (lookup_tyvar tyvars v :: maps (fn class => [" : ", deresolve_class class]) sort)) | 
| 63350 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 185 | NOBR (str s) vs; | 
| 55679 | 186 | fun print_defhead export tyvars vars const vs params tys ty = | 
| 55776 
7dd1971b39c1
amended some slips, rolling back currently dysfunctional export minimimalisation for Scala
 haftmann parents: 
55684diff
changeset | 187 |       concat [str "def", constraint (applify "(" ")" (fn (param, ty) =>
 | 
| 37639 | 188 | constraint ((str o lookup_var vars) param) (print_typ tyvars NOBR ty)) | 
| 63350 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 189 | NOBR (print_context tyvars vs (deresolve_const const)) (params ~~ tys)) (print_typ tyvars NOBR ty), | 
| 55679 | 190 | str "="]; | 
| 191 | fun print_def export const (vs, ty) [] = | |
| 37639 | 192 | let | 
| 193 | val (tys, ty') = Code_Thingol.unfold_fun ty; | |
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43326diff
changeset | 194 | val params = Name.invent (snd reserved) "a" (length tys); | 
| 37639 | 195 | val tyvars = intro_tyvars vs reserved; | 
| 196 | val vars = intro_vars params reserved; | |
| 197 | in | |
| 55679 | 198 | concat [print_defhead export tyvars vars const vs params tys ty', | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 199 |               str ("sys.error(\"" ^ const ^ "\")")]
 | 
| 37639 | 200 | end | 
| 55679 | 201 | | print_def export const (vs, ty) eqs = | 
| 37639 | 202 | let | 
| 203 | val tycos = fold (fn ((ts, t), _) => | |
| 204 | fold Code_Thingol.add_tyconames (t :: ts)) eqs []; | |
| 205 | val tyvars = reserved | |
| 206 | |> intro_base_names | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 207 | (is_none o tyco_syntax) deresolve_tyco tycos | 
| 37639 | 208 | |> intro_tyvars vs; | 
| 209 | val simple = case eqs | |
| 210 | of [((ts, _), _)] => forall Code_Thingol.is_IVar ts | |
| 211 | | _ => false; | |
| 212 | val vars1 = reserved | |
| 55145 
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
 haftmann parents: 
52520diff
changeset | 213 | |> intro_base_names_for (is_none o const_syntax) | 
| 
2bb3cd36bcf7
more abstract declaration of unqualified constant names in code printing context
 haftmann parents: 
52520diff
changeset | 214 | deresolve (map (snd o fst) eqs); | 
| 37639 | 215 | val params = if simple | 
| 216 | then (map (fn IVar (SOME x) => x) o fst o fst o hd) eqs | |
| 217 | else aux_params vars1 (map (fst o fst) eqs); | |
| 218 | val vars2 = intro_vars params vars1; | |
| 219 | val (tys', ty') = Code_Thingol.unfold_fun_n (length params) ty; | |
| 38922 | 220 | fun tuplify [p] = p | 
| 221 |               | tuplify ps = enum "," "(" ")" ps;
 | |
| 37639 | 222 | fun print_rhs vars' ((_, t), (some_thm, _)) = | 
| 223 | print_term tyvars false some_thm vars' NOBR t; | |
| 224 | fun print_clause (eq as ((ts, _), (some_thm, _))) = | |
| 34294 | 225 | let | 
| 37639 | 226 | val vars' = intro_vars ((fold o Code_Thingol.fold_varnames) | 
| 227 | (insert (op =)) ts []) vars1; | |
| 228 | in | |
| 229 | concat [str "case", | |
| 38922 | 230 | tuplify (map (print_term tyvars true some_thm vars' NOBR) ts), | 
| 37639 | 231 | str "=>", print_rhs vars' eq] | 
| 232 | end; | |
| 55679 | 233 | val head = print_defhead export tyvars vars2 const vs params tys' ty'; | 
| 37639 | 234 | in if simple then | 
| 235 | concat [head, print_rhs vars2 (hd eqs)] | |
| 236 | else | |
| 237 | Pretty.block_enclose | |
| 38922 | 238 | (concat [head, tuplify (map (str o lookup_var vars2) params), | 
| 37639 | 239 |                 str "match", str "{"], str "}")
 | 
| 240 | (map print_clause eqs) | |
| 241 | end; | |
| 55150 | 242 | val print_method = str o Library.enclose "`" "`" o deresolve_full o Constant; | 
| 63350 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 243 |     fun print_inst class (tyco, { vs, inst_params, superinst_params }) =
 | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 244 | let | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 245 | val tyvars = intro_tyvars vs reserved; | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 246 | val classtyp = (class, tyco `%% map (ITyVar o fst) vs); | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 247 |         fun print_classparam_instance ((classparam, (const as { dom, ... }, dom_length)), (thm, _)) =
 | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 248 | let | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 249 | val aux_dom = Name.invent_names (snd reserved) "a" dom; | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 250 | val auxs = map fst aux_dom; | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 251 | val vars = intro_vars auxs reserved; | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 252 | val (aux_dom1, aux_dom2) = chop dom_length aux_dom; | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 253 | fun abstract_using [] = [] | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 254 |               | abstract_using aux_dom = [enum "," "(" ")"
 | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 255 | (map (fn (aux, ty) => constraint ((str o lookup_var vars) aux) | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 256 | (print_typ tyvars NOBR ty)) aux_dom), str "=>"]; | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 257 | val aux_abstr1 = abstract_using aux_dom1; | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 258 | val aux_abstr2 = abstract_using aux_dom2; | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 259 | in | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 260 | concat ([str "val", print_method classparam, str "="] | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 261 | @ aux_abstr1 @ aux_abstr2 @| print_app tyvars false (SOME thm) vars NOBR | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 262 | (const, map (IVar o SOME) auxs)) | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 263 | end; | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 264 | in | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 265 | Pretty.block_enclose (concat [str "implicit def", | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 266 | constraint (print_context tyvars vs | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 267 | ((Library.enclose "`" "`" o deresolve_full o Class_Instance) (tyco, class))) | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 268 | (print_dicttyp tyvars classtyp), | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 269 |           str "=", str "new", print_dicttyp tyvars classtyp, str "{"], str "}")
 | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 270 | (map print_classparam_instance (inst_params @ superinst_params)) | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 271 | end; | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 272 | fun print_stmt (Constant const, (export, Fun ((vs, ty), raw_eqs))) = | 
| 55679 | 273 | print_def export const (vs, ty) (filter (snd o snd) raw_eqs) | 
| 63350 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 274 | | print_stmt (Type_Constructor tyco, (export, Datatype (vs, cos))) = | 
| 34294 | 275 | let | 
| 48003 
1d11af40b106
dropped sort constraints on datatype specifications
 haftmann parents: 
47609diff
changeset | 276 | val tyvars = intro_tyvars (map (rpair []) vs) reserved; | 
| 50626 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 haftmann parents: 
48568diff
changeset | 277 | fun print_co ((co, vs_args), tys) = | 
| 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 haftmann parents: 
48568diff
changeset | 278 | concat [Pretty.block ((applify "[" "]" (str o lookup_tyvar tyvars) NOBR | 
| 55776 
7dd1971b39c1
amended some slips, rolling back currently dysfunctional export minimimalisation for Scala
 haftmann parents: 
55684diff
changeset | 279 | ((concat o map str) ["final", "case", "class", deresolve_const co]) vs_args) | 
| 50626 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 haftmann parents: 
48568diff
changeset | 280 |                 @@ enum "," "(" ")" (map (fn (v, arg) => constraint (str v) (print_typ tyvars NOBR arg))
 | 
| 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 haftmann parents: 
48568diff
changeset | 281 | (Name.invent_names (snd reserved) "a" tys))), | 
| 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 haftmann parents: 
48568diff
changeset | 282 | str "extends", | 
| 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 haftmann parents: 
48568diff
changeset | 283 | applify "[" "]" (str o lookup_tyvar tyvars) NOBR | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 284 | ((str o deresolve_tyco) tyco) vs | 
| 50626 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 haftmann parents: 
48568diff
changeset | 285 | ]; | 
| 34294 | 286 | in | 
| 50626 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 haftmann parents: 
48568diff
changeset | 287 | Pretty.chunks (applify "[" "]" (str o lookup_tyvar tyvars) | 
| 55776 
7dd1971b39c1
amended some slips, rolling back currently dysfunctional export minimimalisation for Scala
 haftmann parents: 
55684diff
changeset | 288 | NOBR ((concat o map str) ["abstract", "sealed", "class", deresolve_tyco tyco]) vs | 
| 37639 | 289 | :: map print_co cos) | 
| 34294 | 290 | end | 
| 63350 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 291 | | print_stmt (Type_Class class, (export, Class ((v, (classrels, classparams)), insts))) = | 
| 34294 | 292 | let | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 293 | val tyvars = intro_tyvars [(v, [class])] reserved; | 
| 37639 | 294 | fun add_typarg s = Pretty.block | 
| 295 | [str s, str "[", (str o lookup_tyvar tyvars) v, str "]"]; | |
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37337diff
changeset | 296 | fun print_super_classes [] = NONE | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 297 | | print_super_classes classrels = SOME (concat (str "extends" | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 298 | :: separate (str "with") (map (add_typarg o deresolve_class o snd) classrels))); | 
| 34294 | 299 | fun print_classparam_val (classparam, ty) = | 
| 37639 | 300 | concat [str "val", constraint (print_method classparam) | 
| 301 | ((print_tupled_typ tyvars o Code_Thingol.unfold_fun) ty)]; | |
| 34294 | 302 | fun print_classparam_def (classparam, ty) = | 
| 303 | let | |
| 304 | val (tys, ty) = Code_Thingol.unfold_fun ty; | |
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43326diff
changeset | 305 | val [implicit_name] = Name.invent (snd reserved) (lookup_tyvar tyvars v) 1; | 
| 37639 | 306 | val proto_vars = intro_vars [implicit_name] reserved; | 
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43326diff
changeset | 307 | val auxs = Name.invent (snd proto_vars) "a" (length tys); | 
| 37639 | 308 | val vars = intro_vars auxs proto_vars; | 
| 34294 | 309 | in | 
| 55776 
7dd1971b39c1
amended some slips, rolling back currently dysfunctional export minimimalisation for Scala
 haftmann parents: 
55684diff
changeset | 310 |                 concat [str "def", constraint (Pretty.block [applify "(" ")"
 | 
| 37639 | 311 | (fn (aux, ty) => constraint ((str o lookup_var vars) aux) | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 312 | (print_typ tyvars NOBR ty)) NOBR (add_typarg (deresolve_const classparam)) | 
| 37639 | 313 | (auxs ~~ tys), str "(implicit ", str implicit_name, str ": ", | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 314 | add_typarg (deresolve_class class), str ")"]) (print_typ tyvars NOBR ty), str "=", | 
| 37639 | 315 |                   applify "(" ")" (str o lookup_var vars) NOBR
 | 
| 316 | (Pretty.block [str implicit_name, str ".", print_method classparam]) auxs] | |
| 34294 | 317 | end; | 
| 318 | in | |
| 319 | Pretty.chunks ( | |
| 320 | (Pretty.block_enclose | |
| 55776 
7dd1971b39c1
amended some slips, rolling back currently dysfunctional export minimimalisation for Scala
 haftmann parents: 
55684diff
changeset | 321 | (concat ([str "trait", (add_typarg o deresolve_class) class] | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 322 |                   @ the_list (print_super_classes classrels) @ [str "{"]), str "}")
 | 
| 34294 | 323 | (map print_classparam_val classparams)) | 
| 324 | :: map print_classparam_def classparams | |
| 63350 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 325 | @| Pretty.block_enclose | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 326 |                 ((concat o map str) ["object", deresolve_class class, "{"], str "}")
 | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 327 | (map (print_inst class) insts) | 
| 34294 | 328 | ) | 
| 37639 | 329 | end; | 
| 34294 | 330 | in print_stmt end; | 
| 331 | ||
| 63350 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 332 | fun pickup_instances_for_class program = | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 333 | let | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 334 | val tab = | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 335 | Symtab.empty | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 336 | |> Code_Symbol.Graph.fold | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 337 |         (fn (_, (Code_Thingol.Classinst { class, tyco, vs, inst_params, superinst_params, ... }, _)) =>
 | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 338 | Symtab.map_default (class, []) | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 339 |                 (cons (tyco, { vs = vs, inst_params = inst_params, superinst_params = superinst_params }))
 | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 340 | | _ => I) program; | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 341 | in Symtab.lookup_list tab end; | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 342 | |
| 55683 | 343 | fun scala_program_of_program ctxt module_name reserved identifiers exports program = | 
| 34294 | 344 | let | 
| 58520 
a4d1f8041af0
accomplish potentially case-insenstive file systems for Scala
 haftmann parents: 
58398diff
changeset | 345 | val variant = if Config.get ctxt case_insensitive | 
| 
a4d1f8041af0
accomplish potentially case-insenstive file systems for Scala
 haftmann parents: 
58398diff
changeset | 346 | then Code_Namespace.variant_case_insensitive | 
| 
a4d1f8041af0
accomplish potentially case-insenstive file systems for Scala
 haftmann parents: 
58398diff
changeset | 347 | else Name.variant; | 
| 38809 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 348 | fun namify_module name_fragment ((nsp_class, nsp_object), nsp_common) = | 
| 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 349 | let | 
| 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 350 | val declare = Name.declare name_fragment; | 
| 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 351 | in (name_fragment, ((declare nsp_class, declare nsp_object), declare nsp_common)) end; | 
| 38769 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 352 | fun namify_class base ((nsp_class, nsp_object), nsp_common) = | 
| 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 353 | let | 
| 58520 
a4d1f8041af0
accomplish potentially case-insenstive file systems for Scala
 haftmann parents: 
58398diff
changeset | 354 | val (base', nsp_class') = variant base nsp_class | 
| 38769 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 355 | in (base', ((nsp_class', nsp_object), Name.declare base' nsp_common)) end; | 
| 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 356 | fun namify_object base ((nsp_class, nsp_object), nsp_common) = | 
| 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 357 | let | 
| 58520 
a4d1f8041af0
accomplish potentially case-insenstive file systems for Scala
 haftmann parents: 
58398diff
changeset | 358 | val (base', nsp_object') = variant base nsp_object | 
| 38769 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 359 | in (base', ((nsp_class, nsp_object'), Name.declare base' nsp_common)) end; | 
| 56826 
ba18bd41e510
enforce case of identifiers only to accomodate strict language requirements (or clear separation of constructors from variables in the case of SML)
 haftmann parents: 
56812diff
changeset | 360 | fun namify_common base ((nsp_class, nsp_object), nsp_common) = | 
| 38769 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 361 | let | 
| 58520 
a4d1f8041af0
accomplish potentially case-insenstive file systems for Scala
 haftmann parents: 
58398diff
changeset | 362 | val (base', nsp_common') = variant base nsp_common | 
| 38769 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 363 | in | 
| 56826 
ba18bd41e510
enforce case of identifiers only to accomodate strict language requirements (or clear separation of constructors from variables in the case of SML)
 haftmann parents: 
56812diff
changeset | 364 | (base', ((Name.declare base' nsp_class, Name.declare base' nsp_object), nsp_common')) | 
| 38769 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 365 | end; | 
| 38809 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 366 | fun namify_stmt (Code_Thingol.Fun _) = namify_object | 
| 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 367 | | namify_stmt (Code_Thingol.Datatype _) = namify_class | 
| 56826 
ba18bd41e510
enforce case of identifiers only to accomodate strict language requirements (or clear separation of constructors from variables in the case of SML)
 haftmann parents: 
56812diff
changeset | 368 | | namify_stmt (Code_Thingol.Datatypecons _) = namify_common | 
| 38809 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 369 | | namify_stmt (Code_Thingol.Class _) = namify_class | 
| 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 370 | | namify_stmt (Code_Thingol.Classrel _) = namify_object | 
| 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 371 | | namify_stmt (Code_Thingol.Classparam _) = namify_object | 
| 56826 
ba18bd41e510
enforce case of identifiers only to accomodate strict language requirements (or clear separation of constructors from variables in the case of SML)
 haftmann parents: 
56812diff
changeset | 372 | | namify_stmt (Code_Thingol.Classinst _) = namify_common; | 
| 63350 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 373 | val pickup_instances = pickup_instances_for_class program; | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 374 | fun modify_stmt (_, (_, Code_Thingol.Fun (_, SOME _))) = NONE | 
| 63350 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 375 | | modify_stmt (_, (export, Code_Thingol.Fun (x, NONE))) = SOME (export, Fun x) | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 376 | | modify_stmt (_, (export, Code_Thingol.Datatype x)) = SOME (export, Datatype x) | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 377 | | modify_stmt (_, (_, Code_Thingol.Datatypecons _)) = NONE | 
| 63350 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 378 | | modify_stmt (Type_Class class, (export, Code_Thingol.Class x)) = | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 379 | SOME (export, Class (x, pickup_instances class)) | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 380 | | modify_stmt (_, (_, Code_Thingol.Classrel _)) = NONE | 
| 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 381 | | modify_stmt (_, (_, Code_Thingol.Classparam _)) = NONE | 
| 63350 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 382 | | modify_stmt (_, (_, Code_Thingol.Classinst _)) = NONE | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: 
38966diff
changeset | 383 | in | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 384 | 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 | 385 |       { module_name = module_name, reserved = reserved, identifiers = identifiers,
 | 
| 41939 | 386 | empty_nsp = ((reserved, reserved), reserved), namify_module = namify_module, | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 387 | namify_stmt = namify_stmt, cyclic_modules = true, | 
| 63350 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 388 | class_transitive = true, class_relation_public = false, empty_data = (), | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 389 | memorize_data = K I, modify_stmts = map modify_stmt } | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 390 | exports program | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: 
38966diff
changeset | 391 | end; | 
| 34294 | 392 | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 393 | fun serialize_scala ctxt { module_name, reserved_syms, identifiers,
 | 
| 55683 | 394 | includes, class_syntax, tyco_syntax, const_syntax } exports program = | 
| 34294 | 395 | let | 
| 38769 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 396 | |
| 38809 
7dc73a208722
proper namespace administration for hierarchical modules
 haftmann parents: 
38782diff
changeset | 397 | (* build program *) | 
| 39147 | 398 |     val { deresolver, hierarchical_program = scala_program } =
 | 
| 55681 | 399 | scala_program_of_program ctxt module_name (Name.make_context reserved_syms) | 
| 55683 | 400 | identifiers exports program; | 
| 38769 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 401 | |
| 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 402 | (* print statements *) | 
| 55150 | 403 | fun lookup_constr tyco constr = case Code_Symbol.Graph.get_node program (Type_Constructor tyco) | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 404 | of Code_Thingol.Datatype (_, constrs) => | 
| 37639 | 405 | the (AList.lookup (op = o apsnd fst) constrs constr); | 
| 55150 | 406 | fun classparams_of_class class = case Code_Symbol.Graph.get_node program (Type_Class class) | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 407 | of Code_Thingol.Class (_, (_, classparams)) => classparams; | 
| 55150 | 408 | fun args_num (sym as Constant const) = case Code_Symbol.Graph.get_node program sym | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 409 | of Code_Thingol.Fun (((_, ty), []), _) => | 
| 37464 | 410 | (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 | 411 | | Code_Thingol.Fun ((_, ((ts, _), _) :: _), _) => length ts | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 412 | | Code_Thingol.Datatypecons tyco => length (lookup_constr tyco const) | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 413 | | Code_Thingol.Classparam class => | 
| 37639 | 414 | (length o fst o Code_Thingol.unfold_fun o the o AList.lookup (op =) | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55145diff
changeset | 415 | (classparams_of_class class)) const; | 
| 47609 | 416 | fun print_stmt prefix_fragments = print_scala_stmt | 
| 39147 | 417 | tyco_syntax const_syntax (make_vars reserved_syms) args_num | 
| 50626 
e21485358c56
uniform parentheses for constructor -- necessary to accomodate scala 10
 haftmann parents: 
48568diff
changeset | 418 | (Code_Thingol.is_constr program) (deresolver prefix_fragments, deresolver []); | 
| 38769 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 419 | |
| 39147 | 420 | (* print modules *) | 
| 63350 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 421 | fun print_module _ base _ ps = Pretty.chunks2 | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 422 |       (str ("object " ^ base ^ " {")
 | 
| 
705229ed856e
compiling implicit instances into companion objects for classes avoids ambiguities
 haftmann parents: 
63304diff
changeset | 423 |         :: ps @| str ("} /* object " ^ base ^ " */"));
 | 
| 38769 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 424 | |
| 
317e64c886d2
preliminary implementation of hierarchical module name space
 haftmann parents: 
38059diff
changeset | 425 | (* serialization *) | 
| 39147 | 426 | val p = Pretty.chunks2 (map snd includes | 
| 427 |       @ Code_Namespace.print_hierarchical {
 | |
| 428 | print_module = print_module, print_stmt = print_stmt, | |
| 429 | lift_markup = I } scala_program); | |
| 39056 | 430 | fun write width NONE = writeln o format [] width | 
| 431 | | 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 | 432 |     fun prepare syms width p = ([("", format syms width p)], try (deresolver []));
 | 
| 34294 | 433 | in | 
| 48568 | 434 | Code_Target.serialization write prepare p | 
| 34294 | 435 | end; | 
| 436 | ||
| 38966 | 437 | val serializer : Code_Target.serializer = | 
| 438 | Code_Target.parse_args (Scan.succeed ()) #> K serialize_scala; | |
| 439 | ||
| 34294 | 440 | val literals = let | 
| 37224 | 441 | fun char_scala c = if c = "'" then "\\'" | 
| 442 | else if c = "\"" then "\\\"" | |
| 443 | else if c = "\\" then "\\\\" | |
| 444 | else let val k = ord c | |
| 63304 | 445 | in if k < 32 orelse k > 126 | 
| 64463 
bed02fca80b5
Scala "\u" notation uses hexadecimal, not octal (amending 00a135c0a17f);
 wenzelm parents: 
63350diff
changeset | 446 | then "\\u" ^ align_right "0" 4 (Int.fmt StringCvt.HEX k) else c end | 
| 58398 | 447 | fun numeral_scala k = | 
| 448 | if ~2147483647 < k andalso k <= 2147483647 | |
| 449 | then signed_string_of_int k | |
| 450 | else quote (signed_string_of_int k) | |
| 34294 | 451 | in Literals {
 | 
| 452 | literal_char = Library.enclose "'" "'" o char_scala, | |
| 453 | literal_string = quote o translate_string char_scala, | |
| 34944 
970e1466028d
code literals: distinguish numeral classes by different entries
 haftmann parents: 
34900diff
changeset | 454 |   literal_numeral = fn k => "BigInt(" ^ numeral_scala k ^ ")",
 | 
| 34888 | 455 |   literal_list = fn [] => str "Nil" | ps => Pretty.block [str "List", enum "," "(" ")" ps],
 | 
| 34294 | 456 | infix_cons = (6, "::") | 
| 457 | } end; | |
| 458 | ||
| 459 | ||
| 460 | (** Isar setup **) | |
| 461 | ||
| 59323 | 462 | val _ = Theory.setup | 
| 463 | (Code_Target.add_language | |
| 38966 | 464 |     (target, { serializer = serializer, literals = literals,
 | 
| 41939 | 465 |       check = { env_var = "SCALA_HOME",
 | 
| 466 | make_destination = fn p => Path.append p (Path.explode "ROOT.scala"), | |
| 41940 | 467 | make_command = fn _ => | 
| 65531 | 468 | "isabelle_scala scalac $ISABELLE_SCALAC_OPTIONS ROOT.scala" } }) | 
| 55150 | 469 |   #> 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: 
52229diff
changeset | 470 | [(target, SOME (2, fn print_typ => fn fxy => fn [ty1, ty2] => | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52229diff
changeset | 471 | 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: 
52229diff
changeset | 472 | print_typ BR ty1 (*product type vs. tupled arguments!*), | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52229diff
changeset | 473 | str "=>", | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52229diff
changeset | 474 | 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: 
52229diff
changeset | 475 | )))])) | 
| 34294 | 476 | #> fold (Code_Target.add_reserved target) [ | 
| 477 | "abstract", "case", "catch", "class", "def", "do", "else", "extends", "false", | |
| 478 | "final", "finally", "for", "forSome", "if", "implicit", "import", "lazy", | |
| 479 | "match", "new", "null", "object", "override", "package", "private", "protected", | |
| 480 | "requires", "return", "sealed", "super", "this", "throw", "trait", "try", | |
| 37243 
6e2ac5358d6e
capitalized type variables; added yield as keyword
 haftmann parents: 
37224diff
changeset | 481 | "true", "type", "val", "var", "while", "with", "yield" | 
| 34294 | 482 | ] | 
| 483 | #> fold (Code_Target.add_reserved target) [ | |
| 48073 
1b609a7837ef
prefer sys.error over plain error in Scala to avoid deprecation warning
 haftmann parents: 
48072diff
changeset | 484 | "apply", "sys", "scala", "BigInt", "Nil", "List" | 
| 59323 | 485 | ]); | 
| 34294 | 486 | |
| 487 | end; (*struct*) |