Andreas@58039: (* Title: Code_Test.ML Andreas@58039: Author: Andreas Lochbihler, ETH Zurich Andreas@58039: Andreas@58039: Test infrastructure for the code generator Andreas@58039: *) Andreas@58039: Andreas@58039: signature CODE_TEST = sig Andreas@58039: val add_driver : string * ((Proof.context -> (string * string) list * string -> Path.T -> string) * string) -> theory -> theory Andreas@58039: val get_driver : theory -> string -> ((Proof.context -> (string * string) list * string -> Path.T -> string) * string) option Andreas@58039: val overlord : bool Config.T Andreas@58039: val successN : string Andreas@58039: val failureN : string Andreas@58039: val start_markerN : string Andreas@58039: val end_markerN : string Andreas@58039: val test_terms : Proof.context -> term list -> string -> unit Andreas@58039: val test_targets : Proof.context -> term list -> string list -> unit list Andreas@58039: val test_code_cmd : string list -> string list -> Toplevel.state -> unit Andreas@58039: Andreas@58348: val eval_term : string -> Proof.context -> term -> term Andreas@58039: Andreas@58039: val gen_driver : Andreas@58039: (theory -> Path.T -> string list -> string -> Andreas@58039: {files : (Path.T * string) list, Andreas@58039: compile_cmd : string option, run_cmd : string, mk_code_file : string -> Path.T}) Andreas@58039: -> string -> string -> string Andreas@58039: -> theory -> (string * string) list * string -> Path.T -> string Andreas@58039: wenzelm@58415: val ISABELLE_POLYML : string Andreas@58039: val polymlN : string Andreas@58039: val evaluate_in_polyml : Proof.context -> (string * string) list * string -> Path.T -> string Andreas@58039: Andreas@58039: val mltonN : string Andreas@58039: val ISABELLE_MLTON : string Andreas@58039: val evaluate_in_mlton : Proof.context -> (string * string) list * string -> Path.T -> string Andreas@58039: Andreas@58039: val smlnjN : string Andreas@58039: val ISABELLE_SMLNJ : string Andreas@58039: val evaluate_in_smlnj : Proof.context -> (string * string) list * string -> Path.T -> string Andreas@58039: Andreas@58039: val ocamlN : string Andreas@58039: val ISABELLE_OCAMLC : string Andreas@58039: val evaluate_in_ocaml : Proof.context -> (string * string) list * string -> Path.T -> string Andreas@58039: Andreas@58039: val ghcN : string Andreas@58039: val ISABELLE_GHC : string Andreas@58039: val ghc_options : string Config.T Andreas@58039: val evaluate_in_ghc : Proof.context -> (string * string) list * string -> Path.T -> string Andreas@58039: Andreas@58039: val scalaN : string Andreas@58039: val ISABELLE_SCALA : string Andreas@58039: val evaluate_in_scala : Proof.context -> (string * string) list * string -> Path.T -> string Andreas@58039: end Andreas@58039: Andreas@58039: structure Code_Test : CODE_TEST = struct Andreas@58039: Andreas@58039: (* convert a list of terms into nested tuples and back *) Andreas@58039: fun mk_tuples [] = @{term "()"} Andreas@58039: | mk_tuples [t] = t Andreas@58039: | mk_tuples (t :: ts) = HOLogic.mk_prod (t, mk_tuples ts) Andreas@58039: Andreas@58039: fun dest_tuples (Const (@{const_name Pair}, _) $ l $ r) = l :: dest_tuples r Andreas@58039: | dest_tuples t = [t] Andreas@58039: Andreas@58039: Andreas@58039: fun map_option _ NONE = NONE Andreas@58039: | map_option f (SOME x) = SOME (f x) Andreas@58039: Andreas@58039: fun last_field sep str = Andreas@58039: let Andreas@58039: val n = size sep; Andreas@58039: val len = size str; Andreas@58039: fun find i = Andreas@58039: if i < 0 then NONE Andreas@58039: else if String.substring (str, i, n) = sep then SOME i Andreas@58039: else find (i - 1); Andreas@58039: in Andreas@58039: (case find (len - n) of Andreas@58039: NONE => NONE Andreas@58039: | SOME i => SOME (String.substring (str, 0, i), String.extract (str, i + n, NONE))) Andreas@58039: end; Andreas@58039: Andreas@58039: fun split_first_last start stop s = Andreas@58039: case first_field start s Andreas@58039: of NONE => NONE Andreas@58039: | SOME (initial, rest) => Andreas@58039: case last_field stop rest Andreas@58039: of NONE => NONE Andreas@58039: | SOME (middle, tail) => SOME (initial, middle, tail); Andreas@58039: Andreas@58039: (* Data slot for drivers *) Andreas@58039: Andreas@58039: structure Drivers = Theory_Data Andreas@58039: ( Andreas@58039: type T = (string * ((Proof.context -> (string * string) list * string -> Path.T -> string) * string)) list; Andreas@58039: val empty = []; Andreas@58039: val extend = I; Andreas@58039: fun merge data : T = AList.merge (op =) (K true) data; Andreas@58039: ) Andreas@58039: Andreas@58039: val add_driver = Drivers.map o AList.update (op =); Andreas@58039: val get_driver = AList.lookup (op =) o Drivers.get; Andreas@58039: Andreas@58039: (* Andreas@58039: Test drivers must produce output of the following format: Andreas@58039: Andreas@58039: The start of the relevant data is marked with start_markerN, Andreas@58039: its end with end_markerN. Andreas@58039: Andreas@58039: Between these two markers, every line corresponds to one test. Andreas@58039: Lines of successful tests start with successN, failures start with failureN. Andreas@58039: The failure failureN may continue with the YXML encoding of the evaluated term. Andreas@58039: There must not be any additional whitespace in between. Andreas@58039: *) Andreas@58039: Andreas@58039: (* Parsing of results *) Andreas@58039: Andreas@58039: val successN = "True" Andreas@58039: val failureN = "False" Andreas@58039: val start_markerN = "*@*Isabelle/Code_Test-start*@*" Andreas@58039: val end_markerN = "*@*Isabelle/Code_Test-end*@*" Andreas@58039: Andreas@58039: fun parse_line line = Andreas@58039: if String.isPrefix successN line then (true, NONE) Andreas@58039: else if String.isPrefix failureN line then (false, Andreas@58039: if size line > size failureN then Andreas@58039: String.extract (line, size failureN, NONE) Andreas@58039: |> YXML.parse_body Andreas@58039: |> Term_XML.Decode.term Andreas@58039: |> dest_tuples Andreas@58039: |> SOME Andreas@58039: else NONE) Andreas@58039: else raise Fail ("Cannot parse result of evaluation:\n" ^ line) Andreas@58039: Andreas@58039: fun parse_result target out = Andreas@58039: case split_first_last start_markerN end_markerN out Andreas@58039: of NONE => error ("Evaluation failed for " ^ target ^ "!\nBash output:\n" ^ out) Andreas@58039: | SOME (_, middle, _) => middle |> trim_line |> split_lines |> map parse_line Andreas@58039: Andreas@58039: (* Pretty printing of test results *) Andreas@58039: Andreas@58039: fun pretty_eval _ NONE _ = [] Andreas@58039: | pretty_eval ctxt (SOME evals) ts = Andreas@58039: [Pretty.fbrk, Andreas@58039: Pretty.big_list "Evaluated terms" Andreas@58039: (map (fn (t, eval) => Pretty.block Andreas@58039: [Syntax.pretty_term ctxt t, Pretty.brk 1, Pretty.str "=", Pretty.brk 1, Andreas@58039: Syntax.pretty_term ctxt eval]) Andreas@58039: (ts ~~ evals))] Andreas@58039: Andreas@58039: fun pretty_failure ctxt target (((_, evals), query), eval_ts) = Andreas@58039: Pretty.block (Pretty.text ("Test in " ^ target ^ " failed for") @ [Pretty.brk 1, Pretty.quote (Syntax.pretty_term ctxt query)] Andreas@58039: @ pretty_eval ctxt evals eval_ts) Andreas@58039: Andreas@58039: fun pretty_failures ctxt target failures = Andreas@58039: Pretty.blk (0, Pretty.fbreaks (map (pretty_failure ctxt target) failures)) Andreas@58039: Andreas@58039: (* Driver invocation *) Andreas@58039: Andreas@58039: val overlord = Attrib.setup_config_bool @{binding "code_test_overlord"} (K false); Andreas@58039: Andreas@58039: fun with_overlord_dir name f = Andreas@58039: let Andreas@58039: val path = Path.append (Path.explode "$ISABELLE_HOME_USER") (Path.basic (name ^ serial_string ())) Andreas@58039: val _ = Isabelle_System.mkdirs path; Andreas@58039: in Andreas@58039: Exn.release (Exn.capture f path) Andreas@58039: end; Andreas@58039: Andreas@58039: fun dynamic_value_strict ctxt t compiler = Andreas@58039: let Andreas@58039: val thy = Proof_Context.theory_of ctxt Andreas@58039: val (driver, target) = case get_driver thy compiler Andreas@58039: of NONE => error ("No driver for target " ^ compiler) Andreas@58039: | SOME f => f; Andreas@58039: val debug = Config.get (Proof_Context.init_global thy) overlord Andreas@58039: val with_dir = if debug then with_overlord_dir else Isabelle_System.with_tmp_dir Andreas@58039: fun evaluate f = with_dir "Code_Test" (driver ctxt f) |> parse_result compiler Andreas@58039: fun evaluator program _ vs_ty deps = Andreas@58039: Exn.interruptible_capture evaluate (Code_Target.evaluator ctxt target program deps true vs_ty); Andreas@58039: fun postproc f = map (apsnd (map_option (map f))) Andreas@58039: in Andreas@58039: Exn.release (Code_Thingol.dynamic_value ctxt (Exn.map_result o postproc) evaluator t) Andreas@58039: end; Andreas@58039: Andreas@58039: (* Term preprocessing *) Andreas@58039: Andreas@58039: fun add_eval (Const (@{const_name Trueprop}, _) $ t) = add_eval t Andreas@58039: | add_eval (Const (@{const_name "HOL.eq"}, _) $ lhs $ rhs) = (fn acc => Andreas@58039: acc Andreas@58039: |> add_eval rhs Andreas@58039: |> add_eval lhs Andreas@58039: |> cons rhs Andreas@58039: |> cons lhs) Andreas@58039: | add_eval (Const (@{const_name "Not"}, _) $ t) = add_eval t Andreas@58039: | add_eval (Const (@{const_name "Orderings.ord_class.less_eq"}, _) $ lhs $ rhs) = (fn acc => Andreas@58039: lhs :: rhs :: acc) Andreas@58039: | add_eval (Const (@{const_name "Orderings.ord_class.less"}, _) $ lhs $ rhs) = (fn acc => Andreas@58039: lhs :: rhs :: acc) Andreas@58039: | add_eval _ = I Andreas@58039: Andreas@58039: fun mk_term_of [] = @{term "None :: (unit \ yxml_of_term) option"} Andreas@58039: | mk_term_of ts = Andreas@58039: let Andreas@58039: val tuple = mk_tuples ts Andreas@58039: val T = fastype_of tuple Andreas@58039: in Andreas@58039: @{term "Some :: (unit \ yxml_of_term) \ (unit \ yxml_of_term) option"} $ Andreas@58039: (absdummy @{typ unit} (@{const yxml_string_of_term} $ Andreas@58039: (Const (@{const_name Code_Evaluation.term_of}, T --> @{typ term}) $ tuple))) Andreas@58039: end Andreas@58039: Andreas@58039: fun test_terms ctxt ts target = Andreas@58039: let Andreas@58039: val thy = Proof_Context.theory_of ctxt Andreas@58039: Andreas@58039: fun term_of t = Sign.of_sort thy (fastype_of t, @{sort term_of}) Andreas@58039: Andreas@58039: fun ensure_bool t = case fastype_of t of @{typ bool} => () Andreas@58039: | _ => error ("Test case not of type bool: " ^ Pretty.string_of (Syntax.pretty_term ctxt t)) Andreas@58039: Andreas@58039: val _ = map ensure_bool ts Andreas@58039: Andreas@58039: val evals = map (fn t => filter term_of (add_eval t [])) ts Andreas@58039: val eval = map mk_term_of evals Andreas@58039: Andreas@58039: val T = HOLogic.mk_prodT (@{typ bool}, Type (@{type_name option}, [@{typ unit} --> @{typ yxml_of_term}])) Andreas@58039: val t = HOLogic.mk_list T (map HOLogic.mk_prod (ts ~~ eval)) Andreas@58039: Andreas@58039: val result = dynamic_value_strict ctxt t target; Andreas@58039: Andreas@58039: val failed = Andreas@58039: filter_out (fst o fst o fst) (result ~~ ts ~~ evals) Andreas@58039: handle ListPair.UnequalLengths => Andreas@58039: error ("Evaluation failed!\nWrong number of test results: " ^ Int.toString (length result)) Andreas@58039: val _ = case failed of [] => () Andreas@58039: | _ => error (Pretty.string_of (pretty_failures ctxt target failed)) Andreas@58039: in Andreas@58039: () Andreas@58039: end Andreas@58039: Andreas@58039: fun test_targets ctxt = map o test_terms ctxt Andreas@58039: Andreas@58039: fun test_code_cmd raw_ts targets state = Andreas@58039: let Andreas@58039: val ctxt = Toplevel.context_of state; Andreas@58039: val ts = Syntax.read_terms ctxt raw_ts; Andreas@58039: val frees = fold Term.add_free_names ts [] Andreas@58039: val _ = if frees = [] then () else Andreas@58039: error ("Terms contain free variables: " ^ Andreas@58039: Pretty.string_of (Pretty.block (Pretty.commas (map Pretty.str frees)))) Andreas@58039: in Andreas@58039: test_targets ctxt ts targets; () Andreas@58039: end Andreas@58039: Andreas@58348: fun eval_term target ctxt t = Andreas@58348: let Andreas@58348: val frees = Term.add_free_names t [] Andreas@58348: val _ = if frees = [] then () else Andreas@58348: error ("Term contains free variables: " ^ Andreas@58348: Pretty.string_of (Pretty.block (Pretty.commas (map Pretty.str frees)))) Andreas@58039: Andreas@58039: val thy = Proof_Context.theory_of ctxt Andreas@58039: Andreas@58039: val T_t = fastype_of t Andreas@58039: val _ = if Sign.of_sort thy (T_t, @{sort term_of}) then () else error Andreas@58039: ("Type " ^ Pretty.string_of (Syntax.pretty_typ ctxt T_t) ^ Andreas@58039: " of term not of sort " ^ Pretty.string_of (Syntax.pretty_sort ctxt @{sort term_of})) Andreas@58039: Andreas@58039: val T = HOLogic.mk_prodT (@{typ bool}, Type (@{type_name option}, [@{typ unit} --> @{typ yxml_of_term}])) Andreas@58039: val t' = HOLogic.mk_list T [HOLogic.mk_prod (@{term "False"}, mk_term_of [t])] Andreas@58039: Andreas@58039: val result = dynamic_value_strict ctxt t' target; Andreas@58039: in Andreas@58348: case result of [(_, SOME [t])] => t | _ => error "Evaluation failed" Andreas@58039: end Andreas@58039: Andreas@58039: (* Generic driver *) Andreas@58039: Andreas@58039: fun gen_driver mk_driver env_var env_var_dest compilerN ctxt (code_files, value_name) = Andreas@58039: let Andreas@58039: val compiler = getenv env_var Andreas@58039: val _ = if compiler <> "" then () else error (Pretty.string_of (Pretty.para Andreas@58039: ("Environment variable " ^ env_var ^ " is not set. To test code generation with " ^ Andreas@58039: compilerN ^ ", set this variable to your " ^ env_var_dest ^ " in the settings file."))) Andreas@58039: Andreas@58039: fun compile NONE = () Andreas@58039: | compile (SOME cmd) = Andreas@58039: let Andreas@58039: val (out, ret) = Isabelle_System.bash_output cmd Andreas@58039: in Andreas@58039: if ret = 0 then () else error Andreas@58039: ("Compilation with " ^ compilerN ^ " failed:\n" ^ cmd ^ "\n" ^ out) Andreas@58039: end Andreas@58039: Andreas@58039: fun run (path : Path.T)= Andreas@58039: let Andreas@58039: val modules = map fst code_files Andreas@58039: val {files, compile_cmd, run_cmd, mk_code_file} Andreas@58039: = mk_driver ctxt path modules value_name Andreas@58039: Andreas@58039: val _ = map (fn (name, code) => File.write (mk_code_file name) code) code_files Andreas@58039: val _ = map (fn (name, content) => File.write name content) files Andreas@58039: Andreas@58039: val _ = compile compile_cmd Andreas@58039: Andreas@58039: val (out, res) = Isabelle_System.bash_output run_cmd Andreas@58039: val _ = if res = 0 then () else error Andreas@58039: ("Evaluation for " ^ compilerN ^ " terminated with error code " ^ Int.toString res ^ Andreas@58039: "\nBash output:\n" ^ out) Andreas@58039: in Andreas@58039: out Andreas@58039: end Andreas@58039: in Andreas@58039: run Andreas@58039: end Andreas@58039: Andreas@58039: (* Driver for PolyML *) Andreas@58039: wenzelm@58415: val ISABELLE_POLYML = "ISABELLE_POLYML" Andreas@58039: val polymlN = "PolyML"; Andreas@58039: Andreas@58039: fun mk_driver_polyml _ path _ value_name = Andreas@58039: let Andreas@58039: val generatedN = "generated.sml" Andreas@58039: val driverN = "driver.sml" Andreas@58039: Andreas@58039: val code_path = Path.append path (Path.basic generatedN) Andreas@58039: val driver_path = Path.append path (Path.basic driverN) Andreas@58039: val driver = Andreas@58039: "fun main prog_name = \n" ^ Andreas@58039: " let\n" ^ Andreas@58039: " fun format_term NONE = \"\"\n" ^ Andreas@58039: " | format_term (SOME t) = t ();\n" ^ Andreas@58039: " fun format (true, _) = \"" ^ successN ^ "\\n\"\n" ^ Andreas@58039: " | format (false, to) = \"" ^ failureN ^ "\" ^ format_term to ^ \"\\n\";\n" ^ Andreas@58039: " val result = " ^ value_name ^ " ();\n" ^ Andreas@58039: " val _ = print \"" ^ start_markerN ^ "\";\n" ^ Andreas@58039: " val _ = map (print o format) result;\n" ^ Andreas@58039: " val _ = print \"" ^ end_markerN ^ "\";\n" ^ Andreas@58039: " in\n" ^ Andreas@58039: " ()\n" ^ Andreas@58039: " end;\n" Andreas@58039: val cmd = Andreas@58039: "echo \"use \\\"" ^ Path.implode code_path ^ "\\\"; use \\\"" ^ Andreas@58039: Path.implode driver_path ^ "\\\"; main ();\" | " ^ wenzelm@58415: Path.implode (Path.variable ISABELLE_POLYML) Andreas@58039: in Andreas@58039: {files = [(driver_path, driver)], compile_cmd = NONE, run_cmd = cmd, mk_code_file = K code_path} Andreas@58039: end Andreas@58039: wenzelm@58415: val evaluate_in_polyml = gen_driver mk_driver_polyml ISABELLE_POLYML "PolyML executable" polymlN Andreas@58039: Andreas@58039: (* Driver for mlton *) Andreas@58039: Andreas@58039: val mltonN = "MLton" Andreas@58039: val ISABELLE_MLTON = "ISABELLE_MLTON" Andreas@58039: Andreas@58039: fun mk_driver_mlton _ path _ value_name = Andreas@58039: let Andreas@58039: val generatedN = "generated.sml" Andreas@58039: val driverN = "driver.sml" Andreas@58039: val projectN = "test" Andreas@58039: val ml_basisN = projectN ^ ".mlb" Andreas@58039: Andreas@58039: val code_path = Path.append path (Path.basic generatedN) Andreas@58039: val driver_path = Path.append path (Path.basic driverN) Andreas@58039: val ml_basis_path = Path.append path (Path.basic ml_basisN) Andreas@58039: val driver = Andreas@58039: "fun format_term NONE = \"\"\n" ^ Andreas@58039: " | format_term (SOME t) = t ();\n" ^ Andreas@58039: "fun format (true, _) = \"" ^ successN ^ "\\n\"\n" ^ Andreas@58039: " | format (false, to) = \"" ^ failureN ^ "\" ^ format_term to ^ \"\\n\";\n" ^ Andreas@58039: "val result = " ^ value_name ^ " ();\n" ^ Andreas@58039: "val _ = print \"" ^ start_markerN ^ "\";\n" ^ Andreas@58039: "val _ = map (print o format) result;\n" ^ Andreas@58039: "val _ = print \"" ^ end_markerN ^ "\";\n" Andreas@58039: val ml_basis = Andreas@58039: "$(SML_LIB)/basis/basis.mlb\n" ^ Andreas@58039: generatedN ^ "\n" ^ Andreas@58039: driverN Andreas@58039: Andreas@58039: val compile_cmd = Andreas@58039: File.shell_path (Path.variable ISABELLE_MLTON) ^ Andreas@58039: " -default-type intinf " ^ File.shell_path ml_basis_path Andreas@58039: val run_cmd = File.shell_path (Path.append path (Path.basic projectN)) Andreas@58039: in Andreas@58039: {files = [(driver_path, driver), (ml_basis_path, ml_basis)], Andreas@58039: compile_cmd = SOME compile_cmd, run_cmd = run_cmd, mk_code_file = K code_path} Andreas@58039: end Andreas@58039: Andreas@58039: val evaluate_in_mlton = gen_driver mk_driver_mlton ISABELLE_MLTON "MLton executable" mltonN Andreas@58039: Andreas@58039: (* Driver for SML/NJ *) Andreas@58039: Andreas@58039: val smlnjN = "SMLNJ" Andreas@58039: val ISABELLE_SMLNJ = "ISABELLE_SMLNJ" Andreas@58039: Andreas@58039: fun mk_driver_smlnj _ path _ value_name = Andreas@58039: let Andreas@58039: val generatedN = "generated.sml" Andreas@58039: val driverN = "driver.sml" Andreas@58039: Andreas@58039: val code_path = Path.append path (Path.basic generatedN) Andreas@58039: val driver_path = Path.append path (Path.basic driverN) Andreas@58039: val driver = Andreas@58039: "structure Test = struct\n" ^ Andreas@58039: "fun main prog_name =\n" ^ Andreas@58039: " let\n" ^ Andreas@58039: " fun format_term NONE = \"\"\n" ^ Andreas@58039: " | format_term (SOME t) = t ();\n" ^ Andreas@58039: " fun format (true, _) = \"" ^ successN ^ "\\n\"\n" ^ Andreas@58039: " | format (false, to) = \"" ^ failureN ^ "\" ^ format_term to ^ \"\\n\";\n" ^ Andreas@58039: " val result = " ^ value_name ^ " ();\n" ^ Andreas@58039: " val _ = print \"" ^ start_markerN ^ "\";\n" ^ Andreas@58039: " val _ = map (print o format) result;\n" ^ Andreas@58039: " val _ = print \"" ^ end_markerN ^ "\";\n" ^ Andreas@58039: " in\n" ^ Andreas@58039: " 0\n" ^ Andreas@58039: " end;\n" ^ Andreas@58039: "end;" Andreas@58039: val cmd = Andreas@58039: "echo \"Control.MC.matchRedundantError := false; Control.MC.matchRedundantWarn := false;" ^ Andreas@58039: "use \\\"" ^ Path.implode code_path ^ "\\\"; use \\\"" ^ Path.implode driver_path ^ "\\\";" ^ Andreas@58039: "Test.main ();\" | " ^ Path.implode (Path.variable ISABELLE_SMLNJ) Andreas@58039: in Andreas@58039: {files = [(driver_path, driver)], compile_cmd = NONE, run_cmd = cmd, mk_code_file = K code_path} Andreas@58039: end Andreas@58039: Andreas@58039: val evaluate_in_smlnj = gen_driver mk_driver_smlnj ISABELLE_SMLNJ "SMLNJ executable" smlnjN Andreas@58039: Andreas@58039: (* Driver for OCaml *) Andreas@58039: Andreas@58039: val ocamlN = "OCaml" Andreas@58039: val ISABELLE_OCAMLC = "ISABELLE_OCAMLC" Andreas@58039: Andreas@58039: fun mk_driver_ocaml _ path _ value_name = Andreas@58039: let Andreas@58039: val generatedN = "generated.ml" Andreas@58039: val driverN = "driver.ml" Andreas@58039: Andreas@58039: val code_path = Path.append path (Path.basic generatedN) Andreas@58039: val driver_path = Path.append path (Path.basic driverN) Andreas@58039: val driver = Andreas@58039: "let format_term = function\n" ^ Andreas@58039: " | None -> \"\"\n" ^ Andreas@58039: " | Some t -> t ();;\n" ^ Andreas@58039: "let format = function\n" ^ Andreas@58039: " | (true, _) -> \"" ^ successN ^ "\\n\"\n" ^ Andreas@58039: " | (false, x) -> \"" ^ failureN ^ "\" ^ format_term x ^ \"\\n\";;\n" ^ Andreas@58039: "let result = " ^ ("Generated." ^ value_name) ^ " ();;\n" ^ Andreas@58039: "let main x =\n" ^ Andreas@58039: " let _ = print_string \"" ^ start_markerN ^ "\" in\n" ^ Andreas@58039: " let _ = List.map (fun x -> print_string (format x)) result in\n" ^ Andreas@58039: " print_string \"" ^ end_markerN ^ "\";;\n" ^ Andreas@58039: "main ();;" Andreas@58039: Andreas@58039: val compiled_path = Path.append path (Path.basic "test") Andreas@58039: val compile_cmd = Andreas@58039: Path.implode (Path.variable ISABELLE_OCAMLC) ^ " -w pu -o " ^ Path.implode compiled_path ^ Andreas@58039: " -I " ^ Path.implode path ^ Andreas@58039: " nums.cma " ^ Path.implode code_path ^ " " ^ Path.implode driver_path Andreas@58039: Andreas@58039: val run_cmd = File.shell_path compiled_path Andreas@58039: in Andreas@58039: {files = [(driver_path, driver)], Andreas@58039: compile_cmd = SOME compile_cmd, run_cmd = run_cmd, mk_code_file = K code_path} Andreas@58039: end Andreas@58039: Andreas@58039: val evaluate_in_ocaml = gen_driver mk_driver_ocaml ISABELLE_OCAMLC "ocamlc executable" ocamlN Andreas@58039: Andreas@58039: (* Driver for GHC *) Andreas@58039: Andreas@58039: val ghcN = "GHC" Andreas@58039: val ISABELLE_GHC = "ISABELLE_GHC" Andreas@58039: Andreas@58039: val ghc_options = Attrib.setup_config_string @{binding code_test_ghc} (K "") Andreas@58039: Andreas@58039: fun mk_driver_ghc ctxt path modules value_name = Andreas@58039: let Andreas@58039: val driverN = "Main.hs" Andreas@58039: Andreas@58039: fun mk_code_file name = Path.append path (Path.basic (name ^ ".hs")) Andreas@58039: val driver_path = Path.append path (Path.basic driverN) Andreas@58039: val driver = Andreas@58039: "module Main where {\n" ^ Andreas@58039: String.concat (map (fn module => "import qualified " ^ module ^ ";\n") modules) ^ Andreas@58039: "main = do {\n" ^ Andreas@58039: " let {\n" ^ Andreas@58039: " format_term Nothing = \"\";\n" ^ Andreas@58039: " format_term (Just t) = t ();\n" ^ Andreas@58039: " format (True, _) = \"" ^ successN ^ "\\n\";\n" ^ Andreas@58039: " format (False, to) = \"" ^ failureN ^ "\" ++ format_term to ++ \"\\n\";\n" ^ Andreas@58039: " result = " ^ value_name ^ " ();\n" ^ Andreas@58039: " };\n" ^ Andreas@58039: " Prelude.putStr \"" ^ start_markerN ^ "\";\n" ^ Andreas@58039: " Prelude.mapM_ (putStr . format) result;\n" ^ Andreas@58039: " Prelude.putStr \"" ^ end_markerN ^ "\";\n" ^ Andreas@58039: " }\n" ^ Andreas@58039: "}\n" Andreas@58039: Andreas@58039: val compiled_path = Path.append path (Path.basic "test") Andreas@58039: val compile_cmd = Andreas@58039: Path.implode (Path.variable ISABELLE_GHC) ^ " " ^ Code_Haskell.language_params ^ " " ^ Andreas@58039: Config.get ctxt ghc_options ^ " -o " ^ Path.implode compiled_path ^ " " ^ Andreas@58039: Path.implode driver_path ^ " -i" ^ Path.implode path Andreas@58039: Andreas@58039: val run_cmd = File.shell_path compiled_path Andreas@58039: in Andreas@58039: {files = [(driver_path, driver)], Andreas@58039: compile_cmd = SOME compile_cmd, run_cmd = run_cmd, mk_code_file = mk_code_file} Andreas@58039: end Andreas@58039: Andreas@58039: val evaluate_in_ghc = gen_driver mk_driver_ghc ISABELLE_GHC "GHC executable" ghcN Andreas@58039: Andreas@58039: (* Driver for Scala *) Andreas@58039: Andreas@58039: val scalaN = "Scala" Andreas@58039: val ISABELLE_SCALA = "ISABELLE_SCALA" Andreas@58039: Andreas@58039: fun mk_driver_scala _ path _ value_name = Andreas@58039: let Andreas@58039: val generatedN = "Generated_Code" Andreas@58039: val driverN = "Driver.scala" Andreas@58039: Andreas@58039: val code_path = Path.append path (Path.basic (generatedN ^ ".scala")) Andreas@58039: val driver_path = Path.append path (Path.basic driverN) Andreas@58039: val driver = Andreas@58039: "import " ^ generatedN ^ "._\n" ^ Andreas@58039: "object Test {\n" ^ Andreas@58039: " def format_term(x : Option[Unit => String]) : String = x match {\n" ^ Andreas@58039: " case None => \"\"\n" ^ Andreas@58039: " case Some(x) => x(())\n" ^ Andreas@58039: " }\n" ^ Andreas@58039: " def format(term : (Boolean, Option[Unit => String])) : String = term match {\n" ^ Andreas@58039: " case (true, _) => \"True\\n\"\n" ^ Andreas@58039: " case (false, x) => \"False\" + format_term(x) + \"\\n\"\n" ^ Andreas@58039: " }\n" ^ Andreas@58039: " def main(args:Array[String]) = {\n" ^ Andreas@58039: " val result = " ^ value_name ^ "(());\n" ^ Andreas@58039: " print(\"" ^ start_markerN ^ "\");\n" ^ Andreas@58039: " result.map{test:(Boolean, Option[Unit => String]) => print(format(test))};\n" ^ Andreas@58039: " print(\"" ^ end_markerN ^ "\");\n" ^ Andreas@58039: " }\n" ^ Andreas@58039: "}\n" Andreas@58039: Andreas@58039: val compile_cmd = Andreas@58039: Path.implode (Path.append (Path.variable ISABELLE_SCALA) (Path.basic "scalac")) ^ Andreas@58039: " -d " ^ File.shell_path path ^ " -classpath " ^ File.shell_path path ^ " " ^ Andreas@58039: File.shell_path code_path ^ " " ^ File.shell_path driver_path Andreas@58039: Andreas@58039: val run_cmd = Andreas@58039: Path.implode (Path.append (Path.variable ISABELLE_SCALA) (Path.basic "scala")) ^ Andreas@58039: " -cp " ^ File.shell_path path ^ " Test" Andreas@58039: in Andreas@58039: {files = [(driver_path, driver)], Andreas@58039: compile_cmd = SOME compile_cmd, run_cmd = run_cmd, mk_code_file = K code_path} Andreas@58039: end Andreas@58039: Andreas@58039: val evaluate_in_scala = gen_driver mk_driver_scala ISABELLE_SCALA "Scala directory" scalaN Andreas@58039: Andreas@58039: val test_codeP = Scan.repeat1 Parse.prop -- (@{keyword "in"} |-- Scan.repeat1 Parse.name) Andreas@58039: Andreas@58039: val _ = Andreas@58039: Outer_Syntax.command @{command_spec "test_code"} Andreas@58039: "compile test cases to target languages, execute them and report results" Andreas@58039: (test_codeP >> (fn (raw_ts, targets) => Toplevel.keep (test_code_cmd raw_ts targets))) Andreas@58039: Andreas@58039: val _ = Context.>> (Context.map_theory ( Andreas@58039: fold add_driver Andreas@58039: [(polymlN, (evaluate_in_polyml, Code_ML.target_SML)), Andreas@58039: (mltonN, (evaluate_in_mlton, Code_ML.target_SML)), Andreas@58039: (smlnjN, (evaluate_in_smlnj, Code_ML.target_SML)), Andreas@58039: (ocamlN, (evaluate_in_ocaml, Code_ML.target_OCaml)), Andreas@58039: (ghcN, (evaluate_in_ghc, Code_Haskell.target)), Andreas@58348: (scalaN, (evaluate_in_scala, Code_Scala.target))] Andreas@58348: #> fold (fn target => Value.add_evaluator (target, eval_term target)) Andreas@58348: [polymlN, mltonN, smlnjN, ocamlN, ghcN, scalaN] Andreas@58348: )) Andreas@58039: end Andreas@58039: