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