src/Tools/Code/code_scala.ML
author wenzelm
Mon, 07 Mar 2016 21:09:28 +0100
changeset 62549 9498623b27f0
parent 61130 8e736ce4c6f4
child 63303 7cffe366d333
permissions -rw-r--r--
File.bash_string operations in ML as in Scala -- exclusively for GNU bash, not perl and not user output; clarified treatment of whitespace in some bash scripts;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37745
6315b6426200 checking generated code for various target languages
haftmann
parents: 37669
diff changeset
     1
(*  Title:      Tools/Code/code_scala.ML
6315b6426200 checking generated code for various target languages
haftmann
parents: 37669
diff changeset
     2
    Author:     Florian Haftmann, TU Muenchen
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
     3
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
     4
Serializer for Scala.
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
     5
*)
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
     6
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
     7
signature CODE_SCALA =
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
     8
sig
37745
6315b6426200 checking generated code for various target languages
haftmann
parents: 37669
diff changeset
     9
  val target: string
58520
a4d1f8041af0 accomplish potentially case-insenstive file systems for Scala
haftmann
parents: 58398
diff changeset
    10
  val case_insensitive: bool Config.T;
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
    11
end;
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
    12
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
    13
structure Code_Scala : CODE_SCALA =
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
    14
struct
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
    15
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
    16
val target = "Scala";
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
    17
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55147
diff changeset
    18
open Basic_Code_Symbol;
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
    19
open Basic_Code_Thingol;
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
    20
open Code_Printer;
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
    21
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
    22
infixr 5 @@;
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
    23
infixr 5 @|;
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
    24
58520
a4d1f8041af0 accomplish potentially case-insenstive file systems for Scala
haftmann
parents: 58398
diff changeset
    25
(** preliminary: option to circumvent clashes on case-insensitive file systems **)
a4d1f8041af0 accomplish potentially case-insenstive file systems for Scala
haftmann
parents: 58398
diff changeset
    26
a4d1f8041af0 accomplish potentially case-insenstive file systems for Scala
haftmann
parents: 58398
diff changeset
    27
val case_insensitive = Attrib.setup_config_bool @{binding "scala_case_insensitive"} (K false);
a4d1f8041af0 accomplish potentially case-insenstive file systems for Scala
haftmann
parents: 58398
diff changeset
    28
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
    29
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
    30
(** Scala serializer **)
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
    31
47609
b3dab1892cda dropped dead code
haftmann
parents: 46850
diff changeset
    32
fun print_scala_stmt tyco_syntax const_syntax reserved
50626
e21485358c56 uniform parentheses for constructor -- necessary to accomodate scala 10
haftmann
parents: 48568
diff changeset
    33
    args_num is_constr (deresolve, deresolve_full) =
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
    34
  let
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55147
diff changeset
    35
    val deresolve_const = deresolve o Constant;
0940309ed8f1 less clumsy namespace
haftmann
parents: 55147
diff changeset
    36
    val deresolve_tyco = deresolve o Type_Constructor;
0940309ed8f1 less clumsy namespace
haftmann
parents: 55147
diff changeset
    37
    val deresolve_class = deresolve o Type_Class;
56812
baef1c110f12 centralized upper/lowercase name mangling
haftmann
parents: 55776
diff changeset
    38
    fun lookup_tyvar tyvars = lookup_var tyvars o Name.enforce_case true;
baef1c110f12 centralized upper/lowercase name mangling
haftmann
parents: 55776
diff changeset
    39
    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: 55145
diff changeset
    40
    fun print_tyco_expr tyvars fxy (sym, tys) = applify "[" "]"
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
    41
          (print_typ tyvars NOBR) fxy ((str o deresolve) sym) tys
38923
79d7f2b4cf71 more coherent naming of syntax data structures
haftmann
parents: 38922
diff changeset
    42
    and print_typ tyvars fxy (tyco `%% tys) = (case tyco_syntax tyco
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55147
diff changeset
    43
         of NONE => print_tyco_expr tyvars fxy (Type_Constructor tyco, tys)
47609
b3dab1892cda dropped dead code
haftmann
parents: 46850
diff changeset
    44
          | SOME (_, print) => print (print_typ tyvars) fxy tys)
37243
6e2ac5358d6e capitalized type variables; added yield as keyword
haftmann
parents: 37224
diff changeset
    45
      | print_typ tyvars fxy (ITyVar v) = (str o lookup_tyvar tyvars) v;
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55147
diff changeset
    46
    fun print_dicttyp tyvars (class, ty) = print_tyco_expr tyvars NOBR (Type_Class class, [ty]);
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
    47
    fun print_tupled_typ tyvars ([], ty) =
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
    48
          print_typ tyvars NOBR ty
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
    49
      | print_tupled_typ tyvars ([ty1], ty2) =
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
    50
          concat [print_typ tyvars BR ty1, str "=>", print_typ tyvars NOBR ty2]
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
    51
      | print_tupled_typ tyvars (tys, ty) =
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
    52
          concat [enum "," "(" ")" (map (print_typ tyvars NOBR) tys),
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
    53
            str "=>", print_typ tyvars NOBR ty];
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
    54
    fun constraint p1 p2 = Pretty.block [p1, str ":", Pretty.brk 1, p2];
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
    55
    fun print_var vars NONE = str "_"
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
    56
      | print_var vars (SOME v) = (str o lookup_var vars) v
48072
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
    57
    fun print_term tyvars is_pat some_thm vars fxy (IConst const) =
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
    58
          print_app tyvars is_pat some_thm vars fxy (const, [])
35228
ac2cab4583f4 context theorem is optional
haftmann
parents: 34944
diff changeset
    59
      | print_term tyvars is_pat some_thm vars fxy (t as (t1 `$ t2)) =
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
    60
          (case Code_Thingol.unfold_const_app t
35228
ac2cab4583f4 context theorem is optional
haftmann
parents: 34944
diff changeset
    61
           of SOME app => print_app tyvars is_pat some_thm vars fxy app
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
    62
            | _ => applify "(" ")" (print_term tyvars is_pat some_thm vars NOBR) fxy
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
    63
                (print_term tyvars is_pat some_thm vars BR t1) [t2])
35228
ac2cab4583f4 context theorem is optional
haftmann
parents: 34944
diff changeset
    64
      | print_term tyvars is_pat some_thm vars fxy (IVar v) =
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
    65
          print_var vars v
61130
8e736ce4c6f4 unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents: 59323
diff changeset
    66
      | print_term tyvars is_pat some_thm vars fxy (t as _ `|=> _) =
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
    67
          let
61130
8e736ce4c6f4 unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents: 59323
diff changeset
    68
            val (vs_tys, body) = Code_Thingol.unfold_abs t;
8e736ce4c6f4 unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents: 59323
diff changeset
    69
            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: 59323
diff changeset
    70
            val vars' = intro_vars (map_filter fst vs_tys) vars;
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
    71
          in
61130
8e736ce4c6f4 unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents: 59323
diff changeset
    72
            brackets (ps @| print_term tyvars false some_thm vars' NOBR body)
41939
wenzelm
parents: 41784
diff changeset
    73
          end
48072
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
    74
      | print_term tyvars is_pat some_thm vars fxy (ICase case_expr) =
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
    75
          (case Code_Thingol.unfold_const_app (#primitive case_expr)
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55147
diff changeset
    76
           of SOME (app as ({ sym = Constant const, ... }, _)) =>
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
    77
                if is_none (const_syntax const)
48072
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
    78
                then print_case tyvars some_thm vars fxy case_expr
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
    79
                else print_app tyvars is_pat some_thm vars fxy app
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
    80
            | 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: 59323
diff changeset
    81
    and print_abs_head tyvars (some_v, ty) vars =
8e736ce4c6f4 unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents: 59323
diff changeset
    82
           let
8e736ce4c6f4 unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents: 59323
diff changeset
    83
             val vars' = intro_vars (the_list some_v) vars;
8e736ce4c6f4 unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents: 59323
diff changeset
    84
           in
8e736ce4c6f4 unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents: 59323
diff changeset
    85
             (concat [
8e736ce4c6f4 unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents: 59323
diff changeset
    86
               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: 59323
diff changeset
    87
               str "=>"
8e736ce4c6f4 unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents: 59323
diff changeset
    88
             ], vars')
8e736ce4c6f4 unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents: 59323
diff changeset
    89
           end
37464
9250ad1b98e0 tuned whitespace; dropped dead code
haftmann
parents: 37453
diff changeset
    90
    and print_app tyvars is_pat some_thm vars fxy
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
    91
        (app as ({ sym, typargs, dom, ... }, ts)) =
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
    92
      let
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
    93
        val k = length ts;
50626
e21485358c56 uniform parentheses for constructor -- necessary to accomodate scala 10
haftmann
parents: 48568
diff changeset
    94
        val typargs' = if is_pat then [] else typargs;
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
    95
        val syntax = case sym of
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55147
diff changeset
    96
            Constant const => const_syntax const
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
    97
          | _ => NONE;
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
    98
        val (l, print') = case syntax of
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
    99
            NONE => (args_num sym, fn fxy => fn ts => gen_applify (is_constr sym) "(" ")"
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   100
              (print_term tyvars is_pat some_thm vars NOBR) fxy
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   101
                (applify "[" "]" (print_typ tyvars NOBR)
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   102
                  NOBR ((str o deresolve) sym) typargs') ts)
55153
eedd549de3ef more suitable names, without any notion of "activating"
haftmann
parents: 55150
diff changeset
   103
          | SOME (k, Plain_printer s) => (k, fn fxy => fn ts => applify "(" ")"
37881
096c8397c989 distinguish different classes of const syntax
haftmann
parents: 37822
diff changeset
   104
              (print_term tyvars is_pat some_thm vars NOBR) fxy
096c8397c989 distinguish different classes of const syntax
haftmann
parents: 37822
diff changeset
   105
                (applify "[" "]" (print_typ tyvars NOBR)
48072
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
   106
                  NOBR (str s) typargs') ts)
55153
eedd549de3ef more suitable names, without any notion of "activating"
haftmann
parents: 55150
diff changeset
   107
          | SOME (k, Complex_printer print) =>
38059
72f4630d4c43 tuned printing of applications and let cascades
haftmann
parents: 37958
diff changeset
   108
              (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: 48003
diff changeset
   109
                (ts ~~ take k dom))
38059
72f4630d4c43 tuned printing of applications and let cascades
haftmann
parents: 37958
diff changeset
   110
      in if k = l then print' fxy ts
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   111
      else if k < l then
35228
ac2cab4583f4 context theorem is optional
haftmann
parents: 34944
diff changeset
   112
        print_term tyvars is_pat some_thm vars fxy (Code_Thingol.eta_expand l app)
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   113
      else let
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   114
        val (ts1, ts23) = chop l ts;
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   115
      in
38059
72f4630d4c43 tuned printing of applications and let cascades
haftmann
parents: 37958
diff changeset
   116
        Pretty.block (print' BR ts1 :: map (fn t => Pretty.block
35228
ac2cab4583f4 context theorem is optional
haftmann
parents: 34944
diff changeset
   117
          [str ".apply(", print_term tyvars is_pat some_thm vars NOBR t, str ")"]) ts23)
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   118
      end end
37464
9250ad1b98e0 tuned whitespace; dropped dead code
haftmann
parents: 37453
diff changeset
   119
    and print_bind tyvars some_thm fxy p =
9250ad1b98e0 tuned whitespace; dropped dead code
haftmann
parents: 37453
diff changeset
   120
      gen_print_bind (print_term tyvars true) some_thm fxy p
48072
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
   121
    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: 48072
diff changeset
   122
          (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: 48003
diff changeset
   123
      | print_case tyvars some_thm vars fxy (case_expr as { clauses = [_], ... }) =
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   124
          let
48072
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
   125
            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: 41687
diff changeset
   126
            fun print_match_val ((pat, ty), t) vars =
32a7726d2136 more idiomatic printing of let cascades and type variable constraints
haftmann
parents: 41687
diff changeset
   127
              vars
32a7726d2136 more idiomatic printing of let cascades and type variable constraints
haftmann
parents: 41687
diff changeset
   128
              |> print_bind tyvars some_thm BR pat
32a7726d2136 more idiomatic printing of let cascades and type variable constraints
haftmann
parents: 41687
diff changeset
   129
              |>> (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: 41687
diff changeset
   130
                  str "=", print_term tyvars false some_thm vars NOBR t]));
32a7726d2136 more idiomatic printing of let cascades and type variable constraints
haftmann
parents: 41687
diff changeset
   131
            fun print_match_seq t vars =
32a7726d2136 more idiomatic printing of let cascades and type variable constraints
haftmann
parents: 41687
diff changeset
   132
              ((true, print_term tyvars false some_thm vars NOBR t), vars);
32a7726d2136 more idiomatic printing of let cascades and type variable constraints
haftmann
parents: 41687
diff changeset
   133
            fun print_match is_first ((IVar NONE, ty), t) =
41784
537013b8ba8e dropped redundancy
haftmann
parents: 41781
diff changeset
   134
                  if Code_Thingol.is_IAbs t andalso is_first
41781
32a7726d2136 more idiomatic printing of let cascades and type variable constraints
haftmann
parents: 41687
diff changeset
   135
                    then print_match_val ((IVar NONE, ty), t)
32a7726d2136 more idiomatic printing of let cascades and type variable constraints
haftmann
parents: 41687
diff changeset
   136
                    else print_match_seq t
32a7726d2136 more idiomatic printing of let cascades and type variable constraints
haftmann
parents: 41687
diff changeset
   137
              | print_match _ ((pat, ty), t) =
32a7726d2136 more idiomatic printing of let cascades and type variable constraints
haftmann
parents: 41687
diff changeset
   138
                  print_match_val ((pat, ty), t);
32a7726d2136 more idiomatic printing of let cascades and type variable constraints
haftmann
parents: 41687
diff changeset
   139
            val (seps_ps, vars') =
32a7726d2136 more idiomatic printing of let cascades and type variable constraints
haftmann
parents: 41687
diff changeset
   140
              vars |> print_match true bind ||>> fold_map (print_match false) binds |>> uncurry cons;
38059
72f4630d4c43 tuned printing of applications and let cascades
haftmann
parents: 37958
diff changeset
   141
            val all_seps_ps = seps_ps @ [(true, print_term tyvars false some_thm vars' NOBR body)];
72f4630d4c43 tuned printing of applications and let cascades
haftmann
parents: 37958
diff changeset
   142
            fun insert_seps [(_, p)] = [p]
72f4630d4c43 tuned printing of applications and let cascades
haftmann
parents: 37958
diff changeset
   143
              | insert_seps ((_, p) :: (seps_ps as (sep, _) :: _)) =
72f4630d4c43 tuned printing of applications and let cascades
haftmann
parents: 37958
diff changeset
   144
                  (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: 41687
diff changeset
   145
          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: 48003
diff changeset
   146
      | print_case tyvars some_thm vars fxy { term = t, typ = ty, clauses = clauses as _ :: _, ... } =
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   147
          let
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   148
            fun print_select (pat, body) =
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   149
              let
37464
9250ad1b98e0 tuned whitespace; dropped dead code
haftmann
parents: 37453
diff changeset
   150
                val (p_pat, vars') = print_bind tyvars some_thm NOBR pat vars;
9250ad1b98e0 tuned whitespace; dropped dead code
haftmann
parents: 37453
diff changeset
   151
                val p_body = print_term tyvars false some_thm vars' NOBR body
9250ad1b98e0 tuned whitespace; dropped dead code
haftmann
parents: 37453
diff changeset
   152
              in concat [str "case", p_pat, str "=>", p_body] end;
46850
7b04cfc24eb6 always bracket case expressions in Scala
haftmann
parents: 45009
diff changeset
   153
          in
7b04cfc24eb6 always bracket case expressions in Scala
haftmann
parents: 45009
diff changeset
   154
            map print_select clauses
7b04cfc24eb6 always bracket case expressions in Scala
haftmann
parents: 45009
diff changeset
   155
            |> Pretty.block_enclose (concat [print_term tyvars false some_thm vars NOBR t, str "match", str "{"], str "}")
7b04cfc24eb6 always bracket case expressions in Scala
haftmann
parents: 45009
diff changeset
   156
            |> single
7b04cfc24eb6 always bracket case expressions in Scala
haftmann
parents: 45009
diff changeset
   157
            |> enclose "(" ")"
48072
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
   158
          end;
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   159
    fun print_context tyvars vs sym = applify "[" "]"
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   160
      (fn (v, sort) => (Pretty.block o map str)
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   161
        (lookup_tyvar tyvars v :: maps (fn class => [" : ", deresolve_class class]) sort))
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   162
          NOBR ((str o deresolve) sym) vs;
55679
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55153
diff changeset
   163
    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: 55684
diff changeset
   164
      concat [str "def", constraint (applify "(" ")" (fn (param, ty) =>
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   165
        constraint ((str o lookup_var vars) param) (print_typ tyvars NOBR ty))
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55147
diff changeset
   166
          NOBR (print_context tyvars vs (Constant const)) (params ~~ tys)) (print_typ tyvars NOBR ty),
55679
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55153
diff changeset
   167
            str "="];
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55153
diff changeset
   168
    fun print_def export const (vs, ty) [] =
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   169
          let
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   170
            val (tys, ty') = Code_Thingol.unfold_fun ty;
43329
84472e198515 tuned signature: Name.invent and Name.invent_names;
wenzelm
parents: 43326
diff changeset
   171
            val params = Name.invent (snd reserved) "a" (length tys);
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   172
            val tyvars = intro_tyvars vs reserved;
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   173
            val vars = intro_vars params reserved;
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   174
          in
55679
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55153
diff changeset
   175
            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: 55145
diff changeset
   176
              str ("sys.error(\"" ^ const ^ "\")")]
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   177
          end
55679
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55153
diff changeset
   178
      | print_def export const (vs, ty) eqs =
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   179
          let
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   180
            val tycos = fold (fn ((ts, t), _) =>
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   181
              fold Code_Thingol.add_tyconames (t :: ts)) eqs [];
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   182
            val tyvars = reserved
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   183
              |> intro_base_names
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   184
                   (is_none o tyco_syntax) deresolve_tyco tycos
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   185
              |> intro_tyvars vs;
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   186
            val simple = case eqs
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   187
             of [((ts, _), _)] => forall Code_Thingol.is_IVar ts
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   188
              | _ => false;
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   189
            val vars1 = reserved
55145
2bb3cd36bcf7 more abstract declaration of unqualified constant names in code printing context
haftmann
parents: 52520
diff changeset
   190
              |> intro_base_names_for (is_none o const_syntax)
2bb3cd36bcf7 more abstract declaration of unqualified constant names in code printing context
haftmann
parents: 52520
diff changeset
   191
                   deresolve (map (snd o fst) eqs);
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   192
            val params = if simple
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   193
              then (map (fn IVar (SOME x) => x) o fst o fst o hd) eqs
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   194
              else aux_params vars1 (map (fst o fst) eqs);
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   195
            val vars2 = intro_vars params vars1;
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   196
            val (tys', ty') = Code_Thingol.unfold_fun_n (length params) ty;
38922
ec2a8efd8990 Code_Printer.tuplify
haftmann
parents: 38916
diff changeset
   197
            fun tuplify [p] = p
ec2a8efd8990 Code_Printer.tuplify
haftmann
parents: 38916
diff changeset
   198
              | tuplify ps = enum "," "(" ")" ps;
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   199
            fun print_rhs vars' ((_, t), (some_thm, _)) =
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   200
              print_term tyvars false some_thm vars' NOBR t;
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   201
            fun print_clause (eq as ((ts, _), (some_thm, _))) =
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   202
              let
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   203
                val vars' = intro_vars ((fold o Code_Thingol.fold_varnames)
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   204
                  (insert (op =)) ts []) vars1;
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   205
              in
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   206
                concat [str "case",
38922
ec2a8efd8990 Code_Printer.tuplify
haftmann
parents: 38916
diff changeset
   207
                  tuplify (map (print_term tyvars true some_thm vars' NOBR) ts),
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   208
                  str "=>", print_rhs vars' eq]
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   209
              end;
55679
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55153
diff changeset
   210
            val head = print_defhead export tyvars vars2 const vs params tys' ty';
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   211
          in if simple then
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   212
            concat [head, print_rhs vars2 (hd eqs)]
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   213
          else
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   214
            Pretty.block_enclose
38922
ec2a8efd8990 Code_Printer.tuplify
haftmann
parents: 38916
diff changeset
   215
              (concat [head, tuplify (map (str o lookup_var vars2) params),
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   216
                str "match", str "{"], str "}")
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   217
              (map print_clause eqs)
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   218
          end;
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55147
diff changeset
   219
    val print_method = str o Library.enclose "`" "`" o deresolve_full o Constant;
55679
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55153
diff changeset
   220
    fun print_stmt (Constant const, (export, Code_Thingol.Fun (((vs, ty), raw_eqs), _))) =
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55153
diff changeset
   221
          print_def export const (vs, ty) (filter (snd o snd) raw_eqs)
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55153
diff changeset
   222
      | print_stmt (Type_Constructor tyco, (export, Code_Thingol.Datatype (vs, cos))) =
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   223
          let
48003
1d11af40b106 dropped sort constraints on datatype specifications
haftmann
parents: 47609
diff changeset
   224
            val tyvars = intro_tyvars (map (rpair []) vs) reserved;
50626
e21485358c56 uniform parentheses for constructor -- necessary to accomodate scala 10
haftmann
parents: 48568
diff changeset
   225
            fun print_co ((co, vs_args), tys) =
e21485358c56 uniform parentheses for constructor -- necessary to accomodate scala 10
haftmann
parents: 48568
diff changeset
   226
              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: 55684
diff changeset
   227
                ((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: 48568
diff changeset
   228
                @@ 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: 48568
diff changeset
   229
                  (Name.invent_names (snd reserved) "a" tys))),
e21485358c56 uniform parentheses for constructor -- necessary to accomodate scala 10
haftmann
parents: 48568
diff changeset
   230
                str "extends",
e21485358c56 uniform parentheses for constructor -- necessary to accomodate scala 10
haftmann
parents: 48568
diff changeset
   231
                applify "[" "]" (str o lookup_tyvar tyvars) NOBR
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   232
                  ((str o deresolve_tyco) tyco) vs
50626
e21485358c56 uniform parentheses for constructor -- necessary to accomodate scala 10
haftmann
parents: 48568
diff changeset
   233
              ];
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   234
          in
50626
e21485358c56 uniform parentheses for constructor -- necessary to accomodate scala 10
haftmann
parents: 48568
diff changeset
   235
            Pretty.chunks (applify "[" "]" (str o lookup_tyvar tyvars)
55776
7dd1971b39c1 amended some slips, rolling back currently dysfunctional export minimimalisation for Scala
haftmann
parents: 55684
diff changeset
   236
              NOBR ((concat o map str) ["abstract", "sealed", "class", deresolve_tyco tyco]) vs
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   237
                :: map print_co cos)
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   238
          end
55679
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55153
diff changeset
   239
      | print_stmt (Type_Class class, (export, Code_Thingol.Class (v, (classrels, classparams)))) =
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   240
          let
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   241
            val tyvars = intro_tyvars [(v, [class])] reserved;
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   242
            fun add_typarg s = Pretty.block
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   243
              [str s, str "[", (str o lookup_tyvar tyvars) v, str "]"];
37384
5aba26803073 more consistent naming aroud type classes and instances
haftmann
parents: 37337
diff changeset
   244
            fun print_super_classes [] = NONE
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   245
              | print_super_classes classrels = SOME (concat (str "extends"
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   246
                  :: separate (str "with") (map (add_typarg o deresolve_class o snd) classrels)));
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   247
            fun print_classparam_val (classparam, ty) =
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   248
              concat [str "val", constraint (print_method classparam)
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   249
                ((print_tupled_typ tyvars o Code_Thingol.unfold_fun) ty)];
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   250
            fun print_classparam_def (classparam, ty) =
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   251
              let
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   252
                val (tys, ty) = Code_Thingol.unfold_fun ty;
43329
84472e198515 tuned signature: Name.invent and Name.invent_names;
wenzelm
parents: 43326
diff changeset
   253
                val [implicit_name] = Name.invent (snd reserved) (lookup_tyvar tyvars v) 1;
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   254
                val proto_vars = intro_vars [implicit_name] reserved;
43329
84472e198515 tuned signature: Name.invent and Name.invent_names;
wenzelm
parents: 43326
diff changeset
   255
                val auxs = Name.invent (snd proto_vars) "a" (length tys);
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   256
                val vars = intro_vars auxs proto_vars;
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   257
              in
55776
7dd1971b39c1 amended some slips, rolling back currently dysfunctional export minimimalisation for Scala
haftmann
parents: 55684
diff changeset
   258
                concat [str "def", constraint (Pretty.block [applify "(" ")"
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   259
                  (fn (aux, ty) => constraint ((str o lookup_var vars) aux)
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   260
                  (print_typ tyvars NOBR ty)) NOBR (add_typarg (deresolve_const classparam))
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   261
                  (auxs ~~ tys), str "(implicit ", str implicit_name, str ": ",
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   262
                  add_typarg (deresolve_class class), str ")"]) (print_typ tyvars NOBR ty), str "=",
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   263
                  applify "(" ")" (str o lookup_var vars) NOBR
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   264
                  (Pretty.block [str implicit_name, str ".", print_method classparam]) auxs]
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   265
              end;
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   266
          in
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   267
            Pretty.chunks (
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   268
              (Pretty.block_enclose
55776
7dd1971b39c1 amended some slips, rolling back currently dysfunctional export minimimalisation for Scala
haftmann
parents: 55684
diff changeset
   269
                (concat ([str "trait", (add_typarg o deresolve_class) class]
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   270
                  @ the_list (print_super_classes classrels) @ [str "{"]), str "}")
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   271
                (map print_classparam_val classparams))
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   272
              :: map print_classparam_def classparams
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   273
            )
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   274
          end
55679
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55153
diff changeset
   275
      | print_stmt (sym, (export, Code_Thingol.Classinst
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55153
diff changeset
   276
          { class, tyco, vs, inst_params, superinst_params, ... })) =
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   277
          let
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   278
            val tyvars = intro_tyvars vs reserved;
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   279
            val classtyp = (class, tyco `%% map (ITyVar o fst) vs);
52520
4a884366b0d8 consider explicit hint for domain of class parameters when printing instance statements
haftmann
parents: 52519
diff changeset
   280
            fun print_classparam_instance ((classparam, (const as { dom, ... }, dom_length)), (thm, _)) =
37450
45073611170a first serious draft of a scala code generator
haftmann
parents: 37447
diff changeset
   281
              let
48072
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
   282
                val aux_dom = Name.invent_names (snd reserved) "a" dom;
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
   283
                val auxs = map fst aux_dom;
37450
45073611170a first serious draft of a scala code generator
haftmann
parents: 37447
diff changeset
   284
                val vars = intro_vars auxs reserved;
52520
4a884366b0d8 consider explicit hint for domain of class parameters when printing instance statements
haftmann
parents: 52519
diff changeset
   285
                val (aux_dom1, aux_dom2) = chop dom_length aux_dom;
4a884366b0d8 consider explicit hint for domain of class parameters when printing instance statements
haftmann
parents: 52519
diff changeset
   286
                fun abstract_using [] = []
4a884366b0d8 consider explicit hint for domain of class parameters when printing instance statements
haftmann
parents: 52519
diff changeset
   287
                  | abstract_using aux_dom = [enum "," "(" ")"
4a884366b0d8 consider explicit hint for domain of class parameters when printing instance statements
haftmann
parents: 52519
diff changeset
   288
                      (map (fn (aux, ty) => constraint ((str o lookup_var vars) aux)
4a884366b0d8 consider explicit hint for domain of class parameters when printing instance statements
haftmann
parents: 52519
diff changeset
   289
                      (print_typ tyvars NOBR ty)) aux_dom), str "=>"];
4a884366b0d8 consider explicit hint for domain of class parameters when printing instance statements
haftmann
parents: 52519
diff changeset
   290
                val aux_abstr1 = abstract_using aux_dom1;
4a884366b0d8 consider explicit hint for domain of class parameters when printing instance statements
haftmann
parents: 52519
diff changeset
   291
                val aux_abstr2 = abstract_using aux_dom2;
41939
wenzelm
parents: 41784
diff changeset
   292
              in
55684
ee49b4f7edc8 keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents: 55683
diff changeset
   293
                concat ([str "val", print_method classparam, str "="]
52520
4a884366b0d8 consider explicit hint for domain of class parameters when printing instance statements
haftmann
parents: 52519
diff changeset
   294
                  @ aux_abstr1 @ aux_abstr2 @| print_app tyvars false (SOME thm) vars NOBR
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   295
                    (const, map (IVar o SOME) auxs))
37450
45073611170a first serious draft of a scala code generator
haftmann
parents: 37447
diff changeset
   296
              end;
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   297
          in
55776
7dd1971b39c1 amended some slips, rolling back currently dysfunctional export minimimalisation for Scala
haftmann
parents: 55684
diff changeset
   298
            Pretty.block_enclose (concat [str "implicit def",
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   299
              constraint (print_context tyvars vs sym) (print_dicttyp tyvars classtyp),
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   300
              str "=", str "new", print_dicttyp tyvars classtyp, str "{"], str "}")
48072
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 48003
diff changeset
   301
                (map print_classparam_instance (inst_params @ superinst_params))
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   302
          end;
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   303
  in print_stmt end;
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   304
55683
5732a55b9232 explicit option for "open" code generation
haftmann
parents: 55681
diff changeset
   305
fun scala_program_of_program ctxt module_name reserved identifiers exports program =
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   306
  let
58520
a4d1f8041af0 accomplish potentially case-insenstive file systems for Scala
haftmann
parents: 58398
diff changeset
   307
    val variant = if Config.get ctxt case_insensitive
a4d1f8041af0 accomplish potentially case-insenstive file systems for Scala
haftmann
parents: 58398
diff changeset
   308
      then Code_Namespace.variant_case_insensitive
a4d1f8041af0 accomplish potentially case-insenstive file systems for Scala
haftmann
parents: 58398
diff changeset
   309
      else Name.variant;
38809
7dc73a208722 proper namespace administration for hierarchical modules
haftmann
parents: 38782
diff changeset
   310
    fun namify_module name_fragment ((nsp_class, nsp_object), nsp_common) =
7dc73a208722 proper namespace administration for hierarchical modules
haftmann
parents: 38782
diff changeset
   311
      let
7dc73a208722 proper namespace administration for hierarchical modules
haftmann
parents: 38782
diff changeset
   312
        val declare = Name.declare name_fragment;
7dc73a208722 proper namespace administration for hierarchical modules
haftmann
parents: 38782
diff changeset
   313
      in (name_fragment, ((declare nsp_class, declare nsp_object), declare nsp_common)) end;
38769
317e64c886d2 preliminary implementation of hierarchical module name space
haftmann
parents: 38059
diff changeset
   314
    fun namify_class base ((nsp_class, nsp_object), nsp_common) =
317e64c886d2 preliminary implementation of hierarchical module name space
haftmann
parents: 38059
diff changeset
   315
      let
58520
a4d1f8041af0 accomplish potentially case-insenstive file systems for Scala
haftmann
parents: 58398
diff changeset
   316
        val (base', nsp_class') = variant base nsp_class
38769
317e64c886d2 preliminary implementation of hierarchical module name space
haftmann
parents: 38059
diff changeset
   317
      in (base', ((nsp_class', nsp_object), Name.declare base' nsp_common)) end;
317e64c886d2 preliminary implementation of hierarchical module name space
haftmann
parents: 38059
diff changeset
   318
    fun namify_object base ((nsp_class, nsp_object), nsp_common) =
317e64c886d2 preliminary implementation of hierarchical module name space
haftmann
parents: 38059
diff changeset
   319
      let
58520
a4d1f8041af0 accomplish potentially case-insenstive file systems for Scala
haftmann
parents: 58398
diff changeset
   320
        val (base', nsp_object') = variant base nsp_object
38769
317e64c886d2 preliminary implementation of hierarchical module name space
haftmann
parents: 38059
diff changeset
   321
      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: 56812
diff changeset
   322
    fun namify_common base ((nsp_class, nsp_object), nsp_common) =
38769
317e64c886d2 preliminary implementation of hierarchical module name space
haftmann
parents: 38059
diff changeset
   323
      let
58520
a4d1f8041af0 accomplish potentially case-insenstive file systems for Scala
haftmann
parents: 58398
diff changeset
   324
        val (base', nsp_common') = variant base nsp_common
38769
317e64c886d2 preliminary implementation of hierarchical module name space
haftmann
parents: 38059
diff changeset
   325
      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: 56812
diff changeset
   326
        (base', ((Name.declare base' nsp_class, Name.declare base' nsp_object), nsp_common'))
38769
317e64c886d2 preliminary implementation of hierarchical module name space
haftmann
parents: 38059
diff changeset
   327
      end;
38809
7dc73a208722 proper namespace administration for hierarchical modules
haftmann
parents: 38782
diff changeset
   328
    fun namify_stmt (Code_Thingol.Fun _) = namify_object
7dc73a208722 proper namespace administration for hierarchical modules
haftmann
parents: 38782
diff changeset
   329
      | 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: 56812
diff changeset
   330
      | namify_stmt (Code_Thingol.Datatypecons _) = namify_common
38809
7dc73a208722 proper namespace administration for hierarchical modules
haftmann
parents: 38782
diff changeset
   331
      | namify_stmt (Code_Thingol.Class _) = namify_class
7dc73a208722 proper namespace administration for hierarchical modules
haftmann
parents: 38782
diff changeset
   332
      | namify_stmt (Code_Thingol.Classrel _) = namify_object
7dc73a208722 proper namespace administration for hierarchical modules
haftmann
parents: 38782
diff changeset
   333
      | 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: 56812
diff changeset
   334
      | namify_stmt (Code_Thingol.Classinst _) = namify_common;
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   335
    fun memorize_implicits sym =
38769
317e64c886d2 preliminary implementation of hierarchical module name space
haftmann
parents: 38059
diff changeset
   336
      let
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents: 38966
diff changeset
   337
        fun is_classinst stmt = case stmt
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents: 38966
diff changeset
   338
         of Code_Thingol.Classinst _ => true
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents: 38966
diff changeset
   339
          | _ => false;
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   340
        val implicits = filter (is_classinst o Code_Symbol.Graph.get_node program)
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   341
          (Code_Symbol.Graph.immediate_succs program sym);
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents: 38966
diff changeset
   342
      in union (op =) implicits end;
55684
ee49b4f7edc8 keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents: 55683
diff changeset
   343
    fun modify_stmt (_, (_, Code_Thingol.Fun (_, SOME _))) = NONE
ee49b4f7edc8 keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents: 55683
diff changeset
   344
      | modify_stmt (_, (_, Code_Thingol.Datatypecons _)) = NONE
ee49b4f7edc8 keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents: 55683
diff changeset
   345
      | modify_stmt (_, (_, Code_Thingol.Classrel _)) = NONE
ee49b4f7edc8 keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents: 55683
diff changeset
   346
      | modify_stmt (_, (_, Code_Thingol.Classparam _)) = NONE
ee49b4f7edc8 keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents: 55683
diff changeset
   347
      | modify_stmt (_, export_stmt) = SOME export_stmt;
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents: 38966
diff changeset
   348
  in
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   349
    Code_Namespace.hierarchical_program ctxt
52138
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51143
diff changeset
   350
      { module_name = module_name, reserved = reserved, identifiers = identifiers,
41939
wenzelm
parents: 41784
diff changeset
   351
        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: 55683
diff changeset
   352
        namify_stmt = namify_stmt, cyclic_modules = true,
ee49b4f7edc8 keep only identifiers public which are explicitly requested or demanded by dependencies
haftmann
parents: 55683
diff changeset
   353
        class_transitive = true, class_relation_public = false, empty_data = [],
55683
5732a55b9232 explicit option for "open" code generation
haftmann
parents: 55681
diff changeset
   354
        memorize_data = memorize_implicits, modify_stmts = map modify_stmt } exports program
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents: 38966
diff changeset
   355
  end;
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   356
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   357
fun serialize_scala ctxt { module_name, reserved_syms, identifiers,
55683
5732a55b9232 explicit option for "open" code generation
haftmann
parents: 55681
diff changeset
   358
    includes, class_syntax, tyco_syntax, const_syntax } exports program =
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   359
  let
38769
317e64c886d2 preliminary implementation of hierarchical module name space
haftmann
parents: 38059
diff changeset
   360
38809
7dc73a208722 proper namespace administration for hierarchical modules
haftmann
parents: 38782
diff changeset
   361
    (* build program *)
39147
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39102
diff changeset
   362
    val { deresolver, hierarchical_program = scala_program } =
55681
7714287dc044 more fine-grain notion of export
haftmann
parents: 55679
diff changeset
   363
      scala_program_of_program ctxt module_name (Name.make_context reserved_syms)
55683
5732a55b9232 explicit option for "open" code generation
haftmann
parents: 55681
diff changeset
   364
        identifiers exports program;
38769
317e64c886d2 preliminary implementation of hierarchical module name space
haftmann
parents: 38059
diff changeset
   365
317e64c886d2 preliminary implementation of hierarchical module name space
haftmann
parents: 38059
diff changeset
   366
    (* print statements *)
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55147
diff changeset
   367
    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: 55145
diff changeset
   368
     of Code_Thingol.Datatype (_, constrs) =>
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   369
          the (AList.lookup (op = o apsnd fst) constrs constr);
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55147
diff changeset
   370
    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: 55145
diff changeset
   371
     of Code_Thingol.Class (_, (_, classparams)) => classparams;
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55147
diff changeset
   372
    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: 55145
diff changeset
   373
     of Code_Thingol.Fun (((_, ty), []), _) =>
37464
9250ad1b98e0 tuned whitespace; dropped dead code
haftmann
parents: 37453
diff changeset
   374
          (length o fst o Code_Thingol.unfold_fun) ty
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   375
      | Code_Thingol.Fun ((_, ((ts, _), _) :: _), _) => length ts
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   376
      | Code_Thingol.Datatypecons tyco => length (lookup_constr tyco const)
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   377
      | Code_Thingol.Classparam class =>
37639
5b6733e6e033 pervasive tuning of code
haftmann
parents: 37464
diff changeset
   378
          (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: 55145
diff changeset
   379
            (classparams_of_class class)) const;
47609
b3dab1892cda dropped dead code
haftmann
parents: 46850
diff changeset
   380
    fun print_stmt prefix_fragments = print_scala_stmt
39147
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39102
diff changeset
   381
      tyco_syntax const_syntax (make_vars reserved_syms) args_num
50626
e21485358c56 uniform parentheses for constructor -- necessary to accomodate scala 10
haftmann
parents: 48568
diff changeset
   382
      (Code_Thingol.is_constr program) (deresolver prefix_fragments, deresolver []);
38769
317e64c886d2 preliminary implementation of hierarchical module name space
haftmann
parents: 38059
diff changeset
   383
39147
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39102
diff changeset
   384
    (* print modules *)
38809
7dc73a208722 proper namespace administration for hierarchical modules
haftmann
parents: 38782
diff changeset
   385
    fun print_implicit prefix_fragments implicit =
38782
3865cbe5d2be only print qualified implicits
haftmann
parents: 38780
diff changeset
   386
      let
38809
7dc73a208722 proper namespace administration for hierarchical modules
haftmann
parents: 38782
diff changeset
   387
        val s = deresolver prefix_fragments implicit;
38782
3865cbe5d2be only print qualified implicits
haftmann
parents: 38780
diff changeset
   388
      in if length (Long_Name.explode s) = 1 then NONE else SOME (str s) end;
39147
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39102
diff changeset
   389
    fun print_module prefix_fragments base implicits ps = Pretty.chunks2
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39102
diff changeset
   390
      ([str ("object " ^ base ^ " {")]
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39102
diff changeset
   391
        @ (case map_filter (print_implicit prefix_fragments) implicits
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39102
diff changeset
   392
            of [] => [] | implicit_ps => (single o Pretty.block)
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39102
diff changeset
   393
            (str "import /*implicits*/" :: Pretty.brk 1 :: commas implicit_ps))
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39102
diff changeset
   394
        @ ps @ [str ("} /* object " ^ base ^ " */")]);
38769
317e64c886d2 preliminary implementation of hierarchical module name space
haftmann
parents: 38059
diff changeset
   395
317e64c886d2 preliminary implementation of hierarchical module name space
haftmann
parents: 38059
diff changeset
   396
    (* serialization *)
39147
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39102
diff changeset
   397
    val p = Pretty.chunks2 (map snd includes
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39102
diff changeset
   398
      @ Code_Namespace.print_hierarchical {
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39102
diff changeset
   399
        print_module = print_module, print_stmt = print_stmt,
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39102
diff changeset
   400
        lift_markup = I } scala_program);
39056
fa197571676b formal markup of generated code for statements
haftmann
parents: 39034
diff changeset
   401
    fun write width NONE = writeln o format [] width
fa197571676b formal markup of generated code for statements
haftmann
parents: 39034
diff changeset
   402
      | write width (SOME p) = File.write p o format [] width;
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55145
diff changeset
   403
    fun prepare syms width p = ([("", format syms width p)], try (deresolver []));
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   404
  in
48568
084cd758a8ab evaluation: allow multiple code modules
haftmann
parents: 48073
diff changeset
   405
    Code_Target.serialization write prepare p
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   406
  end;
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   407
38966
68853347ba37 tuned internally and made smlnj happy
haftmann
parents: 38928
diff changeset
   408
val serializer : Code_Target.serializer =
68853347ba37 tuned internally and made smlnj happy
haftmann
parents: 38928
diff changeset
   409
  Code_Target.parse_args (Scan.succeed ()) #> K serialize_scala;
68853347ba37 tuned internally and made smlnj happy
haftmann
parents: 38928
diff changeset
   410
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   411
val literals = let
37224
f4d3c929c526 corrected printing of characters
haftmann
parents: 36535
diff changeset
   412
  fun char_scala c = if c = "'" then "\\'"
f4d3c929c526 corrected printing of characters
haftmann
parents: 36535
diff changeset
   413
    else if c = "\"" then "\\\""
f4d3c929c526 corrected printing of characters
haftmann
parents: 36535
diff changeset
   414
    else if c = "\\" then "\\\\"
f4d3c929c526 corrected printing of characters
haftmann
parents: 36535
diff changeset
   415
    else let val k = ord c
f4d3c929c526 corrected printing of characters
haftmann
parents: 36535
diff changeset
   416
    in if k < 32 orelse k > 126 then "\\" ^ radixstring (8, "0", k) else c end
58398
f38717f175d9 simplified and tuned using signed_string_of_int
haftmann
parents: 56826
diff changeset
   417
  fun numeral_scala k =
f38717f175d9 simplified and tuned using signed_string_of_int
haftmann
parents: 56826
diff changeset
   418
    if ~2147483647 < k andalso k <= 2147483647
f38717f175d9 simplified and tuned using signed_string_of_int
haftmann
parents: 56826
diff changeset
   419
    then signed_string_of_int k
f38717f175d9 simplified and tuned using signed_string_of_int
haftmann
parents: 56826
diff changeset
   420
    else quote (signed_string_of_int k)
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   421
in Literals {
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   422
  literal_char = Library.enclose "'" "'" o char_scala,
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   423
  literal_string = quote o translate_string char_scala,
34944
970e1466028d code literals: distinguish numeral classes by different entries
haftmann
parents: 34900
diff changeset
   424
  literal_numeral = fn k => "BigInt(" ^ numeral_scala k ^ ")",
34888
460ec1a99aa2 being more accurate wrt. list syntax
haftmann
parents: 34308
diff changeset
   425
  literal_list = fn [] => str "Nil" | ps => Pretty.block [str "List", enum "," "(" ")" ps],
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   426
  infix_cons = (6, "::")
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   427
} end;
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   428
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   429
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   430
(** Isar setup **)
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   431
59323
468bd3aedfa1 modernized and more uniform style
haftmann
parents: 59104
diff changeset
   432
val _ = Theory.setup
468bd3aedfa1 modernized and more uniform style
haftmann
parents: 59104
diff changeset
   433
  (Code_Target.add_language
38966
68853347ba37 tuned internally and made smlnj happy
haftmann
parents: 38928
diff changeset
   434
    (target, { serializer = serializer, literals = literals,
41939
wenzelm
parents: 41784
diff changeset
   435
      check = { env_var = "SCALA_HOME",
wenzelm
parents: 41784
diff changeset
   436
        make_destination = fn p => Path.append p (Path.explode "ROOT.scala"),
41940
a3b68a7a0e15 allow spaces in executable names;
wenzelm
parents: 41939
diff changeset
   437
        make_command = fn _ =>
a3b68a7a0e15 allow spaces in executable names;
wenzelm
parents: 41939
diff changeset
   438
          "env JAVA_OPTS='-Xms128m -Xmx512m -Xss2m' \"$SCALA_HOME/bin/scalac\" ROOT.scala" } })
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55147
diff changeset
   439
  #> 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: 52229
diff changeset
   440
    [(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: 52229
diff changeset
   441
      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: 52229
diff changeset
   442
        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: 52229
diff changeset
   443
        str "=>",
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52229
diff changeset
   444
        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: 52229
diff changeset
   445
      )))]))
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   446
  #> fold (Code_Target.add_reserved target) [
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   447
      "abstract", "case", "catch", "class", "def", "do", "else", "extends", "false",
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   448
      "final", "finally", "for", "forSome", "if", "implicit", "import", "lazy",
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   449
      "match", "new", "null", "object", "override", "package", "private", "protected",
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   450
      "requires", "return", "sealed", "super", "this", "throw", "trait", "try",
37243
6e2ac5358d6e capitalized type variables; added yield as keyword
haftmann
parents: 37224
diff changeset
   451
      "true", "type", "val", "var", "while", "with", "yield"
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   452
    ]
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   453
  #> fold (Code_Target.add_reserved target) [
48073
1b609a7837ef prefer sys.error over plain error in Scala to avoid deprecation warning
haftmann
parents: 48072
diff changeset
   454
      "apply", "sys", "scala", "BigInt", "Nil", "List"
59323
468bd3aedfa1 modernized and more uniform style
haftmann
parents: 59104
diff changeset
   455
    ]);
34294
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   456
19c1fd52d6c9 a primitive scala serializer
haftmann
parents:
diff changeset
   457
end; (*struct*)