evaluate takes ml context and ml expression parameter
authorhaftmann
Tue, 31 Aug 2010 16:23:58 +0200
changeset 38931 5e84c11c4b8a
parent 38930 072363be3fd9
child 38932 515059ca8022
evaluate takes ml context and ml expression parameter
src/Pure/ML/ml_context.ML
src/Tools/Code/code_eval.ML
src/Tools/nbe.ML
--- a/src/Pure/ML/ml_context.ML	Tue Aug 31 16:07:30 2010 +0200
+++ b/src/Pure/ML/ml_context.ML	Tue Aug 31 16:23:58 2010 +0200
@@ -35,8 +35,8 @@
   val eval_text_in: Proof.context option -> bool -> Position.T -> Symbol_Pos.text -> unit
   val expression: Position.T -> string -> string -> ML_Lex.token Antiquote.antiquote list ->
     Context.generic -> Context.generic
-  val evaluate:
-    Proof.context -> bool -> string * (unit -> 'a) option Unsynchronized.ref -> string -> 'a
+  val evaluate: Proof.context -> bool ->
+    string * (unit -> 'a) option Unsynchronized.ref -> string * string -> 'a
 end
 
 structure ML_Context: ML_CONTEXT =
@@ -203,10 +203,10 @@
 
 
 (* FIXME not thread-safe -- reference can be accessed directly *)
-fun evaluate ctxt verbose (ref_name, r) txt = NAMED_CRITICAL "ML" (fn () =>
+fun evaluate ctxt verbose (ref_name, r) (prelude, txt) = NAMED_CRITICAL "ML" (fn () =>
   let
-    val ants =
-      ML_Lex.read Position.none ("val _ = (" ^ ref_name ^ " := SOME (fn () => " ^ txt ^ "))");
+    val s = prelude ^ "val _ = (" ^ ref_name ^ " := SOME (fn () => " ^ txt ^ "))"
+    val ants = ML_Lex.read Position.none s;
     val _ = r := NONE;
     val _ =
       Context.setmp_thread_data (SOME (Context.Proof ctxt))
--- a/src/Tools/Code/code_eval.ML	Tue Aug 31 16:07:30 2010 +0200
+++ b/src/Tools/Code/code_eval.ML	Tue Aug 31 16:23:58 2010 +0200
@@ -51,11 +51,11 @@
           |> Graph.new_node (value_name,
               Code_Thingol.Fun (Term.dummy_patternN, ((([], ty), [(([], t), (NONE, true))]), NONE)))
           |> fold (curry Graph.add_edge value_name) deps;
-        val (value_code, [SOME value_name']) = Code_Target.produce_code_for thy
-          (the_default target some_target) NONE (SOME "") [] naming program [value_name];
-        val sml_code = "let\n" ^ value_code ^ "\nin " ^ Long_Name.base_name value_name'
-          ^ space_implode " " (map (enclose "(" ")") args) ^ " end";
-      in ML_Context.evaluate ctxt false reff sml_code end;
+        val (program_code, [SOME value_name']) = Code_Target.produce_code_for thy
+          (the_default target some_target) NONE (SOME "Code") [] naming program' [value_name];
+        val value_code = space_implode " "
+          (value_name' :: map (enclose "(" ")") args);
+      in ML_Context.evaluate ctxt false reff (program_code, value_code) end;
   in Code_Thingol.dynamic_eval_value thy postproc evaluator t end;
 
 
--- a/src/Tools/nbe.ML	Tue Aug 31 16:07:30 2010 +0200
+++ b/src/Tools/nbe.ML	Tue Aug 31 16:23:58 2010 +0200
@@ -388,6 +388,7 @@
       in
         s
         |> traced (fn s => "\n--- code to be evaluated:\n" ^ s)
+        |> pair ""
         |> ML_Context.evaluate ctxt (!trace) univs_cookie
         |> (fn f => f deps_vals)
         |> (fn univs => cs ~~ univs)