| author | wenzelm | 
| Wed, 11 Nov 2020 20:55:25 +0100 | |
| changeset 72572 | e7e93c0f6d96 | 
| parent 72511 | 460d743010bc | 
| child 74383 | 107941e8fa01 | 
| 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  | 
|
| 72284 | 12  | 
val debug: bool Config.T  | 
| 64577 | 13  | 
val test_terms: Proof.context -> term list -> string -> unit  | 
| 64580 | 14  | 
val test_code_cmd: string list -> string list -> Proof.context -> unit  | 
| 64577 | 15  | 
val eval_term: string -> Proof.context -> term -> term  | 
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
16  | 
val check_settings: string -> string -> string -> unit  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
17  | 
val compile: string -> string -> unit  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
18  | 
val evaluate: string -> string -> string  | 
| 64577 | 19  | 
val evaluate_in_polyml: Proof.context -> (string * string) list * string -> Path.T -> string  | 
20  | 
val evaluate_in_mlton: Proof.context -> (string * string) list * string -> Path.T -> string  | 
|
21  | 
val evaluate_in_smlnj: Proof.context -> (string * string) list * string -> Path.T -> string  | 
|
22  | 
val evaluate_in_ocaml: Proof.context -> (string * string) list * string -> Path.T -> string  | 
|
23  | 
val ghc_options: string Config.T  | 
|
24  | 
val evaluate_in_ghc: Proof.context -> (string * string) list * string -> Path.T -> string  | 
|
25  | 
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: 
65905 
diff
changeset
 | 
26  | 
val target_Scala: string  | 
| 
 
378895354604
new derived targets for evaluating Haskell and Scala programs
 
Andreas Lochbihler 
parents: 
65905 
diff
changeset
 | 
27  | 
val target_Haskell: string  | 
| 58039 | 28  | 
end  | 
29  | 
||
| 64577 | 30  | 
structure Code_Test: CODE_TEST =  | 
31  | 
struct  | 
|
| 58039 | 32  | 
|
33  | 
(* convert a list of terms into nested tuples and back *)  | 
|
| 64577 | 34  | 
|
| 69593 | 35  | 
fun mk_tuples [] = \<^term>\<open>()\<close>  | 
| 58039 | 36  | 
| mk_tuples [t] = t  | 
37  | 
| mk_tuples (t :: ts) = HOLogic.mk_prod (t, mk_tuples ts)  | 
|
38  | 
||
| 69593 | 39  | 
fun dest_tuples (Const (\<^const_name>\<open>Pair\<close>, _) $ l $ r) = l :: dest_tuples r  | 
| 58039 | 40  | 
| dest_tuples t = [t]  | 
41  | 
||
42  | 
||
43  | 
fun last_field sep str =  | 
|
44  | 
let  | 
|
| 64577 | 45  | 
val n = size sep  | 
46  | 
val len = size str  | 
|
| 58039 | 47  | 
fun find i =  | 
48  | 
if i < 0 then NONE  | 
|
49  | 
else if String.substring (str, i, n) = sep then SOME i  | 
|
| 64577 | 50  | 
else find (i - 1)  | 
| 58039 | 51  | 
in  | 
52  | 
(case find (len - n) of  | 
|
53  | 
NONE => NONE  | 
|
54  | 
| SOME i => SOME (String.substring (str, 0, i), String.extract (str, i + n, NONE)))  | 
|
| 64577 | 55  | 
end  | 
| 58039 | 56  | 
|
57  | 
fun split_first_last start stop s =  | 
|
| 64577 | 58  | 
(case first_field start s of  | 
59  | 
NONE => NONE  | 
|
60  | 
| SOME (initial, rest) =>  | 
|
61  | 
(case last_field stop rest of  | 
|
62  | 
NONE => NONE  | 
|
63  | 
| SOME (middle, tail) => SOME (initial, middle, tail)))  | 
|
| 58039 | 64  | 
|
| 64577 | 65  | 
|
66  | 
(* data slot for drivers *)  | 
|
| 58039 | 67  | 
|
68  | 
structure Drivers = Theory_Data  | 
|
69  | 
(  | 
|
| 64577 | 70  | 
type T =  | 
71  | 
(string * ((Proof.context -> (string * string) list * string -> Path.T -> string) * string)) list  | 
|
72  | 
val empty = []  | 
|
73  | 
val extend = I  | 
|
74  | 
fun merge data : T = AList.merge (op =) (K true) data  | 
|
| 58039 | 75  | 
)  | 
76  | 
||
| 64577 | 77  | 
val add_driver = Drivers.map o AList.update (op =)  | 
78  | 
val get_driver = AList.lookup (op =) o Drivers.get  | 
|
| 58039 | 79  | 
|
80  | 
(*  | 
|
81  | 
Test drivers must produce output of the following format:  | 
|
| 64577 | 82  | 
|
| 72307 | 83  | 
The start of the relevant data is marked with start_marker,  | 
84  | 
its end with end_marker.  | 
|
| 58039 | 85  | 
|
86  | 
Between these two markers, every line corresponds to one test.  | 
|
| 72307 | 87  | 
Lines of successful tests start with success, failures start with failure.  | 
88  | 
The failure failure may continue with the YXML encoding of the evaluated term.  | 
|
| 58039 | 89  | 
There must not be any additional whitespace in between.  | 
90  | 
*)  | 
|
91  | 
||
| 64577 | 92  | 
|
93  | 
(* parsing of results *)  | 
|
| 58039 | 94  | 
|
| 72307 | 95  | 
val success = "True"  | 
96  | 
val failure = "False"  | 
|
97  | 
val start_marker = "*@*Isabelle/Code_Test-start*@*"  | 
|
98  | 
val end_marker = "*@*Isabelle/Code_Test-end*@*"  | 
|
| 58039 | 99  | 
|
100  | 
fun parse_line line =  | 
|
| 72307 | 101  | 
if String.isPrefix success line then (true, NONE)  | 
102  | 
else if String.isPrefix failure line then (false,  | 
|
103  | 
if size line > size failure then  | 
|
104  | 
String.extract (line, size failure, NONE)  | 
|
| 58039 | 105  | 
|> YXML.parse_body  | 
| 
70784
 
799437173553
Term_XML.Encode/Decode.term uses Const "typargs";
 
wenzelm 
parents: 
69950 
diff
changeset
 | 
106  | 
|> Term_XML.Decode.term_raw  | 
| 58039 | 107  | 
|> dest_tuples  | 
108  | 
|> SOME  | 
|
109  | 
else NONE)  | 
|
110  | 
  else raise Fail ("Cannot parse result of evaluation:\n" ^ line)
 | 
|
111  | 
||
112  | 
fun parse_result target out =  | 
|
| 72307 | 113  | 
(case split_first_last start_marker end_marker out of  | 
| 72273 | 114  | 
    NONE => error ("Evaluation failed for " ^ target ^ "!\nCompiler output:\n" ^ out)
 | 
| 
65904
 
8411f1a2272c
permissive trim_line as in Scala, e.g. relevant for poly/TextIO.print output on Windows;
 
wenzelm 
parents: 
65901 
diff
changeset
 | 
115  | 
| SOME (_, middle, _) => middle |> trim_split_lines |> map parse_line)  | 
| 58039 | 116  | 
|
| 64577 | 117  | 
|
118  | 
(* pretty printing of test results *)  | 
|
| 58039 | 119  | 
|
120  | 
fun pretty_eval _ NONE _ = []  | 
|
| 64577 | 121  | 
| pretty_eval ctxt (SOME evals) ts =  | 
122  | 
[Pretty.fbrk,  | 
|
123  | 
Pretty.big_list "Evaluated terms"  | 
|
124  | 
(map (fn (t, eval) => Pretty.block  | 
|
125  | 
[Syntax.pretty_term ctxt t, Pretty.brk 1, Pretty.str "=", Pretty.brk 1,  | 
|
126  | 
Syntax.pretty_term ctxt eval])  | 
|
127  | 
(ts ~~ evals))]  | 
|
| 58039 | 128  | 
|
129  | 
fun pretty_failure ctxt target (((_, evals), query), eval_ts) =  | 
|
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
130  | 
  Pretty.block (Pretty.text ("Test in " ^ target ^ " failed for") @
 | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
131  | 
[Pretty.brk 1, Pretty.quote (Syntax.pretty_term ctxt query)] @  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
132  | 
pretty_eval ctxt evals eval_ts)  | 
| 58039 | 133  | 
|
134  | 
fun pretty_failures ctxt target failures =  | 
|
135  | 
Pretty.blk (0, Pretty.fbreaks (map (pretty_failure ctxt target) failures))  | 
|
136  | 
||
137  | 
||
| 64577 | 138  | 
(* driver invocation *)  | 
139  | 
||
| 72293 | 140  | 
val debug = Attrib.setup_config_bool \<^binding>\<open>test_code_debug\<close> (K false)  | 
| 58039 | 141  | 
|
| 72284 | 142  | 
fun with_debug_dir name f =  | 
| 
72511
 
460d743010bc
clarified signature: overloaded "+" for Path.append;
 
wenzelm 
parents: 
72376 
diff
changeset
 | 
143  | 
(Path.explode "$ISABELLE_HOME_USER" + Path.basic (name ^ serial_string ()))  | 
| 72376 | 144  | 
|> Isabelle_System.make_directory  | 
145  | 
|> Exn.capture f  | 
|
146  | 
|> Exn.release;  | 
|
| 58039 | 147  | 
|
148  | 
fun dynamic_value_strict ctxt t compiler =  | 
|
149  | 
let  | 
|
150  | 
val thy = Proof_Context.theory_of ctxt  | 
|
| 64577 | 151  | 
val (driver, target) =  | 
152  | 
(case get_driver thy compiler of  | 
|
153  | 
        NONE => error ("No driver for target " ^ compiler)
 | 
|
| 72272 | 154  | 
| SOME drv => drv)  | 
| 72285 | 155  | 
val with_dir = if Config.get ctxt debug then with_debug_dir else Isabelle_System.with_tmp_dir  | 
| 72274 | 156  | 
fun eval result =  | 
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69597 
diff
changeset
 | 
157  | 
with_dir "Code_Test" (driver ctxt ((apfst o map o apfst) Long_Name.implode result))  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69597 
diff
changeset
 | 
158  | 
|> parse_result compiler  | 
| 58039 | 159  | 
fun evaluator program _ vs_ty deps =  | 
| 72274 | 160  | 
Exn.interruptible_capture eval  | 
| 64957 | 161  | 
(Code_Target.compilation_text ctxt target program deps true vs_ty)  | 
| 64578 | 162  | 
fun postproc f = map (apsnd (Option.map (map f)))  | 
| 64577 | 163  | 
in Exn.release (Code_Thingol.dynamic_value ctxt (Exn.map_res o postproc) evaluator t) end  | 
| 58039 | 164  | 
|
| 64577 | 165  | 
|
166  | 
(* term preprocessing *)  | 
|
| 58039 | 167  | 
|
| 69593 | 168  | 
fun add_eval (Const (\<^const_name>\<open>Trueprop\<close>, _) $ t) = add_eval t  | 
169  | 
| add_eval (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ lhs $ rhs) = (fn acc =>  | 
|
| 64577 | 170  | 
acc  | 
171  | 
|> add_eval rhs  | 
|
172  | 
|> add_eval lhs  | 
|
173  | 
|> cons rhs  | 
|
174  | 
|> cons lhs)  | 
|
| 69593 | 175  | 
| add_eval (Const (\<^const_name>\<open>Not\<close>, _) $ t) = add_eval t  | 
176  | 
| add_eval (Const (\<^const_name>\<open>Orderings.ord_class.less_eq\<close>, _) $ lhs $ rhs) = (fn acc =>  | 
|
| 64577 | 177  | 
lhs :: rhs :: acc)  | 
| 69593 | 178  | 
| add_eval (Const (\<^const_name>\<open>Orderings.ord_class.less\<close>, _) $ lhs $ rhs) = (fn acc =>  | 
| 64577 | 179  | 
lhs :: rhs :: acc)  | 
| 58039 | 180  | 
| add_eval _ = I  | 
181  | 
||
| 69593 | 182  | 
fun mk_term_of [] = \<^term>\<open>None :: (unit \<Rightarrow> yxml_of_term) option\<close>  | 
| 58039 | 183  | 
| mk_term_of ts =  | 
| 64577 | 184  | 
let  | 
185  | 
val tuple = mk_tuples ts  | 
|
186  | 
val T = fastype_of tuple  | 
|
187  | 
in  | 
|
| 69593 | 188  | 
\<^term>\<open>Some :: (unit \<Rightarrow> yxml_of_term) \<Rightarrow> (unit \<Rightarrow> yxml_of_term) option\<close> $  | 
| 69597 | 189  | 
(absdummy \<^typ>\<open>unit\<close> (\<^const>\<open>yxml_string_of_term\<close> $  | 
| 69593 | 190  | 
(Const (\<^const_name>\<open>Code_Evaluation.term_of\<close>, T --> \<^typ>\<open>term\<close>) $ tuple)))  | 
| 64577 | 191  | 
end  | 
| 58039 | 192  | 
|
193  | 
fun test_terms ctxt ts target =  | 
|
194  | 
let  | 
|
195  | 
val thy = Proof_Context.theory_of ctxt  | 
|
196  | 
||
| 69593 | 197  | 
fun term_of t = Sign.of_sort thy (fastype_of t, \<^sort>\<open>term_of\<close>)  | 
| 58039 | 198  | 
|
| 64577 | 199  | 
fun ensure_bool t =  | 
200  | 
(case fastype_of t of  | 
|
| 69593 | 201  | 
\<^typ>\<open>bool\<close> => ()  | 
| 64579 | 202  | 
| _ =>  | 
203  | 
error (Pretty.string_of  | 
|
204  | 
(Pretty.block [Pretty.str "Test case not of type bool:", Pretty.brk 1,  | 
|
205  | 
Syntax.pretty_term ctxt t])))  | 
|
| 58039 | 206  | 
|
| 64578 | 207  | 
val _ = List.app ensure_bool ts  | 
| 58039 | 208  | 
|
209  | 
val evals = map (fn t => filter term_of (add_eval t [])) ts  | 
|
210  | 
val eval = map mk_term_of evals  | 
|
211  | 
||
| 64578 | 212  | 
val t =  | 
| 69593 | 213  | 
HOLogic.mk_list \<^typ>\<open>bool \<times> (unit \<Rightarrow> yxml_of_term) option\<close>  | 
| 64578 | 214  | 
(map HOLogic.mk_prod (ts ~~ eval))  | 
| 58039 | 215  | 
|
| 64577 | 216  | 
val result = dynamic_value_strict ctxt t target  | 
| 58039 | 217  | 
|
218  | 
val failed =  | 
|
219  | 
filter_out (fst o fst o fst) (result ~~ ts ~~ evals)  | 
|
| 64577 | 220  | 
handle ListPair.UnequalLengths =>  | 
| 72272 | 221  | 
        error ("Evaluation failed!\nWrong number of test results: " ^ string_of_int (length result))
 | 
| 58039 | 222  | 
in  | 
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
223  | 
(case failed of  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
224  | 
[] => ()  | 
| 64578 | 225  | 
| _ => error (Pretty.string_of (pretty_failures ctxt target failed)))  | 
| 58039 | 226  | 
end  | 
227  | 
||
| 64580 | 228  | 
fun test_code_cmd raw_ts targets ctxt =  | 
| 58039 | 229  | 
let  | 
| 64577 | 230  | 
val ts = Syntax.read_terms ctxt raw_ts  | 
| 64579 | 231  | 
val frees = fold Term.add_frees ts []  | 
| 64578 | 232  | 
val _ =  | 
233  | 
if null frees then ()  | 
|
| 64579 | 234  | 
else error (Pretty.string_of  | 
235  | 
(Pretty.block (Pretty.str "Terms contain free variables:" :: Pretty.brk 1 ::  | 
|
| 72283 | 236  | 
Pretty.commas (map (Syntax.pretty_term ctxt o Free) frees))))  | 
237  | 
in List.app (test_terms ctxt ts) targets end  | 
|
| 58039 | 238  | 
|
| 
58348
 
2d47c7d10b62
add target language evaluators for the value command;
 
Andreas Lochbihler 
parents: 
58039 
diff
changeset
 | 
239  | 
fun eval_term target ctxt t =  | 
| 
 
2d47c7d10b62
add target language evaluators for the value command;
 
Andreas Lochbihler 
parents: 
58039 
diff
changeset
 | 
240  | 
let  | 
| 64579 | 241  | 
val frees = Term.add_frees t []  | 
| 64578 | 242  | 
val _ =  | 
243  | 
if null frees then ()  | 
|
| 64579 | 244  | 
else  | 
245  | 
error (Pretty.string_of  | 
|
246  | 
(Pretty.block (Pretty.str "Term contains free variables:" :: Pretty.brk 1 ::  | 
|
| 72283 | 247  | 
Pretty.commas (map (Syntax.pretty_term ctxt o Free) frees))))  | 
| 58039 | 248  | 
|
| 64579 | 249  | 
val T = fastype_of t  | 
| 64578 | 250  | 
val _ =  | 
| 69593 | 251  | 
if Sign.of_sort (Proof_Context.theory_of ctxt) (T, \<^sort>\<open>term_of\<close>) then ()  | 
| 64579 | 252  | 
      else error ("Type " ^ Syntax.string_of_typ ctxt T ^
 | 
| 69593 | 253  | 
" of term not of sort " ^ Syntax.string_of_sort ctxt \<^sort>\<open>term_of\<close>)  | 
| 58039 | 254  | 
|
| 64578 | 255  | 
val t' =  | 
| 69593 | 256  | 
HOLogic.mk_list \<^typ>\<open>bool \<times> (unit \<Rightarrow> yxml_of_term) option\<close>  | 
257  | 
[HOLogic.mk_prod (\<^term>\<open>False\<close>, mk_term_of [t])]  | 
|
| 58039 | 258  | 
|
| 64577 | 259  | 
val result = dynamic_value_strict ctxt t' target  | 
260  | 
in (case result of [(_, SOME [t])] => t | _ => error "Evaluation failed") end  | 
|
| 58039 | 261  | 
|
| 64577 | 262  | 
|
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
263  | 
(* check and invoke compiler *)  | 
| 58039 | 264  | 
|
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
265  | 
fun check_settings compiler var descr =  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
266  | 
if getenv var = "" then  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
267  | 
error (Pretty.string_of (Pretty.para  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
268  | 
      ("Environment variable " ^ var ^ " is not set. To test code generation with " ^
 | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
269  | 
compiler ^ ", set this variable to your " ^ descr ^  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
270  | 
" in the $ISABELLE_HOME_USER/etc/settings file.")))  | 
| 72306 | 271  | 
else ()  | 
| 58039 | 272  | 
|
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
273  | 
fun compile compiler cmd =  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
274  | 
let val (out, ret) = Isabelle_System.bash_output cmd in  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
275  | 
if ret = 0 then ()  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
276  | 
    else error ("Compilation with " ^ compiler ^ " failed:\n" ^ cmd ^ "\n" ^ out)
 | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
277  | 
end  | 
| 58039 | 278  | 
|
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
279  | 
fun evaluate compiler cmd =  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
280  | 
let val (out, res) = Isabelle_System.bash_output cmd in  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
281  | 
if res = 0 then out  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
282  | 
    else error ("Evaluation for " ^ compiler ^ " terminated with error code " ^
 | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
283  | 
string_of_int res ^ "\nCompiler output:\n" ^ out)  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
284  | 
end  | 
| 58039 | 285  | 
|
| 64577 | 286  | 
|
287  | 
(* driver for PolyML *)  | 
|
| 58039 | 288  | 
|
| 64577 | 289  | 
val polymlN = "PolyML"  | 
| 58039 | 290  | 
|
| 72297 | 291  | 
fun evaluate_in_polyml (_: Proof.context) (code_files, value_name) dir =  | 
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
292  | 
let  | 
| 
72511
 
460d743010bc
clarified signature: overloaded "+" for Path.append;
 
wenzelm 
parents: 
72376 
diff
changeset
 | 
293  | 
val code_path = dir + Path.basic "generated.sml"  | 
| 
 
460d743010bc
clarified signature: overloaded "+" for Path.append;
 
wenzelm 
parents: 
72376 
diff
changeset
 | 
294  | 
val driver_path = dir + Path.basic "driver.sml"  | 
| 
 
460d743010bc
clarified signature: overloaded "+" for Path.append;
 
wenzelm 
parents: 
72376 
diff
changeset
 | 
295  | 
val out_path = dir + Path.basic "out"  | 
| 72297 | 296  | 
|
| 72306 | 297  | 
val code = #2 (the_single code_files)  | 
| 72286 | 298  | 
val driver = \<^verbatim>\<open>  | 
| 72290 | 299  | 
fun main () =  | 
| 72286 | 300  | 
let  | 
| 72307 | 301  | 
fun format (true, _) = \<close> ^ ML_Syntax.print_string success ^ \<^verbatim>\<open> ^ "\n"  | 
302  | 
| format (false, NONE) = \<close> ^ ML_Syntax.print_string failure ^ \<^verbatim>\<open> ^ "\n"  | 
|
303  | 
| format (false, SOME t) = \<close> ^ ML_Syntax.print_string failure ^ \<^verbatim>\<open> ^ t () ^ "\n"  | 
|
| 72286 | 304  | 
val result = \<close> ^ value_name ^ \<^verbatim>\<open> ()  | 
305  | 
val result_text = \<close> ^  | 
|
| 72307 | 306  | 
ML_Syntax.print_string start_marker ^  | 
| 72286 | 307  | 
\<^verbatim>\<open> ^ String.concat (map format result) ^ \<close> ^  | 
| 72307 | 308  | 
ML_Syntax.print_string end_marker ^ \<^verbatim>\<open>  | 
| 72288 | 309  | 
val out = BinIO.openOut \<close> ^ ML_Syntax.print_platform_path out_path ^ \<^verbatim>\<open>  | 
| 72286 | 310  | 
val _ = BinIO.output (out, Byte.stringToBytes result_text)  | 
311  | 
val _ = BinIO.closeOut out  | 
|
312  | 
in () end;  | 
|
313  | 
\<close>  | 
|
| 72306 | 314  | 
val _ = File.write code_path code  | 
315  | 
val _ = File.write driver_path driver  | 
|
316  | 
val _ =  | 
|
317  | 
ML_Context.eval  | 
|
318  | 
        {environment = ML_Env.SML, redirect = false, verbose = false, debug = NONE,
 | 
|
319  | 
writeln = writeln, warning = warning}  | 
|
320  | 
Position.none  | 
|
321  | 
(ML_Lex.read_text (code, Path.position code_path) @  | 
|
322  | 
ML_Lex.read_text (driver, Path.position driver_path) @  | 
|
323  | 
ML_Lex.read "main ()")  | 
|
| 72308 | 324  | 
      handle ERROR msg => error ("Evaluation for " ^ polymlN ^ " failed:\n" ^ msg)
 | 
| 72306 | 325  | 
in File.read out_path end  | 
| 58039 | 326  | 
|
| 64577 | 327  | 
|
328  | 
(* driver for mlton *)  | 
|
| 58039 | 329  | 
|
330  | 
val mltonN = "MLton"  | 
|
331  | 
val ISABELLE_MLTON = "ISABELLE_MLTON"  | 
|
332  | 
||
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
333  | 
fun evaluate_in_mlton (_: Proof.context) (code_files, value_name) dir =  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
334  | 
let  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
335  | 
val compiler = mltonN  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
336  | 
val generatedN = "generated.sml"  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
337  | 
val driverN = "driver.sml"  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
338  | 
val projectN = "test"  | 
| 58039 | 339  | 
|
| 
72511
 
460d743010bc
clarified signature: overloaded "+" for Path.append;
 
wenzelm 
parents: 
72376 
diff
changeset
 | 
340  | 
val code_path = dir + Path.basic generatedN  | 
| 
 
460d743010bc
clarified signature: overloaded "+" for Path.append;
 
wenzelm 
parents: 
72376 
diff
changeset
 | 
341  | 
val driver_path = dir + Path.basic driverN  | 
| 
 
460d743010bc
clarified signature: overloaded "+" for Path.append;
 
wenzelm 
parents: 
72376 
diff
changeset
 | 
342  | 
val basis_path = dir + Path.basic (projectN ^ ".mlb")  | 
| 72291 | 343  | 
val driver = \<^verbatim>\<open>  | 
| 72307 | 344  | 
fun format (true, _) = \<close> ^ ML_Syntax.print_string success ^ \<^verbatim>\<open> ^ "\n"  | 
345  | 
| format (false, NONE) = \<close> ^ ML_Syntax.print_string failure ^ \<^verbatim>\<open> ^ "\n"  | 
|
346  | 
| format (false, SOME t) = \<close> ^ ML_Syntax.print_string failure ^ \<^verbatim>\<open> ^ t () ^ "\n"  | 
|
| 72291 | 347  | 
val result = \<close> ^ value_name ^ \<^verbatim>\<open> ()  | 
| 72307 | 348  | 
val _ = print \<close> ^ ML_Syntax.print_string start_marker ^ \<^verbatim>\<open>  | 
| 72291 | 349  | 
val _ = List.app (print o format) result  | 
| 72307 | 350  | 
val _ = print \<close> ^ ML_Syntax.print_string end_marker ^ \<^verbatim>\<open>  | 
| 72291 | 351  | 
\<close>  | 
| 72306 | 352  | 
val _ = check_settings compiler ISABELLE_MLTON "MLton executable"  | 
353  | 
val _ = List.app (File.write code_path o snd) code_files  | 
|
354  | 
val _ = File.write driver_path driver  | 
|
355  | 
    val _ = File.write basis_path ("$(SML_LIB)/basis/basis.mlb\n" ^ generatedN ^ "\n" ^ driverN)
 | 
|
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
356  | 
in  | 
| 72306 | 357  | 
    compile compiler ("\"$ISABELLE_MLTON\" -default-type intinf " ^ File.bash_path basis_path);
 | 
| 
72511
 
460d743010bc
clarified signature: overloaded "+" for Path.append;
 
wenzelm 
parents: 
72376 
diff
changeset
 | 
358  | 
evaluate compiler (File.bash_path (dir + Path.basic projectN))  | 
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
359  | 
end  | 
| 58039 | 360  | 
|
| 64577 | 361  | 
|
362  | 
(* driver for SML/NJ *)  | 
|
| 58039 | 363  | 
|
364  | 
val smlnjN = "SMLNJ"  | 
|
365  | 
val ISABELLE_SMLNJ = "ISABELLE_SMLNJ"  | 
|
366  | 
||
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
367  | 
fun evaluate_in_smlnj (_: Proof.context) (code_files, value_name) dir =  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
368  | 
let  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
369  | 
val compiler = smlnjN  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
370  | 
val generatedN = "generated.sml"  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
371  | 
val driverN = "driver.sml"  | 
| 58039 | 372  | 
|
| 
72511
 
460d743010bc
clarified signature: overloaded "+" for Path.append;
 
wenzelm 
parents: 
72376 
diff
changeset
 | 
373  | 
val code_path = dir + Path.basic generatedN  | 
| 
 
460d743010bc
clarified signature: overloaded "+" for Path.append;
 
wenzelm 
parents: 
72376 
diff
changeset
 | 
374  | 
val driver_path = dir + Path.basic driverN  | 
| 72291 | 375  | 
val driver = \<^verbatim>\<open>  | 
376  | 
structure Test =  | 
|
377  | 
struct  | 
|
378  | 
fun main () =  | 
|
379  | 
let  | 
|
| 72307 | 380  | 
fun format (true, _) = \<close> ^ ML_Syntax.print_string success ^ \<^verbatim>\<open> ^ "\n"  | 
381  | 
| format (false, NONE) = \<close> ^ ML_Syntax.print_string failure ^ \<^verbatim>\<open> ^ "\n"  | 
|
382  | 
| format (false, SOME t) = \<close> ^ ML_Syntax.print_string failure ^ \<^verbatim>\<open> ^ t () ^ "\n"  | 
|
| 72291 | 383  | 
val result = \<close> ^ value_name ^ \<^verbatim>\<open> ()  | 
| 72307 | 384  | 
val _ = print \<close> ^ ML_Syntax.print_string start_marker ^ \<^verbatim>\<open>  | 
| 72291 | 385  | 
val _ = List.app (print o format) result  | 
| 72307 | 386  | 
val _ = print \<close> ^ ML_Syntax.print_string end_marker ^ \<^verbatim>\<open>  | 
| 72291 | 387  | 
in 0 end  | 
388  | 
end  | 
|
389  | 
\<close>  | 
|
| 72306 | 390  | 
val _ = check_settings compiler ISABELLE_SMLNJ "SMLNJ executable"  | 
391  | 
val _ = List.app (File.write code_path o snd) code_files  | 
|
392  | 
val _ = File.write driver_path driver  | 
|
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
393  | 
val ml_source =  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
394  | 
"Control.MC.matchRedundantError := false; Control.MC.matchRedundantWarn := false;" ^  | 
| 72277 | 395  | 
"use " ^ ML_Syntax.print_string (File.platform_path code_path) ^  | 
396  | 
"; use " ^ ML_Syntax.print_string (File.platform_path driver_path) ^  | 
|
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
397  | 
"; Test.main ();"  | 
| 72306 | 398  | 
  in evaluate compiler ("echo " ^ Bash.string ml_source ^ " | \"$ISABELLE_SMLNJ\"") end
 | 
| 58039 | 399  | 
|
| 64577 | 400  | 
|
401  | 
(* driver for OCaml *)  | 
|
| 58039 | 402  | 
|
403  | 
val ocamlN = "OCaml"  | 
|
| 
69926
 
110fff287217
access OCaml tools and libraries via ISABELLE_OCAMLFIND;
 
wenzelm 
parents: 
69925 
diff
changeset
 | 
404  | 
val ISABELLE_OCAMLFIND = "ISABELLE_OCAMLFIND"  | 
| 58039 | 405  | 
|
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
406  | 
fun evaluate_in_ocaml (_: Proof.context) (code_files, value_name) dir =  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
407  | 
let  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
408  | 
val compiler = ocamlN  | 
| 58039 | 409  | 
|
| 
72511
 
460d743010bc
clarified signature: overloaded "+" for Path.append;
 
wenzelm 
parents: 
72376 
diff
changeset
 | 
410  | 
val code_path = dir + Path.basic "generated.ml"  | 
| 
 
460d743010bc
clarified signature: overloaded "+" for Path.append;
 
wenzelm 
parents: 
72376 
diff
changeset
 | 
411  | 
val driver_path = dir + Path.basic "driver.ml"  | 
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
412  | 
val driver =  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
413  | 
"let format_term = function\n" ^  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
414  | 
" | None -> \"\"\n" ^  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
415  | 
" | Some t -> t ();;\n" ^  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
416  | 
"let format = function\n" ^  | 
| 72307 | 417  | 
" | (true, _) -> \"" ^ success ^ "\\n\"\n" ^  | 
418  | 
" | (false, x) -> \"" ^ failure ^ "\" ^ format_term x ^ \"\\n\";;\n" ^  | 
|
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
419  | 
      "let result = " ^ ("Generated." ^ value_name) ^ " ();;\n" ^
 | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
420  | 
"let main x =\n" ^  | 
| 72307 | 421  | 
" let _ = print_string \"" ^ start_marker ^ "\" in\n" ^  | 
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
422  | 
" let _ = List.map (fun x -> print_string (format x)) result in\n" ^  | 
| 72307 | 423  | 
" print_string \"" ^ end_marker ^ "\";;\n" ^  | 
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
424  | 
"main ();;"  | 
| 
72511
 
460d743010bc
clarified signature: overloaded "+" for Path.append;
 
wenzelm 
parents: 
72376 
diff
changeset
 | 
425  | 
val compiled_path = dir + Path.basic "test"  | 
| 72306 | 426  | 
|
427  | 
val _ = check_settings compiler ISABELLE_OCAMLFIND "ocamlfind executable"  | 
|
428  | 
val _ = List.app (File.write code_path o snd) code_files  | 
|
429  | 
val _ = File.write driver_path driver  | 
|
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
430  | 
val cmd =  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
431  | 
"\"$ISABELLE_OCAMLFIND\" ocamlopt -w pu -package zarith -linkpkg" ^  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
432  | 
" -o " ^ File.bash_path compiled_path ^ " -I " ^ File.bash_path dir ^ " " ^  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
433  | 
File.bash_path code_path ^ " " ^ File.bash_path driver_path ^ " </dev/null"  | 
| 72306 | 434  | 
in compile compiler cmd; evaluate compiler (File.bash_path compiled_path) end  | 
| 58039 | 435  | 
|
| 64577 | 436  | 
|
437  | 
(* driver for GHC *)  | 
|
| 58039 | 438  | 
|
439  | 
val ghcN = "GHC"  | 
|
440  | 
val ISABELLE_GHC = "ISABELLE_GHC"  | 
|
441  | 
||
| 69593 | 442  | 
val ghc_options = Attrib.setup_config_string \<^binding>\<open>code_test_ghc\<close> (K "")  | 
| 58039 | 443  | 
|
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
444  | 
fun evaluate_in_ghc ctxt (code_files, value_name) dir =  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
445  | 
let  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
446  | 
val compiler = ghcN  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
447  | 
val modules = map fst code_files  | 
| 58039 | 448  | 
|
| 
72511
 
460d743010bc
clarified signature: overloaded "+" for Path.append;
 
wenzelm 
parents: 
72376 
diff
changeset
 | 
449  | 
val driver_path = dir + Path.basic "Main.hs"  | 
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
450  | 
val driver =  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
451  | 
      "module Main where {\n" ^
 | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
452  | 
implode (map (fn module => "import qualified " ^ unsuffix ".hs" module ^ ";\n") modules) ^  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
453  | 
      "main = do {\n" ^
 | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
454  | 
      "    let {\n" ^
 | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
455  | 
" format_term Nothing = \"\";\n" ^  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
456  | 
" format_term (Just t) = t ();\n" ^  | 
| 72307 | 457  | 
" format (True, _) = \"" ^ success ^ "\\n\";\n" ^  | 
458  | 
" format (False, to) = \"" ^ failure ^ "\" ++ format_term to ++ \"\\n\";\n" ^  | 
|
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
459  | 
" result = " ^ value_name ^ " ();\n" ^  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
460  | 
" };\n" ^  | 
| 72307 | 461  | 
" Prelude.putStr \"" ^ start_marker ^ "\";\n" ^  | 
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
462  | 
" Prelude.mapM_ (putStr . format) result;\n" ^  | 
| 72307 | 463  | 
" Prelude.putStr \"" ^ end_marker ^ "\";\n" ^  | 
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
464  | 
" }\n" ^  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
465  | 
"}\n"  | 
| 58039 | 466  | 
|
| 72306 | 467  | 
val _ = check_settings compiler ISABELLE_GHC "GHC executable"  | 
| 
72511
 
460d743010bc
clarified signature: overloaded "+" for Path.append;
 
wenzelm 
parents: 
72376 
diff
changeset
 | 
468  | 
val _ = List.app (fn (name, code) => File.write (dir + Path.basic name) code) code_files  | 
| 72306 | 469  | 
val _ = File.write driver_path driver  | 
470  | 
||
| 
72511
 
460d743010bc
clarified signature: overloaded "+" for Path.append;
 
wenzelm 
parents: 
72376 
diff
changeset
 | 
471  | 
val compiled_path = dir + Path.basic "test"  | 
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
472  | 
val cmd =  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
473  | 
"\"$ISABELLE_GHC\" " ^ Code_Haskell.language_params ^ " " ^  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
474  | 
Config.get ctxt ghc_options ^ " -o " ^  | 
| 72278 | 475  | 
File.bash_platform_path compiled_path ^ " " ^  | 
476  | 
File.bash_platform_path driver_path ^ " -i" ^  | 
|
477  | 
File.bash_platform_path dir  | 
|
| 72306 | 478  | 
in compile compiler cmd; evaluate compiler (File.bash_path compiled_path) end  | 
| 58039 | 479  | 
|
| 64577 | 480  | 
|
481  | 
(* driver for Scala *)  | 
|
| 58039 | 482  | 
|
483  | 
val scalaN = "Scala"  | 
|
484  | 
||
| 72297 | 485  | 
fun evaluate_in_scala (_: Proof.context) (code_files, value_name) dir =  | 
| 
72276
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
486  | 
let  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
487  | 
val generatedN = "Generated_Code"  | 
| 
 
dfe150a246e6
misc tuning and clarification: prefer functions over data;
 
wenzelm 
parents: 
72275 
diff
changeset
 | 
488  | 
val driverN = "Driver.scala"  | 
| 58039 | 489  | 
|
| 
72511
 
460d743010bc
clarified signature: overloaded "+" for Path.append;
 
wenzelm 
parents: 
72376 
diff
changeset
 | 
490  | 
val code_path = dir + Path.basic (generatedN ^ ".scala")  | 
| 
 
460d743010bc
clarified signature: overloaded "+" for Path.append;
 
wenzelm 
parents: 
72376 
diff
changeset
 | 
491  | 
val driver_path = dir + Path.basic driverN  | 
| 
 
460d743010bc
clarified signature: overloaded "+" for Path.append;
 
wenzelm 
parents: 
72376 
diff
changeset
 | 
492  | 
val out_path = dir + Path.basic "out"  | 
| 72297 | 493  | 
|
| 72306 | 494  | 
val code = #2 (the_single code_files)  | 
| 72296 | 495  | 
val driver = \<^verbatim>\<open>  | 
496  | 
{
 | 
|
497  | 
val result = \<close> ^ value_name ^ \<^verbatim>\<open>(())  | 
|
498  | 
val result_text =  | 
|
499  | 
result.map(  | 
|
500  | 
      {
 | 
|
501  | 
case (true, _) => "True\n"  | 
|
502  | 
case (false, None) => "False\n"  | 
|
503  | 
case (false, Some(t)) => "False" + t(()) + "\n"  | 
|
504  | 
}).mkString  | 
|
505  | 
isabelle.File.write(  | 
|
506  | 
isabelle.Path.explode(\<close> ^ quote (Path.implode (Path.expand out_path)) ^ \<^verbatim>\<open>),  | 
|
| 72307 | 507  | 
\<close> ^ quote start_marker ^ \<^verbatim>\<open> + result_text + \<close> ^ quote end_marker ^ \<^verbatim>\<open>)  | 
| 72296 | 508  | 
}\<close>  | 
| 72306 | 509  | 
val _ = File.write code_path code  | 
510  | 
val _ = File.write driver_path driver  | 
|
511  | 
val _ = Scala_Compiler.toplevel true (code ^ driver)  | 
|
| 72308 | 512  | 
      handle ERROR msg => error ("Evaluation for " ^ scalaN ^ " failed:\n" ^ msg)
 | 
| 72306 | 513  | 
in File.read out_path end  | 
| 58039 | 514  | 
|
| 64580 | 515  | 
|
516  | 
(* command setup *)  | 
|
| 58039 | 517  | 
|
| 64577 | 518  | 
val _ =  | 
| 69593 | 519  | 
Outer_Syntax.command \<^command_keyword>\<open>test_code\<close>  | 
| 58039 | 520  | 
"compile test cases to target languages, execute them and report results"  | 
| 69593 | 521  | 
(Scan.repeat1 Parse.prop -- (\<^keyword>\<open>in\<close> |-- Scan.repeat1 Parse.name)  | 
| 64580 | 522  | 
>> (fn (ts, targets) => Toplevel.keep (test_code_cmd ts targets o Toplevel.context_of)))  | 
| 58039 | 523  | 
|
| 
66284
 
378895354604
new derived targets for evaluating Haskell and Scala programs
 
Andreas Lochbihler 
parents: 
65905 
diff
changeset
 | 
524  | 
val target_Scala = "Scala_eval"  | 
| 
 
378895354604
new derived targets for evaluating Haskell and Scala programs
 
Andreas Lochbihler 
parents: 
65905 
diff
changeset
 | 
525  | 
val target_Haskell = "Haskell_eval"  | 
| 
 
378895354604
new derived targets for evaluating Haskell and Scala programs
 
Andreas Lochbihler 
parents: 
65905 
diff
changeset
 | 
526  | 
|
| 64580 | 527  | 
val _ =  | 
| 72306 | 528  | 
Theory.setup  | 
529  | 
(Code_Target.add_derived_target (target_Scala, [(Code_Scala.target, I)]) #>  | 
|
530  | 
Code_Target.add_derived_target (target_Haskell, [(Code_Haskell.target, I)]) #>  | 
|
531  | 
fold add_driver  | 
|
532  | 
[(polymlN, (evaluate_in_polyml, Code_ML.target_SML)),  | 
|
533  | 
(mltonN, (evaluate_in_mlton, Code_ML.target_SML)),  | 
|
534  | 
(smlnjN, (evaluate_in_smlnj, Code_ML.target_SML)),  | 
|
535  | 
(ocamlN, (evaluate_in_ocaml, Code_ML.target_OCaml)),  | 
|
536  | 
(ghcN, (evaluate_in_ghc, target_Haskell)),  | 
|
537  | 
(scalaN, (evaluate_in_scala, target_Scala))] #>  | 
|
538  | 
fold (fn target => Value_Command.add_evaluator (Binding.name target, eval_term target) #> #2)  | 
|
| 64580 | 539  | 
[polymlN, mltonN, smlnjN, ocamlN, ghcN, scalaN])  | 
| 59323 | 540  | 
|
| 58039 | 541  | 
end  |