removed obsolete use/update_thy_only;
authorwenzelm
Thu, 19 Jul 2007 23:18:58 +0200
changeset 23871 0dd5696f58b6
parent 23870 dde006281806
child 23872 f449381e2caa
removed obsolete use/update_thy_only; removed unused quiet_update_thy, get_succs; renamed get_preds to get_parents; deps: replaced File.info by File.ident (no comparison of paths!); check_deps: reload (partially qualified) parents for unfinished theory, no reference of previously loaded master paths! require_thy: attempt at purely static path lookup, less permissive; misc cleanup;
src/Pure/Thy/thy_info.ML
--- a/src/Pure/Thy/thy_info.ML	Thu Jul 19 23:18:56 2007 +0200
+++ b/src/Pure/Thy/thy_info.ML	Thu Jul 19 23:18:58 2007 +0200
@@ -16,8 +16,6 @@
   val update_thy: string -> unit
   val remove_thy: string -> unit
   val time_use_thy: string -> unit
-  val use_thy_only: string -> unit
-  val update_thy_only: string -> unit
 end;
 
 signature THY_INFO =
@@ -33,16 +31,14 @@
   val lookup_theory: string -> theory option
   val get_theory: string -> theory
   val the_theory: string -> theory -> theory
-  val get_preds: string -> string list
-  val get_succs: string -> string list
+  val get_parents: string -> string list
   val loaded_files: string -> Path.T list
   val touch_child_thys: string -> unit
   val touch_all_thys: unit -> unit
   val load_file: bool -> Path.T -> unit
   val use: string -> unit
-  val quiet_update_thy: bool -> string -> unit
   val pretend_use_thy_only: string -> unit
-  val begin_theory: (Path.T option -> string -> string list ->
+  val begin_theory: (Path.T list -> string -> string list ->
       (Path.T * bool) list -> theory -> theory) ->
     string -> string list -> (Path.T * bool) list -> bool -> theory
   val end_theory: theory -> theory
@@ -54,7 +50,6 @@
 structure ThyInfo: THY_INFO =
 struct
 
-
 (** theory loader actions and hooks **)
 
 datatype action = Update | Outdate | Remove;
@@ -73,10 +68,8 @@
 
 (* messages *)
 
-fun gen_msg txt [] = txt
-  | gen_msg txt names = txt ^ " " ^ commas_quote names;
-
-fun loader_msg txt names = gen_msg ("Theory loader: " ^ txt) names;
+fun loader_msg txt [] = "Theory loader: " ^ txt
+  | loader_msg txt names = "Theory loader: " ^ txt ^ " " ^ commas_quote names;
 
 val show_path = space_implode " via " o map quote;
 fun cycle_msg names = loader_msg ("cyclic dependency of " ^ show_path names) [];
@@ -98,16 +91,27 @@
 
 (* thy database *)
 
+type master = (Path.T * File.ident) * (Path.T * File.ident) option;
+
 type deps =
-  {present: bool, outdated: bool,
-    master: ThyLoad.master option, files: (Path.T * (Path.T * File.info) option) list};
+  {present: bool,
+    outdated: bool,
+    master: master option,
+    files: (Path.T * (Path.T * File.ident) option) list};
+
+fun master_idents (NONE: master option) = []
+  | master_idents (SOME ((_, thy_id), NONE)) = [thy_id]
+  | master_idents (SOME ((_, thy_id), SOME (_, ml_id))) = [thy_id, ml_id];
+
+fun master_path (m: master option) = map (Path.dir o fst o fst) (the_list m);
+fun master_path' (d: deps option) = maps (master_path o #master) (the_list d);
+fun master_path'' d = maps master_path' (the_list d);
 
 fun make_deps present outdated master files =
   {present = present, outdated = outdated, master = master, files = files}: deps;
 
 fun init_deps master files = SOME (make_deps false true master (map (rpair NONE) files));
 
-
 type thy = deps option * theory option;
 
 local
@@ -158,10 +162,10 @@
   (case get_deps name of
     NONE => []
   | SOME {master, files, ...} =>
-      (case master of SOME m => [#1 (ThyLoad.get_thy m)] | NONE => []) @
+      (case master of SOME ((thy_path, _), _) => [thy_path] | NONE => []) @
       (map_filter (Option.map #1 o #2) files));
 
-fun get_preds name =
+fun get_parents name =
   (thy_graph Graph.imm_preds name) handle Graph.UNDEF _ =>
     error (loader_msg "nothing known about theory" [name]);
 
@@ -194,7 +198,7 @@
 
 fun get_theory name =
   (case lookup_theory name of
-    (SOME theory) => theory
+    SOME theory => theory
   | _ => error (loader_msg "undefined theory entry for" [name]));
 
 fun the_theory name thy =
@@ -238,10 +242,6 @@
 fun touch_thy name = touch_thys [name];
 fun touch_child_thys name = touch_thys (thy_graph Graph.imm_succs name);
 
-fun get_succs name =
-  (thy_graph Graph.imm_succs name) handle Graph.UNDEF _ =>
-    error (loader_msg "nothing known about theory" [name]);
-
 fun touch_all_thys () = List.app outdate_thy (get_names ());
 
 end;
@@ -257,14 +257,10 @@
 
 (* load_file *)
 
-val opt_path = Option.map (Path.dir o fst o ThyLoad.get_thy);
-fun opt_path' (d: deps option) = the_default NONE (Option.map (opt_path o #master) d);
-fun opt_path'' d = the_default NONE (Option.map opt_path' d);
-
 local
 
 fun provide path name info (deps as SOME {present, outdated, master, files}) =
-     (if exists (equal path o #1) files then ()
+     (if AList.defined (op =) files path then ()
       else warning (loader_msg "undeclared dependency of theory" [name] ^
         " on file: " ^ quote (Path.implode path));
       SOME (make_deps present outdated master (AList.update (op =) (path, SOME info) files)))
@@ -272,11 +268,12 @@
 
 fun run_file path =
   (case Option.map (Context.theory_name o Context.the_theory) (ML_Context.get_context ()) of
-    NONE => (ThyLoad.load_file NONE path; ())
-  | SOME name => (case lookup_deps name of
-      SOME deps => change_deps name (provide path name
-        (ThyLoad.load_file (opt_path' deps) path))
-    | NONE => (ThyLoad.load_file NONE path; ())));
+    NONE => (ThyLoad.load_ml [] path; ())
+  | SOME name =>
+      (case lookup_deps name of
+        SOME deps =>
+          change_deps name (provide path name (ThyLoad.load_ml (master_path' deps) path))
+      | NONE => (ThyLoad.load_ml [] path; ())));
 
 in
 
@@ -301,7 +298,7 @@
 fun required_by _ [] = ""
   | required_by s initiators = s ^ "(required by " ^ show_path (rev initiators) ^ ")";
 
-fun load_thy really ml time initiators dir name =
+fun load_thy really ml time initiators dirs name =
   let
     val _ =
       if really then priority ("Loading theory " ^ quote name ^ required_by " " initiators)
@@ -309,8 +306,8 @@
 
     val _ = touch_thy name;
     val master =
-      if really then ThyLoad.load_thy dir name ml time
-      else #1 (ThyLoad.deps_thy dir name ml);
+      if really then ThyLoad.load_thy dirs name ml time
+      else #1 (ThyLoad.deps_thy dirs name ml);
 
     val files = get_files name;
     val missing_files = map_filter (fn (path, NONE) => SOME (Path.implode path) | _ => NONE) files;
@@ -325,90 +322,84 @@
 
 (* require_thy *)
 
-fun base_of_path s = Path.implode (Path.base (Path.explode s));
+fun base_name s = Path.implode (Path.base (Path.explode s));
 
 local
 
-fun load_deps ml dir name =
-  let val (master, (parents, files)) = ThyLoad.deps_thy dir name ml
-  in (SOME (init_deps (SOME master) files), parents) end;
+fun check_ml master (path, info) =
+  let val info' =
+    (case info of NONE => NONE
+    | SOME (_, id) =>
+        (case ThyLoad.check_ml (master_path master) path of NONE => NONE
+        | SOME (path', id') => if id <> id' then NONE else SOME (path', id')))
+  in (path, info') end;
 
-fun file_current master (_, NONE) = false
-  | file_current master (path, info) =
-      (info = ThyLoad.check_file (opt_path master) path);
-
-fun current_deps ml strict dir name =
+fun check_deps ml strict dirs name =
   (case lookup_deps name of
-    NONE => (false, load_deps ml dir name)
-  | SOME deps =>
-      let
-        fun get_name s = (case opt_path'' (lookup_deps s) of NONE => s
-          | SOME p => Path.implode (Path.append p (Path.basic s)));
-        val same_deps = (NONE, map get_name (thy_graph Graph.imm_preds name))
-      in
-        (case deps of
-          NONE => (true, same_deps)
-        | SOME {present, outdated, master, files} =>
-            if present andalso not strict then (true, same_deps)
-            else
-              let val master' = SOME (ThyLoad.check_thy dir name ml) in
-                if master <> master' then (false, load_deps ml dir name)
-                else (not outdated andalso forall (file_current master') files, same_deps)
-              end)
-      end);
+    NONE =>
+      let val (master, (parents, files)) = ThyLoad.deps_thy dirs name ml |>> SOME
+      in (false, (init_deps master files, parents)) end
+  | SOME NONE => (true, (NONE, get_parents name))
+  | SOME (deps as SOME {present, outdated, master, files}) =>
+      if present andalso not strict then (true, (deps, get_parents name))
+      else
+        let val (master', (parents', files')) = ThyLoad.deps_thy dirs name ml |>> SOME in
+          if master_idents master <> master_idents master'
+          then (false, (init_deps master' files', parents'))
+          else
+            let
+              val checked_files = map (check_ml master') files;
+              val current = not outdated andalso forall (is_some o snd) checked_files;
+              val deps' = SOME (make_deps present (not current) master' checked_files);
+            in (current, (deps', parents')) end
+        end);
 
-fun require_thy really ml time strict keep_strict initiators prfx (visited, str) =
+fun require_thy really ml time strict keep_strict initiators dirs str visited =
   let
     val path = Path.expand (Path.explode str);
     val name = Path.implode (Path.base path);
+    val dirs' = Path.dir path :: dirs;
+    val _ = member (op =) initiators name andalso error (cycle_msg initiators);
   in
-    if member (op =) initiators name then error (cycle_msg initiators) else ();
-    if known_thy name andalso is_finished name orelse member (op =) visited name then
-      (visited, (true, name))
+    if known_thy name andalso is_finished name orelse member (op =) visited name
+    then ((true, name), visited)
     else
       let
-        val dir = Path.append prfx (Path.dir path);
-        val req_parent = require_thy true true time (strict andalso keep_strict) keep_strict
-          (name :: initiators);
-
-        val (current, (new_deps, parents)) = current_deps ml strict dir name
+        val (current, (deps, parents)) = check_deps ml strict dirs' name
           handle ERROR msg => cat_error msg
             (loader_msg "the error(s) above occurred while examining theory" [name] ^
               (if null initiators then "" else required_by "\n" initiators));
-        val dir' = the_default dir (opt_path'' new_deps);
-        val (visited', parent_results) = foldl_map (req_parent dir') (name :: visited, parents);
+
+        val dirs'' = master_path' deps @ dirs;
+        val req_parent = require_thy true true time (strict andalso keep_strict) keep_strict
+          (name :: initiators) dirs'';
+        val (parent_results, visited') = fold_map req_parent parents (name :: visited);
 
-        val thy = if not really then SOME (get_theory name) else NONE;
-        val result =
-          if current andalso forall fst parent_results then true
-          else
-            ((case new_deps of
-              SOME deps => change_thys (update_node name (map base_of_path parents) (deps, thy))
-            | NONE => ());
-             load_thy really ml (time orelse ! Output.timing) initiators dir name;
-             false);
-      in (visited', (result, name)) end
+        val all_current = current andalso forall fst parent_results;
+        val thy = if all_current orelse not really then SOME (get_theory name) else NONE;
+        val _ = change_thys (update_node name (map base_name parents) (deps, thy));
+        val _ =
+          if all_current then ()
+          else load_thy really ml (time orelse ! Output.timing) initiators dirs' name;
+      in ((all_current, name), visited') end
   end;
 
-fun gen_use_thy' req prfx s = Output.toplevel_errors (fn () =>
-  let val (_, (_, name)) = req [] prfx ([], s)
+fun gen_use_thy' req dirs s = Output.toplevel_errors (fn () =>
+  let val ((_, name), _) = req [] dirs s []
   in ML_Context.set_context (SOME (Context.Theory (get_theory name))) end) ();
 
-fun gen_use_thy req = gen_use_thy' req Path.current;
+fun gen_use_thy req = gen_use_thy' req [];
 
 fun warn_finished f name = (check_unfinished warning name; f name);
 
 in
 
-val weak_use_thy         = gen_use_thy' (require_thy true true false false false);
-fun quiet_update_thy' ml = gen_use_thy' (require_thy true ml false true true);
-fun quiet_update_thy ml  = gen_use_thy (require_thy true ml false true true);
+val weak_use_thy     = gen_use_thy' (require_thy true true false false false);
+val quiet_update_thy = gen_use_thy' (require_thy true true false true true);
 
 val use_thy          = warn_finished (gen_use_thy (require_thy true true false true false));
 val time_use_thy     = warn_finished (gen_use_thy (require_thy true true true true false));
-val use_thy_only     = warn_finished (gen_use_thy (require_thy true false false true false));
 val update_thy       = warn_finished (gen_use_thy (require_thy true true false true true));
-val update_thy_only  = warn_finished (gen_use_thy (require_thy true false false true true));
 val pretend_use_thy_only = warn_finished (gen_use_thy (require_thy false false false true false));
 
 end;
@@ -429,7 +420,7 @@
 (* begin / end theory *)
 
 fun check_uses name uses =
-  let val illegal = map (fn ext => Path.ext ext (Path.basic name)) ThyLoad.ml_exts in
+  let val illegal = map (fn ext => Path.ext ext (Path.basic name)) ("" :: ThyLoad.ml_exts) in
     (case find_first (member (op =) illegal o Path.base o Path.expand o #1) uses of
       NONE => ()
     | SOME (path, _) => error ("Illegal use of theory ML file: " ^ quote (Path.implode path)))
@@ -437,22 +428,21 @@
 
 fun begin_theory present name parents uses int =
   let
-    val bparents = map base_of_path parents;
-    val dir' = opt_path'' (lookup_deps name);
-    val dir = the_default Path.current dir';
-    val assert_thy = if int then quiet_update_thy' true dir else weak_use_thy dir;
+    val bparents = map base_name parents;
+    val dirs' = master_path'' (lookup_deps name);
+    val dirs = dirs' @ [Path.current];
     val _ = check_unfinished error name;
-    val _ = List.app assert_thy parents;
+    val _ = List.app ((if int then quiet_update_thy else weak_use_thy) dirs) parents;
     val _ = check_uses name uses;
 
     val theory = Theory.begin_theory name (map get_theory bparents);
     val deps =
       if known_thy name then get_deps name
-      else (init_deps NONE (map #1 uses));   (*records additional ML files only!*)
+      else (init_deps NONE (map #1 uses));
     val uses_now = map_filter (fn (x, true) => SOME x | _ => NONE) uses;
 
     val _ = change_thys (update_node name bparents (deps, SOME (Theory.copy theory)));
-    val theory' = theory |> present dir' name bparents uses;
+    val theory' = theory |> present dirs' name bparents uses;
     val _ = put_theory name (Theory.copy theory');
     val ((), theory'') =
       ML_Context.pass_context (Context.Theory theory') (List.app (load_file false)) uses_now
@@ -481,7 +471,7 @@
     val parent_names = map Context.theory_name parents;
 
     fun err txt bads =
-      error (loader_msg txt bads ^ "\n" ^ gen_msg "cannot register theory" [name]);
+      error (loader_msg txt bads ^ "\ncannot register theory " ^ quote name);
 
     val nonfinished = filter_out is_finished parent_names;
     fun get_variant (x, y_name) =