explicit handling of ML environment within generic context;
authorwenzelm
Wed, 17 Sep 2008 21:27:32 +0200
changeset 28270 7ada24ebea94
parent 28269 b1e5e6c4c10e
child 28271 0e1b07ded55f
explicit handling of ML environment within generic context; eval_wrapper: non-critical, struct_name is always Isabelle, tuned comments; eval_in/evaluate expect immutable Proof.context value; use_text/use_file now depend on explicit ML name space;
src/Pure/ML/ml_context.ML
--- a/src/Pure/ML/ml_context.ML	Wed Sep 17 21:27:31 2008 +0200
+++ b/src/Pure/ML/ml_context.ML	Wed Sep 17 21:27:32 2008 +0200
@@ -20,6 +20,7 @@
   val the_global_context: unit -> theory
   val the_local_context: unit -> Proof.context
   val exec: (unit -> unit) -> Context.generic -> Context.generic
+  val name_space: ML_NameSpace.nameSpace
   val stored_thms: thm list ref
   val ml_store_thm: string * thm -> unit
   val ml_store_thms: string * thm list -> unit
@@ -32,8 +33,8 @@
   val eval_wrapper: (string -> unit) * (string -> 'a) -> bool -> Position.T -> string -> unit
   val eval: bool -> Position.T -> string -> unit
   val eval_file: Path.T -> unit
-  val eval_in: Context.generic option -> bool -> Position.T -> string -> unit
-  val evaluate: (string -> unit) * (string -> 'b) -> bool ->
+  val eval_in: Proof.context option -> bool -> Position.T -> string -> unit
+  val evaluate: Proof.context -> (string -> unit) * (string -> 'b) -> bool ->
     string * (unit -> 'a) option ref -> string -> 'a
   val expression: Position.T -> string -> string -> string -> Context.generic -> Context.generic
 end
@@ -53,6 +54,69 @@
   | NONE => error "Missing context after execution");
 
 
+(* ML name space *)
+
+structure ML_Env = GenericDataFun
+(
+  type T =
+    ML_NameSpace.valueVal Symtab.table *
+    ML_NameSpace.typeVal Symtab.table *
+    ML_NameSpace.fixityVal Symtab.table *
+    ML_NameSpace.structureVal Symtab.table *
+    ML_NameSpace.signatureVal Symtab.table *
+    ML_NameSpace.functorVal Symtab.table;
+  val empty = (Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty);
+  val extend = I;
+  fun merge _
+   ((val1, type1, fixity1, structure1, signature1, functor1),
+    (val2, type2, fixity2, structure2, signature2, functor2)) : T =
+   (Symtab.merge (K true) (val1, val2),
+    Symtab.merge (K true) (type1, type2),
+    Symtab.merge (K true) (fixity1, fixity2),
+    Symtab.merge (K true) (structure1, structure2),
+    Symtab.merge (K true) (signature1, signature2),
+    Symtab.merge (K true) (functor1, functor2));
+);
+
+val name_space: ML_NameSpace.nameSpace =
+  let
+    fun lookup sel1 sel2 name =
+      Context.thread_data ()
+      |> (fn NONE => NONE | SOME context => Symtab.lookup (sel1 (ML_Env.get context)) name)
+      |> (fn NONE => sel2 ML_NameSpace.global name | some => some);
+
+    fun all sel1 sel2 () =
+      Context.thread_data ()
+      |> (fn NONE => [] | SOME context => Symtab.dest (sel1 (ML_Env.get context)))
+      |> append (sel2 ML_NameSpace.global ())
+      |> sort_distinct (string_ord o pairself #1);
+
+    fun enter ap1 sel2 entry =
+      if is_some (Context.thread_data ()) then
+        Context.>> (ML_Env.map (ap1 (Symtab.update entry)))
+      else sel2 ML_NameSpace.global entry;
+  in
+   {lookupVal    = lookup #1 #lookupVal,
+    lookupType   = lookup #2 #lookupType,
+    lookupFix    = lookup #3 #lookupFix,
+    lookupStruct = lookup #4 #lookupStruct,
+    lookupSig    = lookup #5 #lookupSig,
+    lookupFunct  = lookup #6 #lookupFunct,
+    enterVal     = enter (fn h => fn (a, b, c, d, e, f) => (h a, b, c, d, e, f)) #enterVal,
+    enterType    = enter (fn h => fn (a, b, c, d, e, f) => (a, h b, c, d, e, f)) #enterType,
+    enterFix     = enter (fn h => fn (a, b, c, d, e, f) => (a, b, h c, d, e, f)) #enterFix,
+    enterStruct  = enter (fn h => fn (a, b, c, d, e, f) => (a, b, c, h d, e, f)) #enterStruct,
+    enterSig     = enter (fn h => fn (a, b, c, d, e, f) => (a, b, c, d, h e, f)) #enterSig,
+    enterFunct   = enter (fn h => fn (a, b, c, d, e, f) => (a, b, c, d, e, h f)) #enterFunct,
+    allVal       = all #1 #allVal,
+    allType      = all #2 #allType,
+    allFix       = all #3 #allFix,
+    allStruct    = all #4 #allStruct,
+    allSig       = all #5 #allSig,
+    allFunct     = all #6 #allFunct}
+  end;
+
+
 (* theorem bindings *)
 
 val stored_thms: thm list ref = ref [];
@@ -65,7 +129,7 @@
       else if not (ML_Syntax.is_identifier name) then
         error ("Cannot bind theorem(s) " ^ quote name ^ " as ML value")
       else setmp stored_thms ths' (fn () =>
-        use_text (0, "") Output.ml_output true
+        use_text name_space (0, "") Output.ml_output true
           ("val " ^ name ^ " = " ^ sel ^ "(! ML_Context.stored_thms);")) ();
   in () end;
 
@@ -123,10 +187,13 @@
   P.!!! (P.position P.xname -- Args.parse --| Scan.ahead P.eof)
   >> (fn ((x, pos), y) => Args.src ((x, y), pos));
 
-fun eval_antiquotes struct_name (txt, pos) opt_ctxt =
+val struct_name = "Isabelle";
+
+fun eval_antiquotes (txt, pos) opt_context =
   let
     val syms = SymbolPos.explode (txt, pos);
     val ants = Antiquote.read (syms, pos);
+    val opt_ctxt = Option.map (Context.Proof o Context.proof_of) opt_context;
     val ((ml_env, ml_body), opt_ctxt') =
       if not (exists Antiquote.is_antiq ants)
       then (("", Symbol.escape (SymbolPos.content syms)), opt_ctxt)
@@ -136,7 +203,7 @@
             (case opt_ctxt of
               NONE => error ("Unknown context -- cannot expand ML antiquotations" ^
                 Position.str_of pos)
-            | SOME context => Context.proof_of context);
+            | SOME ctxt => Context.proof_of ctxt);
 
           val lex = #1 (OuterKeyword.get_lexicons ());
           fun no_decl _ = ("", "");
@@ -164,19 +231,25 @@
 
 fun eval_wrapper pr verbose pos txt =
   let
+    fun eval_raw p = use_text name_space
+      (the_default 1 (Position.line_of p), the_default "ML" (Position.file_of p)) pr;
+
+    (*prepare source text*)
     val _ = Position.report Markup.ML_source pos;
-    val struct_name =
-      if Multithreading.available then "Isabelle" ^ serial_string ()
-      else "Isabelle";
-    val ((env, body), env_ctxt) = eval_antiquotes struct_name (txt, pos) (Context.thread_data ());
+    val ((env, body), env_ctxt) = eval_antiquotes (txt, pos) (Context.thread_data ());
     val _ = if ! trace then tracing (cat_lines [env, body]) else ();
-    fun eval_raw p =
-      use_text (the_default 1 (Position.line_of p), the_default "ML" (Position.file_of p)) pr;
-  in
-    Context.setmp_thread_data env_ctxt (fn () => eval_raw Position.none false env) ();
-    NAMED_CRITICAL "ML" (fn () => eval_raw pos verbose body); (* FIXME non-critical with local ML env *)
-    forget_structure struct_name
-  end;
+
+    (*prepare static ML environment*)
+    val _ = Context.setmp_thread_data env_ctxt
+        (fn () => (eval_raw Position.none false env; Context.thread_data ())) ()
+      |> (fn NONE => () | SOME context' => Context.>> (ML_Env.put (ML_Env.get context')));
+
+    (*eval ML*)
+    val _ = eval_raw pos verbose body;
+
+    (*reset static ML environment*)
+    val _ = eval_raw Position.none false "structure Isabelle = struct end";
+  in () end;
 
 end;
 
@@ -186,14 +259,15 @@
 val eval = eval_wrapper Output.ml_output;
 fun eval_file path = eval true (Path.position path) (File.read path);
 
-fun eval_in context verbose pos txt =
-  Context.setmp_thread_data context (fn () => eval verbose pos txt) ();
+fun eval_in ctxt verbose pos txt =
+  Context.setmp_thread_data (Option.map Context.Proof ctxt) (fn () => eval verbose pos txt) ();
 
-fun evaluate pr verbose (ref_name, r) txt = NAMED_CRITICAL "ML" (fn () =>
+fun evaluate ctxt pr verbose (ref_name, r) txt = NAMED_CRITICAL "ML" (fn () =>
   let
     val _ = r := NONE;
-    val _ = eval_wrapper pr verbose Position.none
-      ("val _ = (" ^ ref_name ^ " := SOME (fn () => " ^ txt ^ "))");
+    val _ = Context.setmp_thread_data (SOME (Context.Proof ctxt)) (fn () =>
+      eval_wrapper pr verbose Position.none
+        ("val _ = (" ^ ref_name ^ " := SOME (fn () => " ^ txt ^ "))")) ();
   in (case ! r of NONE => error ("Bad evaluation for " ^ ref_name) | SOME e => e) end) ();
 
 fun expression pos bind body txt =
@@ -205,4 +279,3 @@
 
 structure Basic_ML_Context: BASIC_ML_CONTEXT = ML_Context;
 open Basic_ML_Context;
-