src/HOL/Library/code_test.ML
author wenzelm
Sat, 17 Dec 2016 13:42:25 +0100
changeset 64578 7b20f9f94f4e
parent 64577 0288a566c966
child 64579 38a1d8b41189
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
59720
f893472fff31 proper headers;
wenzelm
parents: 59323
diff changeset
     1
(*  Title:      HOL/Library/code_test.ML
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
     2
    Author:     Andreas Lochbihler, ETH Zürich
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
     3
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
     4
Test infrastructure for the code generator.
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
     5
*)
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
     6
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
     7
signature CODE_TEST =
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
     8
sig
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
     9
  val add_driver:
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    10
    string * ((Proof.context -> (string * string) list * string -> Path.T -> string) * string) ->
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    11
    theory -> theory
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    12
  val get_driver: theory -> string ->
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    13
    ((Proof.context -> (string * string) list * string -> Path.T -> string) * string) option
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    14
  val overlord: bool Config.T
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    15
  val successN: string
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    16
  val failureN: string
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    17
  val start_markerN: string
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    18
  val end_markerN: string
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    19
  val test_terms: Proof.context -> term list -> string -> unit
64578
wenzelm
parents: 64577
diff changeset
    20
  val test_targets: Proof.context -> term list -> string list -> unit
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    21
  val test_code_cmd: string list -> string list -> Toplevel.state -> unit
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    22
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    23
  val eval_term: string -> Proof.context -> term -> term
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    24
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    25
  val gen_driver:
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    26
   (theory -> Path.T -> string list -> string ->
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    27
    {files: (Path.T * string) list,
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    28
     compile_cmd: string option,
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    29
     run_cmd: string,
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    30
     mk_code_file: string -> Path.T}) ->
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    31
   string -> string -> string -> theory -> (string * string) list * string -> Path.T -> string
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    32
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    33
  val ISABELLE_POLYML: string
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    34
  val polymlN: string
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    35
  val evaluate_in_polyml: Proof.context -> (string * string) list * string -> Path.T -> string
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    36
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    37
  val mltonN: string
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    38
  val ISABELLE_MLTON: string
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    39
  val evaluate_in_mlton: Proof.context -> (string * string) list * string -> Path.T -> string
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    40
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    41
  val smlnjN: string
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    42
  val ISABELLE_SMLNJ: string
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    43
  val evaluate_in_smlnj: Proof.context -> (string * string) list * string -> Path.T -> string
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    44
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    45
  val ocamlN: string
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    46
  val ISABELLE_OCAMLC: string
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    47
  val evaluate_in_ocaml: Proof.context -> (string * string) list * string -> Path.T -> string
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    48
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    49
  val ghcN: string
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    50
  val ISABELLE_GHC: string
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    51
  val ghc_options: string Config.T
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    52
  val evaluate_in_ghc: Proof.context -> (string * string) list * string -> Path.T -> string
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    53
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    54
  val scalaN: string
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    55
  val ISABELLE_SCALA: string
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    56
  val evaluate_in_scala: Proof.context -> (string * string) list * string -> Path.T -> string
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    57
end
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    58
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    59
structure Code_Test: CODE_TEST =
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    60
struct
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    61
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    62
(* convert a list of terms into nested tuples and back *)
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    63
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    64
fun mk_tuples [] = @{term "()"}
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    65
  | mk_tuples [t] = t
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    66
  | mk_tuples (t :: ts) = HOLogic.mk_prod (t, mk_tuples ts)
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    67
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    68
fun dest_tuples (Const (@{const_name Pair}, _) $ l $ r) = l :: dest_tuples r
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    69
  | dest_tuples t = [t]
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    70
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    71
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    72
fun last_field sep str =
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    73
  let
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    74
    val n = size sep
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    75
    val len = size str
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    76
    fun find i =
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    77
      if i < 0 then NONE
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    78
      else if String.substring (str, i, n) = sep then SOME i
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    79
      else find (i - 1)
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    80
  in
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    81
    (case find (len - n) of
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    82
      NONE => NONE
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    83
    | SOME i => SOME (String.substring (str, 0, i), String.extract (str, i + n, NONE)))
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    84
  end
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    85
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    86
fun split_first_last start stop s =
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    87
  (case first_field start s of
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    88
    NONE => NONE
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    89
  | SOME (initial, rest) =>
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    90
      (case last_field stop rest of
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    91
        NONE => NONE
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    92
      | SOME (middle, tail) => SOME (initial, middle, tail)))
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    93
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    94
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    95
(* data slot for drivers *)
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    96
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    97
structure Drivers = Theory_Data
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    98
(
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
    99
  type T =
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   100
    (string * ((Proof.context -> (string * string) list * string -> Path.T -> string) * string)) list
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   101
  val empty = []
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   102
  val extend = I
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   103
  fun merge data : T = AList.merge (op =) (K true) data
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   104
)
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   105
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   106
val add_driver = Drivers.map o AList.update (op =)
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   107
val get_driver = AList.lookup (op =) o Drivers.get
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   108
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   109
(*
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   110
  Test drivers must produce output of the following format:
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   111
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   112
  The start of the relevant data is marked with start_markerN,
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   113
  its end with end_markerN.
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   114
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   115
  Between these two markers, every line corresponds to one test.
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   116
  Lines of successful tests start with successN, failures start with failureN.
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   117
  The failure failureN may continue with the YXML encoding of the evaluated term.
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   118
  There must not be any additional whitespace in between.
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   119
*)
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   120
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   121
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   122
(* parsing of results *)
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   123
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   124
val successN = "True"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   125
val failureN = "False"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   126
val start_markerN = "*@*Isabelle/Code_Test-start*@*"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   127
val end_markerN = "*@*Isabelle/Code_Test-end*@*"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   128
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   129
fun parse_line line =
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   130
  if String.isPrefix successN line then (true, NONE)
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   131
  else if String.isPrefix failureN line then (false,
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   132
    if size line > size failureN then
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   133
      String.extract (line, size failureN, NONE)
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   134
      |> YXML.parse_body
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   135
      |> Term_XML.Decode.term
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   136
      |> dest_tuples
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   137
      |> SOME
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   138
    else NONE)
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   139
  else raise Fail ("Cannot parse result of evaluation:\n" ^ line)
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   140
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   141
fun parse_result target out =
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   142
  (case split_first_last start_markerN end_markerN out of
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   143
    NONE => error ("Evaluation failed for " ^ target ^ "!\nBash output:\n" ^ out)
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   144
  | SOME (_, middle, _) => middle |> trim_line |> split_lines |> map parse_line)
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   145
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   146
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   147
(* pretty printing of test results *)
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   148
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   149
fun pretty_eval _ NONE _ = []
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   150
  | pretty_eval ctxt (SOME evals) ts =
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   151
      [Pretty.fbrk,
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   152
       Pretty.big_list "Evaluated terms"
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   153
         (map (fn (t, eval) => Pretty.block
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   154
           [Syntax.pretty_term ctxt t, Pretty.brk 1, Pretty.str "=", Pretty.brk 1,
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   155
            Syntax.pretty_term ctxt eval])
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   156
         (ts ~~ evals))]
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   157
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   158
fun pretty_failure ctxt target (((_, evals), query), eval_ts) =
60022
ea987317a785 tuned whitespace;
wenzelm
parents: 59936
diff changeset
   159
  Pretty.block (Pretty.text ("Test in " ^ target ^ " failed for")
ea987317a785 tuned whitespace;
wenzelm
parents: 59936
diff changeset
   160
    @ [Pretty.brk 1, Pretty.quote (Syntax.pretty_term ctxt query)]
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   161
    @ pretty_eval ctxt evals eval_ts)
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   162
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   163
fun pretty_failures ctxt target failures =
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   164
  Pretty.blk (0, Pretty.fbreaks (map (pretty_failure ctxt target) failures))
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   165
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   166
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   167
(* driver invocation *)
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   168
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   169
val overlord = Attrib.setup_config_bool @{binding "code_test_overlord"} (K false)
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   170
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   171
fun with_overlord_dir name f =
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   172
  let
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   173
    val path =
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   174
      Path.append (Path.explode "$ISABELLE_HOME_USER") (Path.basic (name ^ serial_string ()))
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   175
    val _ = Isabelle_System.mkdirs path
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   176
  in Exn.release (Exn.capture f path) end
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   177
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   178
fun dynamic_value_strict ctxt t compiler =
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   179
  let
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   180
    val thy = Proof_Context.theory_of ctxt
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   181
    val (driver, target) =
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   182
      (case get_driver thy compiler of
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   183
        NONE => error ("No driver for target " ^ compiler)
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   184
      | SOME f => f)
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   185
    val debug = Config.get (Proof_Context.init_global thy) overlord
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   186
    val with_dir = if debug then with_overlord_dir else Isabelle_System.with_tmp_dir
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   187
    fun evaluate f = with_dir "Code_Test" (driver ctxt f) |> parse_result compiler
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   188
    fun evaluator program _ vs_ty deps =
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   189
      Exn.interruptible_capture evaluate
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   190
        (Code_Target.computation_text ctxt target program deps true vs_ty)
64578
wenzelm
parents: 64577
diff changeset
   191
    fun postproc f = map (apsnd (Option.map (map f)))
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   192
  in Exn.release (Code_Thingol.dynamic_value ctxt (Exn.map_res o postproc) evaluator t) end
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   193
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   194
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   195
(* term preprocessing *)
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   196
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   197
fun add_eval (Const (@{const_name Trueprop}, _) $ t) = add_eval t
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   198
  | add_eval (Const (@{const_name "HOL.eq"}, _) $ lhs $ rhs) = (fn acc =>
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   199
      acc
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   200
      |> add_eval rhs
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   201
      |> add_eval lhs
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   202
      |> cons rhs
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   203
      |> cons lhs)
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   204
  | add_eval (Const (@{const_name "Not"}, _) $ t) = add_eval t
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   205
  | add_eval (Const (@{const_name "Orderings.ord_class.less_eq"}, _) $ lhs $ rhs) = (fn acc =>
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   206
      lhs :: rhs :: acc)
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   207
  | add_eval (Const (@{const_name "Orderings.ord_class.less"}, _) $ lhs $ rhs) = (fn acc =>
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   208
      lhs :: rhs :: acc)
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   209
  | add_eval _ = I
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   210
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   211
fun mk_term_of [] = @{term "None :: (unit \<Rightarrow> yxml_of_term) option"}
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   212
  | mk_term_of ts =
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   213
      let
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   214
        val tuple = mk_tuples ts
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   215
        val T = fastype_of tuple
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   216
      in
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   217
        @{term "Some :: (unit \<Rightarrow> yxml_of_term) \<Rightarrow> (unit \<Rightarrow> yxml_of_term) option"} $
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   218
          (absdummy @{typ unit} (@{const yxml_string_of_term} $
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   219
            (Const (@{const_name Code_Evaluation.term_of}, T --> @{typ term}) $ tuple)))
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   220
      end
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   221
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   222
fun test_terms ctxt ts target =
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   223
  let
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   224
    val thy = Proof_Context.theory_of ctxt
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   225
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   226
    fun term_of t = Sign.of_sort thy (fastype_of t, @{sort term_of})
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   227
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   228
    fun ensure_bool t =
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   229
      (case fastype_of t of
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   230
        @{typ bool} => ()
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   231
      | _ => error ("Test case not of type bool: " ^ Pretty.string_of (Syntax.pretty_term ctxt t)))
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   232
64578
wenzelm
parents: 64577
diff changeset
   233
    val _ = List.app ensure_bool ts
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   234
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   235
    val evals = map (fn t => filter term_of (add_eval t [])) ts
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   236
    val eval = map mk_term_of evals
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   237
64578
wenzelm
parents: 64577
diff changeset
   238
    val t =
wenzelm
parents: 64577
diff changeset
   239
      HOLogic.mk_list @{typ "bool \<times> (unit \<Rightarrow> yxml_of_term) option"}
wenzelm
parents: 64577
diff changeset
   240
        (map HOLogic.mk_prod (ts ~~ eval))
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   241
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   242
    val result = dynamic_value_strict ctxt t target
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   243
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   244
    val failed =
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   245
      filter_out (fst o fst o fst) (result ~~ ts ~~ evals)
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   246
      handle ListPair.UnequalLengths =>
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   247
        error ("Evaluation failed!\nWrong number of test results: " ^ Int.toString (length result))
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   248
  in
64578
wenzelm
parents: 64577
diff changeset
   249
    (case failed of [] =>
wenzelm
parents: 64577
diff changeset
   250
      ()
wenzelm
parents: 64577
diff changeset
   251
    | _ => error (Pretty.string_of (pretty_failures ctxt target failed)))
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   252
  end
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   253
64578
wenzelm
parents: 64577
diff changeset
   254
fun test_targets ctxt = List.app o test_terms ctxt
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   255
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   256
fun test_code_cmd raw_ts targets state =
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   257
  let
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   258
    val ctxt = Toplevel.context_of state
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   259
    val ts = Syntax.read_terms ctxt raw_ts
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   260
    val frees = fold Term.add_free_names ts []
64578
wenzelm
parents: 64577
diff changeset
   261
    val _ =
wenzelm
parents: 64577
diff changeset
   262
      if null frees then ()
wenzelm
parents: 64577
diff changeset
   263
      else error ("Terms contain free variables: " ^
wenzelm
parents: 64577
diff changeset
   264
        Pretty.string_of (Pretty.block (Pretty.commas (map Pretty.str frees))))
wenzelm
parents: 64577
diff changeset
   265
  in test_targets ctxt ts targets end
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   266
58348
2d47c7d10b62 add target language evaluators for the value command;
Andreas Lochbihler
parents: 58039
diff changeset
   267
fun eval_term target ctxt t =
2d47c7d10b62 add target language evaluators for the value command;
Andreas Lochbihler
parents: 58039
diff changeset
   268
  let
2d47c7d10b62 add target language evaluators for the value command;
Andreas Lochbihler
parents: 58039
diff changeset
   269
    val frees = Term.add_free_names t []
64578
wenzelm
parents: 64577
diff changeset
   270
    val _ =
wenzelm
parents: 64577
diff changeset
   271
      if null frees then ()
wenzelm
parents: 64577
diff changeset
   272
      else error ("Term contains free variables: " ^
wenzelm
parents: 64577
diff changeset
   273
        Pretty.string_of (Pretty.block (Pretty.commas (map Pretty.str frees))))
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   274
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   275
    val thy = Proof_Context.theory_of ctxt
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   276
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   277
    val T_t = fastype_of t
64578
wenzelm
parents: 64577
diff changeset
   278
    val _ =
wenzelm
parents: 64577
diff changeset
   279
      if Sign.of_sort thy (T_t, @{sort term_of}) then ()
wenzelm
parents: 64577
diff changeset
   280
      else error ("Type " ^ Pretty.string_of (Syntax.pretty_typ ctxt T_t) ^
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   281
       " of term not of sort " ^ Pretty.string_of (Syntax.pretty_sort ctxt @{sort term_of}))
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   282
64578
wenzelm
parents: 64577
diff changeset
   283
    val t' =
wenzelm
parents: 64577
diff changeset
   284
      HOLogic.mk_list @{typ "bool \<times> (unit \<Rightarrow> yxml_of_term) option"}
wenzelm
parents: 64577
diff changeset
   285
        [HOLogic.mk_prod (@{term "False"}, mk_term_of [t])]
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   286
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   287
    val result = dynamic_value_strict ctxt t' target
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   288
  in (case result of [(_, SOME [t])] => t | _ => error "Evaluation failed") end
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   289
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   290
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   291
(* generic driver *)
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   292
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   293
fun gen_driver mk_driver env_var env_var_dest compilerN ctxt (code_files, value_name) =
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   294
  let
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   295
    val compiler = getenv env_var
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   296
    val _ =
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   297
      if compiler <> "" then ()
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   298
      else error (Pretty.string_of (Pretty.para
64578
wenzelm
parents: 64577
diff changeset
   299
        ("Environment variable " ^ env_var ^ " is not set. To test code generation with " ^
wenzelm
parents: 64577
diff changeset
   300
          compilerN ^ ", set this variable to your " ^ env_var_dest ^ " in the settings file.")))
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   301
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   302
    fun compile NONE = ()
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   303
      | compile (SOME cmd) =
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   304
          let
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   305
            val (out, ret) = Isabelle_System.bash_output cmd
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   306
          in
64578
wenzelm
parents: 64577
diff changeset
   307
            if ret = 0 then ()
wenzelm
parents: 64577
diff changeset
   308
            else error ("Compilation with " ^ compilerN ^ " failed:\n" ^ cmd ^ "\n" ^ out)
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   309
          end
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   310
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   311
    fun run path =
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   312
      let
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   313
        val modules = map fst code_files
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   314
        val {files, compile_cmd, run_cmd, mk_code_file} = mk_driver ctxt path modules value_name
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   315
64578
wenzelm
parents: 64577
diff changeset
   316
        val _ = List.app (fn (name, code) => File.write (mk_code_file name) code) code_files
wenzelm
parents: 64577
diff changeset
   317
        val _ = List.app (fn (name, content) => File.write name content) files
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   318
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   319
        val _ = compile compile_cmd
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   320
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   321
        val (out, res) = Isabelle_System.bash_output run_cmd
64578
wenzelm
parents: 64577
diff changeset
   322
        val _ =
wenzelm
parents: 64577
diff changeset
   323
          if res = 0 then ()
wenzelm
parents: 64577
diff changeset
   324
          else error ("Evaluation for " ^ compilerN ^ " terminated with error code " ^
wenzelm
parents: 64577
diff changeset
   325
            Int.toString res ^ "\nBash output:\n" ^ out)
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   326
      in out end
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   327
  in run end
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   328
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   329
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   330
(* driver for PolyML *)
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   331
58415
8392d221bd91 clarified ISABELLE_POLYML;
wenzelm
parents: 58348
diff changeset
   332
val ISABELLE_POLYML = "ISABELLE_POLYML"
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   333
val polymlN = "PolyML"
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   334
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   335
fun mk_driver_polyml _ path _ value_name =
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   336
  let
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   337
    val generatedN = "generated.sml"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   338
    val driverN = "driver.sml"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   339
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   340
    val code_path = Path.append path (Path.basic generatedN)
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   341
    val driver_path = Path.append path (Path.basic driverN)
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   342
    val driver =
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   343
      "fun main prog_name = \n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   344
      "  let\n" ^
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   345
      "    fun format_term NONE = \"\"\n" ^
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   346
      "      | format_term (SOME t) = t ();\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   347
      "    fun format (true, _) = \"" ^ successN ^ "\\n\"\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   348
      "      | format (false, to) = \"" ^ failureN ^ "\" ^ format_term to ^ \"\\n\";\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   349
      "    val result = " ^ value_name ^ " ();\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   350
      "    val _ = print \"" ^ start_markerN ^ "\";\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   351
      "    val _ = map (print o format) result;\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   352
      "    val _ = print \"" ^ end_markerN ^ "\";\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   353
      "  in\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   354
      "    ()\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   355
      "  end;\n"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   356
    val cmd =
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   357
      "echo \"use \\\"" ^ Path.implode code_path ^ "\\\"; use \\\"" ^
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   358
      Path.implode driver_path ^ "\\\"; main ();\" | " ^
58415
8392d221bd91 clarified ISABELLE_POLYML;
wenzelm
parents: 58348
diff changeset
   359
      Path.implode (Path.variable ISABELLE_POLYML)
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   360
  in
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   361
    {files = [(driver_path, driver)], compile_cmd = NONE, run_cmd = cmd, mk_code_file = K code_path}
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   362
  end
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   363
58832
ec9550bd5fd7 make SML/NJ more happy;
wenzelm
parents: 58626
diff changeset
   364
fun evaluate_in_polyml ctxt =
ec9550bd5fd7 make SML/NJ more happy;
wenzelm
parents: 58626
diff changeset
   365
  gen_driver mk_driver_polyml ISABELLE_POLYML "PolyML executable" polymlN ctxt
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   366
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   367
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   368
(* driver for mlton *)
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   369
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   370
val mltonN = "MLton"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   371
val ISABELLE_MLTON = "ISABELLE_MLTON"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   372
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   373
fun mk_driver_mlton _ path _ value_name =
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   374
  let
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   375
    val generatedN = "generated.sml"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   376
    val driverN = "driver.sml"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   377
    val projectN = "test"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   378
    val ml_basisN = projectN ^ ".mlb"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   379
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   380
    val code_path = Path.append path (Path.basic generatedN)
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   381
    val driver_path = Path.append path (Path.basic driverN)
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   382
    val ml_basis_path = Path.append path (Path.basic ml_basisN)
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   383
    val driver =
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   384
      "fun format_term NONE = \"\"\n" ^
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   385
      "  | format_term (SOME t) = t ();\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   386
      "fun format (true, _) = \"" ^ successN ^ "\\n\"\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   387
      "  | format (false, to) = \"" ^ failureN ^ "\" ^ format_term to ^ \"\\n\";\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   388
      "val result = " ^ value_name ^ " ();\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   389
      "val _ = print \"" ^ start_markerN ^ "\";\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   390
      "val _ = map (print o format) result;\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   391
      "val _ = print \"" ^ end_markerN ^ "\";\n"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   392
    val ml_basis =
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   393
      "$(SML_LIB)/basis/basis.mlb\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   394
      generatedN ^ "\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   395
      driverN
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   396
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   397
    val compile_cmd =
62549
9498623b27f0 File.bash_string operations in ML as in Scala -- exclusively for GNU bash, not perl and not user output;
wenzelm
parents: 61077
diff changeset
   398
      File.bash_path (Path.variable ISABELLE_MLTON) ^
9498623b27f0 File.bash_string operations in ML as in Scala -- exclusively for GNU bash, not perl and not user output;
wenzelm
parents: 61077
diff changeset
   399
      " -default-type intinf " ^ File.bash_path ml_basis_path
9498623b27f0 File.bash_string operations in ML as in Scala -- exclusively for GNU bash, not perl and not user output;
wenzelm
parents: 61077
diff changeset
   400
    val run_cmd = File.bash_path (Path.append path (Path.basic projectN))
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   401
  in
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   402
    {files = [(driver_path, driver), (ml_basis_path, ml_basis)],
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   403
     compile_cmd = SOME compile_cmd, run_cmd = run_cmd, mk_code_file = K code_path}
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   404
  end
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   405
58832
ec9550bd5fd7 make SML/NJ more happy;
wenzelm
parents: 58626
diff changeset
   406
fun evaluate_in_mlton ctxt =
ec9550bd5fd7 make SML/NJ more happy;
wenzelm
parents: 58626
diff changeset
   407
  gen_driver mk_driver_mlton ISABELLE_MLTON "MLton executable" mltonN ctxt
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   408
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   409
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   410
(* driver for SML/NJ *)
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   411
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   412
val smlnjN = "SMLNJ"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   413
val ISABELLE_SMLNJ = "ISABELLE_SMLNJ"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   414
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   415
fun mk_driver_smlnj _ path _ value_name =
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   416
  let
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   417
    val generatedN = "generated.sml"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   418
    val driverN = "driver.sml"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   419
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   420
    val code_path = Path.append path (Path.basic generatedN)
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   421
    val driver_path = Path.append path (Path.basic driverN)
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   422
    val driver =
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   423
      "structure Test = struct\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   424
      "fun main prog_name =\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   425
      "  let\n" ^
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   426
      "    fun format_term NONE = \"\"\n" ^
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   427
      "      | format_term (SOME t) = t ();\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   428
      "    fun format (true, _) = \"" ^ successN ^ "\\n\"\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   429
      "      | format (false, to) = \"" ^ failureN ^ "\" ^ format_term to ^ \"\\n\";\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   430
      "    val result = " ^ value_name ^ " ();\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   431
      "    val _ = print \"" ^ start_markerN ^ "\";\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   432
      "    val _ = map (print o format) result;\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   433
      "    val _ = print \"" ^ end_markerN ^ "\";\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   434
      "  in\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   435
      "    0\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   436
      "  end;\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   437
      "end;"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   438
    val cmd =
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   439
      "echo \"Control.MC.matchRedundantError := false; Control.MC.matchRedundantWarn := false;" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   440
      "use \\\"" ^ Path.implode code_path ^ "\\\"; use \\\"" ^ Path.implode driver_path ^ "\\\";" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   441
      "Test.main ();\" | " ^ Path.implode (Path.variable ISABELLE_SMLNJ)
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   442
  in
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   443
    {files = [(driver_path, driver)], compile_cmd = NONE, run_cmd = cmd, mk_code_file = K code_path}
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   444
  end
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   445
58832
ec9550bd5fd7 make SML/NJ more happy;
wenzelm
parents: 58626
diff changeset
   446
fun evaluate_in_smlnj ctxt =
ec9550bd5fd7 make SML/NJ more happy;
wenzelm
parents: 58626
diff changeset
   447
  gen_driver mk_driver_smlnj ISABELLE_SMLNJ "SMLNJ executable" smlnjN ctxt
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   448
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   449
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   450
(* driver for OCaml *)
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   451
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   452
val ocamlN = "OCaml"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   453
val ISABELLE_OCAMLC = "ISABELLE_OCAMLC"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   454
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   455
fun mk_driver_ocaml _ path _ value_name =
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   456
  let
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   457
    val generatedN = "generated.ml"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   458
    val driverN = "driver.ml"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   459
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   460
    val code_path = Path.append path (Path.basic generatedN)
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   461
    val driver_path = Path.append path (Path.basic driverN)
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   462
    val driver =
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   463
      "let format_term = function\n" ^
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   464
      "  | None -> \"\"\n" ^
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   465
      "  | Some t -> t ();;\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   466
      "let format = function\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   467
      "  | (true, _) -> \"" ^ successN ^ "\\n\"\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   468
      "  | (false, x) -> \"" ^ failureN ^ "\" ^ format_term x ^ \"\\n\";;\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   469
      "let result = " ^ ("Generated." ^ value_name) ^ " ();;\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   470
      "let main x =\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   471
      "  let _ = print_string \"" ^ start_markerN ^ "\" in\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   472
      "  let _ = List.map (fun x -> print_string (format x)) result in\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   473
      "  print_string \"" ^ end_markerN ^ "\";;\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   474
      "main ();;"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   475
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   476
    val compiled_path = Path.append path (Path.basic "test")
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   477
    val compile_cmd =
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   478
      Path.implode (Path.variable ISABELLE_OCAMLC) ^ " -w pu -o " ^ Path.implode compiled_path ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   479
      " -I " ^ Path.implode path ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   480
      " nums.cma " ^ Path.implode code_path ^ " " ^ Path.implode driver_path
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   481
62549
9498623b27f0 File.bash_string operations in ML as in Scala -- exclusively for GNU bash, not perl and not user output;
wenzelm
parents: 61077
diff changeset
   482
    val run_cmd = File.bash_path compiled_path
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   483
  in
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   484
    {files = [(driver_path, driver)],
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   485
     compile_cmd = SOME compile_cmd, run_cmd = run_cmd, mk_code_file = K code_path}
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   486
  end
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   487
58832
ec9550bd5fd7 make SML/NJ more happy;
wenzelm
parents: 58626
diff changeset
   488
fun evaluate_in_ocaml ctxt =
ec9550bd5fd7 make SML/NJ more happy;
wenzelm
parents: 58626
diff changeset
   489
  gen_driver mk_driver_ocaml ISABELLE_OCAMLC "ocamlc executable" ocamlN ctxt
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   490
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   491
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   492
(* driver for GHC *)
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   493
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   494
val ghcN = "GHC"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   495
val ISABELLE_GHC = "ISABELLE_GHC"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   496
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   497
val ghc_options = Attrib.setup_config_string @{binding code_test_ghc} (K "")
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   498
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   499
fun mk_driver_ghc ctxt path modules value_name =
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   500
  let
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   501
    val driverN = "Main.hs"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   502
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   503
    fun mk_code_file name = Path.append path (Path.basic (name ^ ".hs"))
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   504
    val driver_path = Path.append path (Path.basic driverN)
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   505
    val driver =
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   506
      "module Main where {\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   507
      String.concat (map (fn module => "import qualified " ^ module ^ ";\n") modules) ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   508
      "main = do {\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   509
      "    let {\n" ^
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   510
      "      format_term Nothing = \"\";\n" ^
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   511
      "      format_term (Just t) = t ();\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   512
      "      format (True, _) = \"" ^ successN ^ "\\n\";\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   513
      "      format (False, to) = \"" ^ failureN ^ "\" ++ format_term to ++ \"\\n\";\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   514
      "      result = " ^ value_name ^ " ();\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   515
      "    };\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   516
      "    Prelude.putStr \"" ^ start_markerN ^ "\";\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   517
      "    Prelude.mapM_ (putStr . format) result;\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   518
      "    Prelude.putStr \"" ^ end_markerN ^ "\";\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   519
      "  }\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   520
      "}\n"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   521
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   522
    val compiled_path = Path.append path (Path.basic "test")
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   523
    val compile_cmd =
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   524
      Path.implode (Path.variable ISABELLE_GHC) ^ " " ^ Code_Haskell.language_params ^ " " ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   525
      Config.get ctxt ghc_options ^ " -o " ^ Path.implode compiled_path ^ " " ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   526
      Path.implode driver_path ^ " -i" ^ Path.implode path
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   527
62549
9498623b27f0 File.bash_string operations in ML as in Scala -- exclusively for GNU bash, not perl and not user output;
wenzelm
parents: 61077
diff changeset
   528
    val run_cmd = File.bash_path compiled_path
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   529
  in
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   530
    {files = [(driver_path, driver)],
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   531
     compile_cmd = SOME compile_cmd, run_cmd = run_cmd, mk_code_file = mk_code_file}
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   532
  end
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   533
58832
ec9550bd5fd7 make SML/NJ more happy;
wenzelm
parents: 58626
diff changeset
   534
fun evaluate_in_ghc ctxt =
ec9550bd5fd7 make SML/NJ more happy;
wenzelm
parents: 58626
diff changeset
   535
  gen_driver mk_driver_ghc ISABELLE_GHC "GHC executable" ghcN ctxt
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   536
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   537
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   538
(* driver for Scala *)
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   539
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   540
val scalaN = "Scala"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   541
val ISABELLE_SCALA = "ISABELLE_SCALA"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   542
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   543
fun mk_driver_scala _ path _ value_name =
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   544
  let
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   545
    val generatedN = "Generated_Code"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   546
    val driverN = "Driver.scala"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   547
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   548
    val code_path = Path.append path (Path.basic (generatedN ^ ".scala"))
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   549
    val driver_path = Path.append path (Path.basic driverN)
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   550
    val driver =
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   551
      "import " ^ generatedN ^ "._\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   552
      "object Test {\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   553
      "  def format_term(x : Option[Unit => String]) : String = x match {\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   554
      "    case None => \"\"\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   555
      "    case Some(x) => x(())\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   556
      "  }\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   557
      "  def format(term : (Boolean, Option[Unit => String])) : String = term match {\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   558
      "      case (true, _) => \"True\\n\"\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   559
      "      case (false, x) => \"False\" + format_term(x) + \"\\n\"\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   560
      "  }\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   561
      "  def main(args:Array[String]) = {\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   562
      "    val result = " ^ value_name ^ "(());\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   563
      "    print(\"" ^ start_markerN ^ "\");\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   564
      "    result.map{test:(Boolean, Option[Unit => String]) => print(format(test))};\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   565
      "    print(\"" ^ end_markerN ^ "\");\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   566
      "  }\n" ^
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   567
      "}\n"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   568
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   569
    val compile_cmd =
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   570
      Path.implode (Path.append (Path.variable ISABELLE_SCALA) (Path.basic "scalac")) ^
62549
9498623b27f0 File.bash_string operations in ML as in Scala -- exclusively for GNU bash, not perl and not user output;
wenzelm
parents: 61077
diff changeset
   571
      " -d " ^ File.bash_path path ^ " -classpath " ^ File.bash_path path ^ " " ^
9498623b27f0 File.bash_string operations in ML as in Scala -- exclusively for GNU bash, not perl and not user output;
wenzelm
parents: 61077
diff changeset
   572
      File.bash_path code_path ^ " " ^ File.bash_path driver_path
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   573
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   574
    val run_cmd =
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   575
      Path.implode (Path.append (Path.variable ISABELLE_SCALA) (Path.basic "scala")) ^
62549
9498623b27f0 File.bash_string operations in ML as in Scala -- exclusively for GNU bash, not perl and not user output;
wenzelm
parents: 61077
diff changeset
   576
      " -cp " ^ File.bash_path path ^ " Test"
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   577
  in
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   578
    {files = [(driver_path, driver)],
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   579
     compile_cmd = SOME compile_cmd, run_cmd = run_cmd, mk_code_file = K code_path}
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   580
  end
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   581
58832
ec9550bd5fd7 make SML/NJ more happy;
wenzelm
parents: 58626
diff changeset
   582
fun evaluate_in_scala ctxt =
ec9550bd5fd7 make SML/NJ more happy;
wenzelm
parents: 58626
diff changeset
   583
  gen_driver mk_driver_scala ISABELLE_SCALA "Scala directory" scalaN ctxt
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   584
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   585
val test_codeP = Scan.repeat1 Parse.prop -- (@{keyword "in"} |-- Scan.repeat1 Parse.name)
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   586
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   587
val _ =
59936
b8ffc3dc9e24 @{command_spec} is superseded by @{command_keyword};
wenzelm
parents: 59720
diff changeset
   588
  Outer_Syntax.command @{command_keyword test_code}
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   589
    "compile test cases to target languages, execute them and report results"
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   590
      (test_codeP >> (fn (raw_ts, targets) => Toplevel.keep (test_code_cmd raw_ts targets)))
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   591
59323
468bd3aedfa1 modernized and more uniform style
haftmann
parents: 58832
diff changeset
   592
val _ = Theory.setup (fold add_driver
468bd3aedfa1 modernized and more uniform style
haftmann
parents: 58832
diff changeset
   593
  [(polymlN, (evaluate_in_polyml, Code_ML.target_SML)),
468bd3aedfa1 modernized and more uniform style
haftmann
parents: 58832
diff changeset
   594
   (mltonN, (evaluate_in_mlton, Code_ML.target_SML)),
468bd3aedfa1 modernized and more uniform style
haftmann
parents: 58832
diff changeset
   595
   (smlnjN, (evaluate_in_smlnj, Code_ML.target_SML)),
468bd3aedfa1 modernized and more uniform style
haftmann
parents: 58832
diff changeset
   596
   (ocamlN, (evaluate_in_ocaml, Code_ML.target_OCaml)),
468bd3aedfa1 modernized and more uniform style
haftmann
parents: 58832
diff changeset
   597
   (ghcN, (evaluate_in_ghc, Code_Haskell.target)),
468bd3aedfa1 modernized and more uniform style
haftmann
parents: 58832
diff changeset
   598
   (scalaN, (evaluate_in_scala, Code_Scala.target))]
63806
c54a53ef1873 clarified modules;
wenzelm
parents: 63157
diff changeset
   599
  #> fold (fn target => Value_Command.add_evaluator (target, eval_term target))
64577
0288a566c966 tuned whitespace;
wenzelm
parents: 63806
diff changeset
   600
    [polymlN, mltonN, smlnjN, ocamlN, ghcN, scalaN])
59323
468bd3aedfa1 modernized and more uniform style
haftmann
parents: 58832
diff changeset
   601
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
   602
end