src/Pure/Thy/thy_load.ML
author wenzelm
Fri, 14 Jan 2011 13:58:07 +0100
changeset 41548 bd0bebf93fa6
parent 41414 00b2b6716ed8
child 41886 aa8dce9ab8a9
permissions -rw-r--r--
Thy_Load.begin_theory: maintain source specification of imports; Thy_Info.register_thy: reconstruct deps using original imports (important for ProofGeneral.inform_file_processed);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
6168
9d5b74068bf9 Theory loader primitives.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Thy/thy_load.ML
9d5b74068bf9 Theory loader primitives.
wenzelm
parents:
diff changeset
     2
    Author:     Markus Wenzel, TU Muenchen
9d5b74068bf9 Theory loader primitives.
wenzelm
parents:
diff changeset
     3
37939
965537d86fcc discontinued special treatment of ML files -- no longer complete extensions on demand;
wenzelm
parents: 37938
diff changeset
     4
Loading files that contribute to a theory, including global load path
965537d86fcc discontinued special treatment of ML files -- no longer complete extensions on demand;
wenzelm
parents: 37938
diff changeset
     5
management.
6168
9d5b74068bf9 Theory loader primitives.
wenzelm
parents:
diff changeset
     6
*)
9d5b74068bf9 Theory loader primitives.
wenzelm
parents:
diff changeset
     7
9d5b74068bf9 Theory loader primitives.
wenzelm
parents:
diff changeset
     8
signature THY_LOAD =
9d5b74068bf9 Theory loader primitives.
wenzelm
parents:
diff changeset
     9
sig
40741
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    10
  eqtype file_ident
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    11
  val pretty_file_ident: file_ident -> Pretty.T
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    12
  val file_ident: Path.T -> file_ident option
37977
3ceccd415145 simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents: 37952
diff changeset
    13
  val set_master_path: Path.T -> unit
3ceccd415145 simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents: 37952
diff changeset
    14
  val get_master_path: unit -> Path.T
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
    15
  val master_directory: theory -> Path.T
41548
bd0bebf93fa6 Thy_Load.begin_theory: maintain source specification of imports;
wenzelm
parents: 41414
diff changeset
    16
  val imports_of: theory -> string list
40741
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    17
  val provide: Path.T * (Path.T * file_ident) -> theory -> theory
41414
00b2b6716ed8 theory loader: implicit load path is considered legacy;
wenzelm
parents: 41412
diff changeset
    18
  val legacy_show_path: unit -> string list
00b2b6716ed8 theory loader: implicit load path is considered legacy;
wenzelm
parents: 41412
diff changeset
    19
  val legacy_add_path: string -> unit
00b2b6716ed8 theory loader: implicit load path is considered legacy;
wenzelm
parents: 41412
diff changeset
    20
  val legacy_path_add: string -> unit
00b2b6716ed8 theory loader: implicit load path is considered legacy;
wenzelm
parents: 41412
diff changeset
    21
  val legacy_del_path: string -> unit
00b2b6716ed8 theory loader: implicit load path is considered legacy;
wenzelm
parents: 41412
diff changeset
    22
  val legacy_reset_path: unit -> unit
40741
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    23
  val check_file: Path.T list -> Path.T -> Path.T * file_ident
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    24
  val check_thy: Path.T -> string -> Path.T * file_ident
24189
1fa9852643a3 simplified ThyLoad.deps_thy etc.: discontinued attached ML files;
wenzelm
parents: 24065
diff changeset
    25
  val deps_thy: Path.T -> string ->
40741
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    26
   {master: Path.T * file_ident, text: string, imports: string list, uses: Path.T list}
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
    27
  val loaded_files: theory -> Path.T list
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
    28
  val all_current: theory -> bool
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
    29
  val provide_file: Path.T -> theory -> theory
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
    30
  val use_ml: Path.T -> unit
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
    31
  val exec_ml: Path.T -> generic_theory -> generic_theory
41548
bd0bebf93fa6 Thy_Load.begin_theory: maintain source specification of imports;
wenzelm
parents: 41414
diff changeset
    32
  val begin_theory: Path.T -> string -> string list -> theory list -> (Path.T * bool) list -> theory
6168
9d5b74068bf9 Theory loader primitives.
wenzelm
parents:
diff changeset
    33
end;
9d5b74068bf9 Theory loader primitives.
wenzelm
parents:
diff changeset
    34
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 33221
diff changeset
    35
structure Thy_Load: THY_LOAD =
6168
9d5b74068bf9 Theory loader primitives.
wenzelm
parents:
diff changeset
    36
struct
9d5b74068bf9 Theory loader primitives.
wenzelm
parents:
diff changeset
    37
40741
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    38
(* file identification *)
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    39
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    40
local
40742
dc6439c0b8b1 prefer Synchronized.var over CRITICAL/Unsynchronized.ref;
wenzelm
parents: 40741
diff changeset
    41
  val file_ident_cache =
dc6439c0b8b1 prefer Synchronized.var over CRITICAL/Unsynchronized.ref;
wenzelm
parents: 40741
diff changeset
    42
    Synchronized.var "file_ident_cache"
dc6439c0b8b1 prefer Synchronized.var over CRITICAL/Unsynchronized.ref;
wenzelm
parents: 40741
diff changeset
    43
      (Symtab.empty: {time_stamp: string, id: string} Symtab.table);
40741
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    44
in
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    45
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    46
fun check_cache (path, time) =
40742
dc6439c0b8b1 prefer Synchronized.var over CRITICAL/Unsynchronized.ref;
wenzelm
parents: 40741
diff changeset
    47
  (case Symtab.lookup (Synchronized.value file_ident_cache) path of
40741
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    48
    NONE => NONE
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    49
  | SOME {time_stamp, id} => if time_stamp = time then SOME id else NONE);
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    50
40742
dc6439c0b8b1 prefer Synchronized.var over CRITICAL/Unsynchronized.ref;
wenzelm
parents: 40741
diff changeset
    51
fun update_cache (path, (time, id)) =
dc6439c0b8b1 prefer Synchronized.var over CRITICAL/Unsynchronized.ref;
wenzelm
parents: 40741
diff changeset
    52
  Synchronized.change file_ident_cache
dc6439c0b8b1 prefer Synchronized.var over CRITICAL/Unsynchronized.ref;
wenzelm
parents: 40741
diff changeset
    53
    (Symtab.update (path, {time_stamp = time, id = id}));
40741
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    54
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    55
end;
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    56
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    57
datatype file_ident = File_Ident of string;
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    58
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    59
fun pretty_file_ident (File_Ident s) = Pretty.str (quote s);
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    60
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    61
fun file_ident path =
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    62
  let val physical_path = perhaps (try OS.FileSys.fullPath) (File.platform_path path) in
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    63
    (case try (Time.toString o OS.FileSys.modTime) physical_path of
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    64
      NONE => NONE
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    65
    | SOME mod_time => SOME (File_Ident
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    66
        (case getenv "ISABELLE_FILE_IDENT" of
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    67
          "" => physical_path ^ ": " ^ mod_time
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    68
        | cmd =>
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    69
            (case check_cache (physical_path, mod_time) of
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    70
              SOME id => id
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    71
            | NONE =>
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    72
                let
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    73
                  val (id, rc) =  (*potentially slow*)
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    74
                    bash_output
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    75
                      ("\"$ISABELLE_HOME/lib/scripts/fileident\" " ^ File.shell_quote physical_path);
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    76
                in
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    77
                  if id <> "" andalso rc = 0 then (update_cache (physical_path, (mod_time, id)); id)
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    78
                  else error ("Failed to identify file " ^ quote (Path.implode path) ^ " by " ^ cmd)
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    79
                end))))
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    80
  end;
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    81
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    82
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
    83
(* manage source files *)
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
    84
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
    85
type files =
40741
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    86
 {master_dir: Path.T,  (*master directory of theory source*)
41548
bd0bebf93fa6 Thy_Load.begin_theory: maintain source specification of imports;
wenzelm
parents: 41414
diff changeset
    87
  imports: string list,  (*source specification of imports*)
40741
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    88
  required: Path.T list,  (*source path*)
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
    89
  provided: (Path.T * (Path.T * file_ident)) list};  (*source path, physical path, identifier*)
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
    90
41548
bd0bebf93fa6 Thy_Load.begin_theory: maintain source specification of imports;
wenzelm
parents: 41414
diff changeset
    91
fun make_files (master_dir, imports, required, provided): files =
bd0bebf93fa6 Thy_Load.begin_theory: maintain source specification of imports;
wenzelm
parents: 41414
diff changeset
    92
 {master_dir = master_dir, imports = imports, required = required, provided = provided};
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
    93
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
    94
structure Files = Theory_Data
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
    95
(
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
    96
  type T = files;
41548
bd0bebf93fa6 Thy_Load.begin_theory: maintain source specification of imports;
wenzelm
parents: 41414
diff changeset
    97
  val empty = make_files (Path.current, [], [], []);
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
    98
  fun extend _ = empty;
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
    99
  fun merge _ = empty;
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   100
);
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   101
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   102
fun map_files f =
41548
bd0bebf93fa6 Thy_Load.begin_theory: maintain source specification of imports;
wenzelm
parents: 41414
diff changeset
   103
  Files.map (fn {master_dir, imports, required, provided} =>
bd0bebf93fa6 Thy_Load.begin_theory: maintain source specification of imports;
wenzelm
parents: 41414
diff changeset
   104
    make_files (f (master_dir, imports, required, provided)));
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   105
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   106
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   107
val master_directory = #master_dir o Files.get;
41548
bd0bebf93fa6 Thy_Load.begin_theory: maintain source specification of imports;
wenzelm
parents: 41414
diff changeset
   108
val imports_of = #imports o Files.get;
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   109
41548
bd0bebf93fa6 Thy_Load.begin_theory: maintain source specification of imports;
wenzelm
parents: 41414
diff changeset
   110
fun put_deps dir imports = map_files (fn _ => (dir, imports, [], []));
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   111
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   112
fun require src_path =
41548
bd0bebf93fa6 Thy_Load.begin_theory: maintain source specification of imports;
wenzelm
parents: 41414
diff changeset
   113
  map_files (fn (master_dir, imports, required, provided) =>
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   114
    if member (op =) required src_path then
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   115
      error ("Duplicate source file dependency: " ^ Path.implode src_path)
41548
bd0bebf93fa6 Thy_Load.begin_theory: maintain source specification of imports;
wenzelm
parents: 41414
diff changeset
   116
    else (master_dir, imports, src_path :: required, provided));
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   117
37978
548f3f165d05 simplified Thy_Header.read -- include Source.of_string_limited here;
wenzelm
parents: 37977
diff changeset
   118
fun provide (src_path, path_id) =
41548
bd0bebf93fa6 Thy_Load.begin_theory: maintain source specification of imports;
wenzelm
parents: 41414
diff changeset
   119
  map_files (fn (master_dir, imports, required, provided) =>
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   120
    if AList.defined (op =) provided src_path then
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   121
      error ("Duplicate resolution of source file dependency: " ^ Path.implode src_path)
41548
bd0bebf93fa6 Thy_Load.begin_theory: maintain source specification of imports;
wenzelm
parents: 41414
diff changeset
   122
    else (master_dir, imports, required, (src_path, path_id) :: provided));
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   123
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   124
37977
3ceccd415145 simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents: 37952
diff changeset
   125
(* maintain default paths *)
6168
9d5b74068bf9 Theory loader primitives.
wenzelm
parents:
diff changeset
   126
37977
3ceccd415145 simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents: 37952
diff changeset
   127
local
40742
dc6439c0b8b1 prefer Synchronized.var over CRITICAL/Unsynchronized.ref;
wenzelm
parents: 40741
diff changeset
   128
  val load_path = Synchronized.var "load_path" [Path.current];
37977
3ceccd415145 simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents: 37952
diff changeset
   129
  val master_path = Unsynchronized.ref Path.current;
3ceccd415145 simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents: 37952
diff changeset
   130
in
6168
9d5b74068bf9 Theory loader primitives.
wenzelm
parents:
diff changeset
   131
41414
00b2b6716ed8 theory loader: implicit load path is considered legacy;
wenzelm
parents: 41412
diff changeset
   132
fun legacy_show_path () = map Path.implode (Synchronized.value load_path);
23944
2ea068548a83 marked some CRITICAL sections;
wenzelm
parents: 23892
diff changeset
   133
41414
00b2b6716ed8 theory loader: implicit load path is considered legacy;
wenzelm
parents: 41412
diff changeset
   134
fun legacy_del_path s = Synchronized.change load_path (remove (op =) (Path.explode s));
33221
wenzelm
parents: 32966
diff changeset
   135
41414
00b2b6716ed8 theory loader: implicit load path is considered legacy;
wenzelm
parents: 41412
diff changeset
   136
fun legacy_add_path s = Synchronized.change load_path (update (op =) (Path.explode s));
33221
wenzelm
parents: 32966
diff changeset
   137
41414
00b2b6716ed8 theory loader: implicit load path is considered legacy;
wenzelm
parents: 41412
diff changeset
   138
fun legacy_path_add s =
40742
dc6439c0b8b1 prefer Synchronized.var over CRITICAL/Unsynchronized.ref;
wenzelm
parents: 40741
diff changeset
   139
  Synchronized.change load_path (fn path =>
dc6439c0b8b1 prefer Synchronized.var over CRITICAL/Unsynchronized.ref;
wenzelm
parents: 40741
diff changeset
   140
    let val p = Path.explode s
dc6439c0b8b1 prefer Synchronized.var over CRITICAL/Unsynchronized.ref;
wenzelm
parents: 40741
diff changeset
   141
    in remove (op =) p path @ [p] end);
33221
wenzelm
parents: 32966
diff changeset
   142
41414
00b2b6716ed8 theory loader: implicit load path is considered legacy;
wenzelm
parents: 41412
diff changeset
   143
fun legacy_reset_path () = Synchronized.change load_path (K [Path.current]);
6168
9d5b74068bf9 Theory loader primitives.
wenzelm
parents:
diff changeset
   144
23944
2ea068548a83 marked some CRITICAL sections;
wenzelm
parents: 23892
diff changeset
   145
fun search_path dir path =
40742
dc6439c0b8b1 prefer Synchronized.var over CRITICAL/Unsynchronized.ref;
wenzelm
parents: 40741
diff changeset
   146
  distinct (op =)
dc6439c0b8b1 prefer Synchronized.var over CRITICAL/Unsynchronized.ref;
wenzelm
parents: 40741
diff changeset
   147
    (dir :: (if Path.is_basic path then (Synchronized.value load_path) else [Path.current]));
23944
2ea068548a83 marked some CRITICAL sections;
wenzelm
parents: 23892
diff changeset
   148
37977
3ceccd415145 simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents: 37952
diff changeset
   149
fun set_master_path path = master_path := path;
3ceccd415145 simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents: 37952
diff changeset
   150
fun get_master_path () = ! master_path;
3ceccd415145 simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents: 37952
diff changeset
   151
23944
2ea068548a83 marked some CRITICAL sections;
wenzelm
parents: 23892
diff changeset
   152
end;
2ea068548a83 marked some CRITICAL sections;
wenzelm
parents: 23892
diff changeset
   153
6168
9d5b74068bf9 Theory loader primitives.
wenzelm
parents:
diff changeset
   154
23872
f449381e2caa tuned signature;
wenzelm
parents: 22145
diff changeset
   155
(* check files *)
6232
4336add1c251 include full paths in file info;
wenzelm
parents: 6219
diff changeset
   156
38133
987680d2e77d simplified/clarified Thy_Load path: search for master only, lookup other files relative to that;
wenzelm
parents: 37978
diff changeset
   157
fun get_file dirs src_path =
6232
4336add1c251 include full paths in file info;
wenzelm
parents: 6219
diff changeset
   158
  let
4336add1c251 include full paths in file info;
wenzelm
parents: 6219
diff changeset
   159
    val path = Path.expand src_path;
23872
f449381e2caa tuned signature;
wenzelm
parents: 22145
diff changeset
   160
    val _ = Path.is_current path andalso error "Bad file specification";
37939
965537d86fcc discontinued special treatment of ML files -- no longer complete extensions on demand;
wenzelm
parents: 37938
diff changeset
   161
  in
38133
987680d2e77d simplified/clarified Thy_Load path: search for master only, lookup other files relative to that;
wenzelm
parents: 37978
diff changeset
   162
    dirs |> get_first (fn dir =>
987680d2e77d simplified/clarified Thy_Load path: search for master only, lookup other files relative to that;
wenzelm
parents: 37978
diff changeset
   163
      let val full_path = File.full_path (Path.append dir path) in
40741
17d6293a1e26 moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents: 40625
diff changeset
   164
        (case file_ident full_path of
38133
987680d2e77d simplified/clarified Thy_Load path: search for master only, lookup other files relative to that;
wenzelm
parents: 37978
diff changeset
   165
          NONE => NONE
41412
35f30e07fe0d check_file: secondary load path is legacy feature;
wenzelm
parents: 40742
diff changeset
   166
        | SOME id => SOME (dir, (full_path, id)))
38133
987680d2e77d simplified/clarified Thy_Load path: search for master only, lookup other files relative to that;
wenzelm
parents: 37978
diff changeset
   167
      end)
37939
965537d86fcc discontinued special treatment of ML files -- no longer complete extensions on demand;
wenzelm
parents: 37938
diff changeset
   168
  end;
6168
9d5b74068bf9 Theory loader primitives.
wenzelm
parents:
diff changeset
   169
38133
987680d2e77d simplified/clarified Thy_Load path: search for master only, lookup other files relative to that;
wenzelm
parents: 37978
diff changeset
   170
fun check_file dirs file =
987680d2e77d simplified/clarified Thy_Load path: search for master only, lookup other files relative to that;
wenzelm
parents: 37978
diff changeset
   171
  (case get_file dirs file of
41412
35f30e07fe0d check_file: secondary load path is legacy feature;
wenzelm
parents: 40742
diff changeset
   172
    SOME (_, path_id) =>
35f30e07fe0d check_file: secondary load path is legacy feature;
wenzelm
parents: 40742
diff changeset
   173
     (if is_some (get_file [hd dirs] file) then ()
35f30e07fe0d check_file: secondary load path is legacy feature;
wenzelm
parents: 40742
diff changeset
   174
      else
35f30e07fe0d check_file: secondary load path is legacy feature;
wenzelm
parents: 40742
diff changeset
   175
        legacy_feature ("File " ^ quote (Path.implode file) ^
35f30e07fe0d check_file: secondary load path is legacy feature;
wenzelm
parents: 40742
diff changeset
   176
          " located via secondary search path: " ^
35f30e07fe0d check_file: secondary load path is legacy feature;
wenzelm
parents: 40742
diff changeset
   177
          quote (Path.implode (#1 (the (get_file (tl dirs) file)))));
35f30e07fe0d check_file: secondary load path is legacy feature;
wenzelm
parents: 40742
diff changeset
   178
      path_id)
38133
987680d2e77d simplified/clarified Thy_Load path: search for master only, lookup other files relative to that;
wenzelm
parents: 37978
diff changeset
   179
  | NONE => error ("Could not find file " ^ quote (Path.implode file) ^
987680d2e77d simplified/clarified Thy_Load path: search for master only, lookup other files relative to that;
wenzelm
parents: 37978
diff changeset
   180
      "\nin " ^ commas_quote (map Path.implode dirs)));
37939
965537d86fcc discontinued special treatment of ML files -- no longer complete extensions on demand;
wenzelm
parents: 37938
diff changeset
   181
38133
987680d2e77d simplified/clarified Thy_Load path: search for master only, lookup other files relative to that;
wenzelm
parents: 37978
diff changeset
   182
fun check_thy master_dir name =
987680d2e77d simplified/clarified Thy_Load path: search for master only, lookup other files relative to that;
wenzelm
parents: 37978
diff changeset
   183
  let
987680d2e77d simplified/clarified Thy_Load path: search for master only, lookup other files relative to that;
wenzelm
parents: 37978
diff changeset
   184
    val thy_file = Thy_Header.thy_path name;
987680d2e77d simplified/clarified Thy_Load path: search for master only, lookup other files relative to that;
wenzelm
parents: 37978
diff changeset
   185
    val dirs = search_path master_dir thy_file;
987680d2e77d simplified/clarified Thy_Load path: search for master only, lookup other files relative to that;
wenzelm
parents: 37978
diff changeset
   186
  in check_file dirs thy_file end;
6232
4336add1c251 include full paths in file info;
wenzelm
parents: 6219
diff changeset
   187
4336add1c251 include full paths in file info;
wenzelm
parents: 6219
diff changeset
   188
23872
f449381e2caa tuned signature;
wenzelm
parents: 22145
diff changeset
   189
(* theory deps *)
7940
def6db239934 improved ml handling;
wenzelm
parents: 7901
diff changeset
   190
38133
987680d2e77d simplified/clarified Thy_Load path: search for master only, lookup other files relative to that;
wenzelm
parents: 37978
diff changeset
   191
fun deps_thy master_dir name =
23872
f449381e2caa tuned signature;
wenzelm
parents: 22145
diff changeset
   192
  let
38133
987680d2e77d simplified/clarified Thy_Load path: search for master only, lookup other files relative to that;
wenzelm
parents: 37978
diff changeset
   193
    val master as (thy_path, _) = check_thy master_dir name;
37978
548f3f165d05 simplified Thy_Header.read -- include Source.of_string_limited here;
wenzelm
parents: 37977
diff changeset
   194
    val text = File.read thy_path;
548f3f165d05 simplified Thy_Header.read -- include Source.of_string_limited here;
wenzelm
parents: 37977
diff changeset
   195
    val (name', imports, uses) = Thy_Header.read (Path.position thy_path) text;
37950
bc285d91041e moved basic thy file name operations from Thy_Load to Thy_Header;
wenzelm
parents: 37949
diff changeset
   196
    val _ = Thy_Header.consistent_name name name';
24189
1fa9852643a3 simplified ThyLoad.deps_thy etc.: discontinued attached ML files;
wenzelm
parents: 24065
diff changeset
   197
    val uses = map (Path.explode o #1) uses;
24065
21483400c2ca load_thy: avoid reloading of text;
wenzelm
parents: 23944
diff changeset
   198
  in {master = master, text = text, imports = imports, uses = uses} end;
6484
3f098b0ec683 use_thy etc.: may specify path prefix, which is temporarily used as load path;
wenzelm
parents: 6363
diff changeset
   199
6168
9d5b74068bf9 Theory loader primitives.
wenzelm
parents:
diff changeset
   200
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   201
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   202
(* loaded files *)
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   203
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   204
val loaded_files = map (#1 o #2) o #provided o Files.get;
6168
9d5b74068bf9 Theory loader primitives.
wenzelm
parents:
diff changeset
   205
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   206
fun check_loaded thy =
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   207
  let
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   208
    val {required, provided, ...} = Files.get thy;
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   209
    val provided_paths = map #1 provided;
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   210
    val _ =
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   211
      (case subtract (op =) provided_paths required of
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   212
        [] => NONE
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   213
      | bad => error ("Pending source file dependencies: " ^ commas (map Path.implode (rev bad))));
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   214
    val _ =
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   215
      (case subtract (op =) required provided_paths of
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   216
        [] => NONE
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   217
      | bad => error ("Undeclared source file dependencies: " ^ commas (map Path.implode (rev bad))));
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   218
  in () end;
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   219
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   220
fun all_current thy =
37939
965537d86fcc discontinued special treatment of ML files -- no longer complete extensions on demand;
wenzelm
parents: 37938
diff changeset
   221
  let
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   222
    val {master_dir, provided, ...} = Files.get thy;
37978
548f3f165d05 simplified Thy_Header.read -- include Source.of_string_limited here;
wenzelm
parents: 37977
diff changeset
   223
    fun current (src_path, (_, id)) =
38133
987680d2e77d simplified/clarified Thy_Load path: search for master only, lookup other files relative to that;
wenzelm
parents: 37978
diff changeset
   224
      (case get_file [master_dir] src_path of
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   225
        NONE => false
41412
35f30e07fe0d check_file: secondary load path is legacy feature;
wenzelm
parents: 40742
diff changeset
   226
      | SOME (_, (_, id')) => id = id');
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   227
  in can check_loaded thy andalso forall current provided end;
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   228
37952
f6c40213675b Thy_Load.check_loaded via Theory.at_end;
wenzelm
parents: 37950
diff changeset
   229
val _ = Context.>> (Context.map_theory (Theory.at_end (fn thy => (check_loaded thy; NONE))));
f6c40213675b Thy_Load.check_loaded via Theory.at_end;
wenzelm
parents: 37950
diff changeset
   230
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   231
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   232
(* provide files *)
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   233
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   234
fun provide_file src_path thy =
38133
987680d2e77d simplified/clarified Thy_Load path: search for master only, lookup other files relative to that;
wenzelm
parents: 37978
diff changeset
   235
  provide (src_path, check_file [master_directory thy] src_path) thy;
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   236
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   237
fun use_ml src_path =
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   238
  if is_none (Context.thread_data ()) then
38133
987680d2e77d simplified/clarified Thy_Load path: search for master only, lookup other files relative to that;
wenzelm
parents: 37978
diff changeset
   239
    ML_Context.eval_file (#1 (check_file [Path.current] src_path))
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   240
  else
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   241
    let
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   242
      val thy = ML_Context.the_global_context ();
38133
987680d2e77d simplified/clarified Thy_Load path: search for master only, lookup other files relative to that;
wenzelm
parents: 37978
diff changeset
   243
      val (path, id) = check_file [master_directory thy] src_path;
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   244
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   245
      val _ = ML_Context.eval_file path;
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   246
      val _ = Context.>> Local_Theory.propagate_ml_env;
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   247
37978
548f3f165d05 simplified Thy_Header.read -- include Source.of_string_limited here;
wenzelm
parents: 37977
diff changeset
   248
      val provide = provide (src_path, (path, id));
38757
2b3e054ae6fc renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
wenzelm
parents: 38133
diff changeset
   249
      val _ = Context.>> (Context.mapping provide (Local_Theory.background_theory provide));
37978
548f3f165d05 simplified Thy_Header.read -- include Source.of_string_limited here;
wenzelm
parents: 37977
diff changeset
   250
    in () end;
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   251
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 37943
diff changeset
   252
fun exec_ml src_path = ML_Context.exec (fn () => use_ml src_path);
6168
9d5b74068bf9 Theory loader primitives.
wenzelm
parents:
diff changeset
   253
37977
3ceccd415145 simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents: 37952
diff changeset
   254
3ceccd415145 simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents: 37952
diff changeset
   255
(* begin theory *)
3ceccd415145 simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents: 37952
diff changeset
   256
41548
bd0bebf93fa6 Thy_Load.begin_theory: maintain source specification of imports;
wenzelm
parents: 41414
diff changeset
   257
fun begin_theory dir name imports parents uses =
37977
3ceccd415145 simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents: 37952
diff changeset
   258
  Theory.begin_theory name parents
41548
bd0bebf93fa6 Thy_Load.begin_theory: maintain source specification of imports;
wenzelm
parents: 41414
diff changeset
   259
  |> put_deps dir imports
37977
3ceccd415145 simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents: 37952
diff changeset
   260
  |> fold (require o fst) uses
3ceccd415145 simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents: 37952
diff changeset
   261
  |> fold (fn (path, true) => Context.theory_map (exec_ml path) o Theory.checkpoint | _ => I) uses
3ceccd415145 simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents: 37952
diff changeset
   262
  |> Theory.checkpoint;
3ceccd415145 simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents: 37952
diff changeset
   263
6168
9d5b74068bf9 Theory loader primitives.
wenzelm
parents:
diff changeset
   264
end;
9d5b74068bf9 Theory loader primitives.
wenzelm
parents:
diff changeset
   265