src/Pure/Thy/thy_load.ML
changeset 37949 48a874444164
parent 37943 3cbd7fa164b1
child 37950 bc285d91041e
--- a/src/Pure/Thy/thy_load.ML	Fri Jul 23 18:42:46 2010 +0200
+++ b/src/Pure/Thy/thy_load.ML	Sat Jul 24 12:14:53 2010 +0200
@@ -17,6 +17,10 @@
 signature THY_LOAD =
 sig
   include BASIC_THY_LOAD
+  val master_directory: theory -> Path.T
+  val init: Path.T -> theory -> theory
+  val require: Path.T -> theory -> theory
+  val provide: Path.T * (Path.T * File.ident) -> theory -> theory
   val ml_path: string -> Path.T
   val thy_path: string -> Path.T
   val split_thy_path: Path.T -> Path.T * string
@@ -26,12 +30,57 @@
   val check_thy: Path.T -> string -> Path.T * File.ident
   val deps_thy: Path.T -> string ->
    {master: Path.T * File.ident, text: string, imports: string list, uses: Path.T list}
-  val load_ml: Path.T -> Path.T -> Path.T * File.ident
+  val loaded_files: theory -> Path.T list
+  val check_loaded: theory -> unit
+  val all_current: theory -> bool
+  val provide_file: Path.T -> theory -> theory
+  val use_ml: Path.T -> unit
+  val exec_ml: Path.T -> generic_theory -> generic_theory
 end;
 
 structure Thy_Load: THY_LOAD =
 struct
 
+(* manage source files *)
+
+type files =
+ {master_dir: Path.T,                                 (*master directory of theory source*)
+  required: Path.T list,                              (*source path*)
+  provided: (Path.T * (Path.T * File.ident)) list};   (*source path, physical path, identifier*)
+
+fun make_files (master_dir, required, provided): files =
+ {master_dir = master_dir, required = required, provided = provided};
+
+structure Files = Theory_Data
+(
+  type T = files;
+  val empty = make_files (Path.current, [], []);
+  fun extend _ = empty;
+  fun merge _ = empty;
+);
+
+fun map_files f =
+  Files.map (fn {master_dir, required, provided} =>
+    make_files (f (master_dir, required, provided)));
+
+
+val master_directory = #master_dir o Files.get;
+
+fun init master_dir = map_files (fn _ => (master_dir, [], []));
+
+fun require src_path =
+  map_files (fn (master_dir, required, provided) =>
+    if member (op =) required src_path then
+      error ("Duplicate source file dependency: " ^ Path.implode src_path)
+    else (master_dir, src_path :: required, provided));
+
+fun provide (src_path, path_info) =
+  map_files (fn (master_dir, required, provided) =>
+    if AList.defined (op =) provided src_path then
+      error ("Duplicate resolution of source file dependency: " ^ Path.implode src_path)
+    else (master_dir, required, (src_path, path_info) :: provided));
+
+
 (* maintain load path *)
 
 local val load_path = Unsynchronized.ref [Path.current] in
@@ -121,13 +170,56 @@
   in {master = master, text = text, imports = imports, uses = uses} end;
 
 
-(* ML files *)
+
+(* loaded files *)
+
+val loaded_files = map (#1 o #2) o #provided o Files.get;
 
-fun load_ml dir raw_path =
+fun check_loaded thy =
+  let
+    val {required, provided, ...} = Files.get thy;
+    val provided_paths = map #1 provided;
+    val _ =
+      (case subtract (op =) provided_paths required of
+        [] => NONE
+      | bad => error ("Pending source file dependencies: " ^ commas (map Path.implode (rev bad))));
+    val _ =
+      (case subtract (op =) required provided_paths of
+        [] => NONE
+      | bad => error ("Undeclared source file dependencies: " ^ commas (map Path.implode (rev bad))));
+  in () end;
+
+fun all_current thy =
   let
-    val (path, id) = check_file dir raw_path;
-    val _ = ML_Context.eval_file path;
-  in (path, id) end;
+    val {master_dir, provided, ...} = Files.get thy;
+    fun current (src_path, path_info) =
+      (case test_file master_dir src_path of
+        NONE => false
+      | SOME path_info' => eq_snd (op =) (path_info, path_info'));
+  in can check_loaded thy andalso forall current provided end;
+
+
+(* provide files *)
+
+fun provide_file src_path thy =
+  provide (src_path, check_file (master_directory thy) src_path) thy;
+
+fun use_ml src_path =
+  if is_none (Context.thread_data ()) then
+    ML_Context.eval_file (#1 (check_file Path.current src_path))
+  else
+    let
+      val thy = ML_Context.the_global_context ();
+      val (path, info) = check_file (master_directory thy) src_path;
+
+      val _ = ML_Context.eval_file path;
+      val _ = Context.>> Local_Theory.propagate_ml_env;
+
+      val provide = provide (src_path, (path, info));
+      val _ = Context.>> (Context.mapping provide (Local_Theory.theory provide));
+    in () end
+
+fun exec_ml src_path = ML_Context.exec (fn () => use_ml src_path);
 
 end;