| author | paulson | 
| Thu, 05 Dec 2013 23:13:54 +0000 | |
| changeset 54679 | 88adcd3b34e8 | 
| parent 54559 | 39d91cac6e91 | 
| child 54722 | 5f5608bfe230 | 
| permissions | -rw-r--r-- | 
| 3604 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 1 | (* Title: Pure/Thy/thy_info.ML | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 2 | Author: Markus Wenzel, TU Muenchen | 
| 3604 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 3 | |
| 37978 
548f3f165d05
simplified Thy_Header.read -- include Source.of_string_limited here;
 wenzelm parents: 
37977diff
changeset | 4 | Global theory info database, with auto-loading according to theory and | 
| 15801 | 5 | file dependencies. | 
| 3976 | 6 | *) | 
| 3604 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 7 | |
| 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 8 | signature THY_INFO = | 
| 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 9 | sig | 
| 37981 
9a15982f41fe
theory loader: removed obsolete touch/outdate operations (require_thy no longer changes the database implicitly);
 wenzelm parents: 
37980diff
changeset | 10 | datatype action = Update | Remove | 
| 7099 | 11 | val add_hook: (action -> string -> unit) -> unit | 
| 26614 | 12 | val get_names: unit -> string list | 
| 7288 | 13 | val lookup_theory: string -> theory option | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 14 | val get_theory: string -> theory | 
| 24563 | 15 | val is_finished: string -> bool | 
| 26983 | 16 | val master_directory: string -> Path.T | 
| 24080 
8c67d869531b
added register_thy (replaces pretend_use_thy_only and really flag);
 wenzelm parents: 
24071diff
changeset | 17 | val loaded_files: string -> Path.T list | 
| 29421 
db532e37cda2
use_thys: perform consolidate_thy on loaded theories, which removes failed nodes in post-hoc fashion;
 wenzelm parents: 
29118diff
changeset | 18 | val remove_thy: string -> unit | 
| 
db532e37cda2
use_thys: perform consolidate_thy on loaded theories, which removes failed nodes in post-hoc fashion;
 wenzelm parents: 
29118diff
changeset | 19 | val kill_thy: string -> unit | 
| 54458 
96ccc8972fc7
prefer explicit "document" flag -- eliminated stateful Present.no_document;
 wenzelm parents: 
54446diff
changeset | 20 | val use_theories: | 
| 
96ccc8972fc7
prefer explicit "document" flag -- eliminated stateful Present.no_document;
 wenzelm parents: 
54446diff
changeset | 21 |     {document: bool, last_timing: Toplevel.transition -> Time.time option, master_dir: Path.T} ->
 | 
| 51217 
65ab2c4f4c32
support for prescient timing information within command transactions;
 wenzelm parents: 
50862diff
changeset | 22 | (string * Position.T) list -> unit | 
| 48927 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 wenzelm parents: 
48898diff
changeset | 23 | val use_thys: (string * Position.T) list -> unit | 
| 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 wenzelm parents: 
48898diff
changeset | 24 | val use_thy: string * Position.T -> unit | 
| 46938 
cda018294515
some support for outer syntax keyword declarations within theory header;
 wenzelm parents: 
46774diff
changeset | 25 | val toplevel_begin_theory: Path.T -> Thy_Header.header -> theory | 
| 37954 
a2e73df0b1e0
simplified/clarified register_thy: more precise treatment of new dependencies, remove descendants;
 wenzelm parents: 
37953diff
changeset | 26 | val register_thy: theory -> unit | 
| 24080 
8c67d869531b
added register_thy (replaces pretend_use_thy_only and really flag);
 wenzelm parents: 
24071diff
changeset | 27 | val finish: unit -> unit | 
| 3604 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 28 | end; | 
| 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 29 | |
| 37216 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 wenzelm parents: 
36953diff
changeset | 30 | structure Thy_Info: THY_INFO = | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 31 | struct | 
| 3604 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 32 | |
| 7099 | 33 | (** theory loader actions and hooks **) | 
| 34 | ||
| 37981 
9a15982f41fe
theory loader: removed obsolete touch/outdate operations (require_thy no longer changes the database implicitly);
 wenzelm parents: 
37980diff
changeset | 35 | datatype action = Update | Remove; | 
| 7099 | 36 | |
| 37 | local | |
| 43640 | 38 | val hooks = Synchronized.var "Thy_Info.hooks" ([]: (action -> string -> unit) list); | 
| 7099 | 39 | in | 
| 43640 | 40 | fun add_hook f = Synchronized.change hooks (cons f); | 
| 41 | fun perform action name = | |
| 42 | List.app (fn f => (try (fn () => f action name) (); ())) (Synchronized.value hooks); | |
| 7099 | 43 | end; | 
| 44 | ||
| 45 | ||
| 46 | ||
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 47 | (** thy database **) | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 48 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 49 | (* messages *) | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 50 | |
| 23871 | 51 | fun loader_msg txt [] = "Theory loader: " ^ txt | 
| 52 | | loader_msg txt names = "Theory loader: " ^ txt ^ " " ^ commas_quote names; | |
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 53 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 54 | val show_path = space_implode " via " o map quote; | 
| 9332 | 55 | fun cycle_msg names = loader_msg ("cyclic dependency of " ^ show_path names) [];
 | 
| 3604 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 56 | |
| 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 57 | |
| 6666 | 58 | (* derived graph operations *) | 
| 3604 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 59 | |
| 50862 
5fc8b83322f5
more sensible order of theory nodes (correspondance to Scala version), e.g. relevant to theory progress;
 wenzelm parents: 
50845diff
changeset | 60 | fun add_deps name parents G = String_Graph.add_deps_acyclic (name, parents) G | 
| 
5fc8b83322f5
more sensible order of theory nodes (correspondance to Scala version), e.g. relevant to theory progress;
 wenzelm parents: 
50845diff
changeset | 61 | handle String_Graph.CYCLES namess => error (cat_lines (map cycle_msg namess)); | 
| 3604 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 62 | |
| 37984 
f26352bd82cf
clarified register_thy: clean slate via kill_thy, more precise CRITICAL section;
 wenzelm parents: 
37981diff
changeset | 63 | fun new_entry name parents entry = | 
| 50862 
5fc8b83322f5
more sensible order of theory nodes (correspondance to Scala version), e.g. relevant to theory progress;
 wenzelm parents: 
50845diff
changeset | 64 | String_Graph.new_node (name, entry) #> add_deps name parents; | 
| 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: 
37955diff
changeset | 65 | |
| 3604 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 66 | |
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 67 | (* thy database *) | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 68 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 69 | type deps = | 
| 41955 
703ea96b13c6
files are identified via SHA1 digests -- discontinued ISABELLE_FILE_IDENT;
 wenzelm parents: 
41678diff
changeset | 70 |  {master: (Path.T * SHA1.digest),  (*master dependencies for thy file*)
 | 
| 48927 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 wenzelm parents: 
48898diff
changeset | 71 | imports: (string * Position.T) list}; (*source specification of imports (partially qualified)*) | 
| 23886 
f40fba467384
simplified ThyLoad interfaces: only one additional directory;
 wenzelm parents: 
23871diff
changeset | 72 | |
| 38148 
d2f3c8d4a89f
uniform naming of imports (source specification) vs. parents (thy node names) vs. parent_thys (theory values);
 wenzelm parents: 
38147diff
changeset | 73 | fun make_deps master imports : deps = {master = master, imports = imports};
 | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 74 | |
| 37980 
b8c3d4dc1e3e
avoid repeated File.read for theory text (as before);
 wenzelm parents: 
37979diff
changeset | 75 | fun master_dir (d: deps option) = the_default Path.current (Option.map (Path.dir o #1 o #master) d); | 
| 44225 
a8f921e6484f
more robust Thy_Header.base_name, with minimal assumptions about path syntax;
 wenzelm parents: 
44222diff
changeset | 76 | fun base_name s = Path.implode (Path.base (Path.explode s)); | 
| 23967 | 77 | |
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 78 | local | 
| 50862 
5fc8b83322f5
more sensible order of theory nodes (correspondance to Scala version), e.g. relevant to theory progress;
 wenzelm parents: 
50845diff
changeset | 79 | val database = | 
| 
5fc8b83322f5
more sensible order of theory nodes (correspondance to Scala version), e.g. relevant to theory progress;
 wenzelm parents: 
50845diff
changeset | 80 | Unsynchronized.ref (String_Graph.empty: (deps option * theory option) String_Graph.T); | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 81 | in | 
| 6362 | 82 | fun get_thys () = ! database; | 
| 38145 
675827e61eb9
more precise CRITICAL sections, using NAMED_CRITICAL uniformly;
 wenzelm parents: 
38143diff
changeset | 83 | fun change_thys f = NAMED_CRITICAL "Thy_Info" (fn () => Unsynchronized.change database f); | 
| 3604 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 84 | end; | 
| 5209 | 85 | |
| 86 | ||
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 87 | (* access thy graph *) | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 88 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 89 | fun thy_graph f x = f (get_thys ()) x; | 
| 9417 | 90 | |
| 50862 
5fc8b83322f5
more sensible order of theory nodes (correspondance to Scala version), e.g. relevant to theory progress;
 wenzelm parents: 
50845diff
changeset | 91 | fun get_names () = String_Graph.topological_order (get_thys ()); | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 92 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 93 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 94 | (* access thy *) | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 95 | |
| 7935 | 96 | fun lookup_thy name = | 
| 50862 
5fc8b83322f5
more sensible order of theory nodes (correspondance to Scala version), e.g. relevant to theory progress;
 wenzelm parents: 
50845diff
changeset | 97 | SOME (thy_graph String_Graph.get_node name) handle String_Graph.UNDEF _ => NONE; | 
| 7935 | 98 | |
| 16047 | 99 | val known_thy = is_some o lookup_thy; | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 100 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 101 | fun get_thy name = | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 102 | (case lookup_thy name of | 
| 15531 | 103 | SOME thy => thy | 
| 104 | | NONE => error (loader_msg "nothing known about theory" [name])); | |
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 105 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 106 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 107 | (* access deps *) | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 108 | |
| 15570 | 109 | val lookup_deps = Option.map #1 o lookup_thy; | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 110 | val get_deps = #1 o get_thy; | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 111 | |
| 26983 | 112 | val is_finished = is_none o get_deps; | 
| 37980 
b8c3d4dc1e3e
avoid repeated File.read for theory text (as before);
 wenzelm parents: 
37979diff
changeset | 113 | val master_directory = master_dir o get_deps; | 
| 7191 | 114 | |
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 115 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 116 | (* access theory *) | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 117 | |
| 7687 | 118 | fun lookup_theory name = | 
| 119 | (case lookup_thy name of | |
| 38142 | 120 | SOME (_, SOME theory) => SOME theory | 
| 15531 | 121 | | _ => NONE); | 
| 7288 | 122 | |
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 123 | fun get_theory name = | 
| 7288 | 124 | (case lookup_theory name of | 
| 23871 | 125 | SOME theory => theory | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 126 | | _ => error (loader_msg "undefined theory entry for" [name])); | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 127 | |
| 44337 
d453faed4815
clarified get_imports -- should not rely on accidental order within graph;
 wenzelm parents: 
44302diff
changeset | 128 | val get_imports = Thy_Load.imports_of o get_theory; | 
| 
d453faed4815
clarified get_imports -- should not rely on accidental order within graph;
 wenzelm parents: 
44302diff
changeset | 129 | |
| 54446 | 130 | (*Proof General legacy*) | 
| 38145 
675827e61eb9
more precise CRITICAL sections, using NAMED_CRITICAL uniformly;
 wenzelm parents: 
38143diff
changeset | 131 | fun loaded_files name = NAMED_CRITICAL "Thy_Info" (fn () => | 
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37942diff
changeset | 132 | (case get_deps name of | 
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37942diff
changeset | 133 | NONE => [] | 
| 38145 
675827e61eb9
more precise CRITICAL sections, using NAMED_CRITICAL uniformly;
 wenzelm parents: 
38143diff
changeset | 134 |   | SOME {master = (thy_path, _), ...} => thy_path :: Thy_Load.loaded_files (get_theory name)));
 | 
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37942diff
changeset | 135 | |
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 136 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 137 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 138 | (** thy operations **) | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 139 | |
| 43640 | 140 | (* main loader actions *) | 
| 29421 
db532e37cda2
use_thys: perform consolidate_thy on loaded theories, which removes failed nodes in post-hoc fashion;
 wenzelm parents: 
29118diff
changeset | 141 | |
| 38145 
675827e61eb9
more precise CRITICAL sections, using NAMED_CRITICAL uniformly;
 wenzelm parents: 
38143diff
changeset | 142 | fun remove_thy name = NAMED_CRITICAL "Thy_Info" (fn () => | 
| 54559 | 143 | if is_finished name then error (loader_msg "cannot update finished theory" [name]) | 
| 29421 
db532e37cda2
use_thys: perform consolidate_thy on loaded theories, which removes failed nodes in post-hoc fashion;
 wenzelm parents: 
29118diff
changeset | 144 | else | 
| 29434 | 145 | let | 
| 50862 
5fc8b83322f5
more sensible order of theory nodes (correspondance to Scala version), e.g. relevant to theory progress;
 wenzelm parents: 
50845diff
changeset | 146 | val succs = thy_graph String_Graph.all_succs [name]; | 
| 40132 
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
 wenzelm parents: 
39557diff
changeset | 147 | val _ = Output.urgent_message (loader_msg "removing" succs); | 
| 38142 | 148 | val _ = List.app (perform Remove) succs; | 
| 50862 
5fc8b83322f5
more sensible order of theory nodes (correspondance to Scala version), e.g. relevant to theory progress;
 wenzelm parents: 
50845diff
changeset | 149 | val _ = change_thys (fold String_Graph.del_node succs); | 
| 38142 | 150 | in () end); | 
| 29421 
db532e37cda2
use_thys: perform consolidate_thy on loaded theories, which removes failed nodes in post-hoc fashion;
 wenzelm parents: 
29118diff
changeset | 151 | |
| 38145 
675827e61eb9
more precise CRITICAL sections, using NAMED_CRITICAL uniformly;
 wenzelm parents: 
38143diff
changeset | 152 | fun kill_thy name = NAMED_CRITICAL "Thy_Info" (fn () => | 
| 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: 
37955diff
changeset | 153 | if known_thy name then remove_thy name | 
| 38142 | 154 | else ()); | 
| 7099 | 155 | |
| 43640 | 156 | fun update_thy deps theory = NAMED_CRITICAL "Thy_Info" (fn () => | 
| 157 | let | |
| 158 | val name = Context.theory_name theory; | |
| 159 | val parents = map Context.theory_name (Theory.parents_of theory); | |
| 160 | val _ = kill_thy name; | |
| 161 | val _ = map get_theory parents; | |
| 162 | val _ = change_thys (new_entry name parents (SOME deps, SOME theory)); | |
| 163 | val _ = perform Update name; | |
| 164 | in () end); | |
| 165 | ||
| 7099 | 166 | |
| 24209 | 167 | (* scheduling loader tasks *) | 
| 168 | ||
| 51330 | 169 | datatype result = | 
| 53375 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 170 |   Result of {theory: theory, exec_id: Document_ID.exec,
 | 
| 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 171 | present: unit -> unit, commit: unit -> unit, weight: int}; | 
| 51331 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 172 | |
| 53375 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 173 | fun theory_result theory = | 
| 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 174 |   Result {theory = theory, exec_id = Document_ID.none, present = I, commit = I, weight = 0};
 | 
| 51330 | 175 | |
| 176 | fun result_theory (Result {theory, ...}) = theory;
 | |
| 51331 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 177 | fun result_present (Result {present, ...}) = present;
 | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 178 | fun result_commit (Result {commit, ...}) = commit;
 | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 179 | fun result_ord (Result {weight = i, ...}, Result {weight = j, ...}) = int_ord (j, i);
 | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 180 | |
| 53375 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 181 | fun join_theory (Result {theory, exec_id, ...}) =
 | 
| 54370 
39ac1a02c60c
join all theory body forks, notably Toplevel.atom_result (diagnostic commands), before peeking at full status;
 wenzelm parents: 
53375diff
changeset | 182 | let | 
| 
39ac1a02c60c
join all theory body forks, notably Toplevel.atom_result (diagnostic commands), before peeking at full status;
 wenzelm parents: 
53375diff
changeset | 183 | (*toplevel proofs and diags*) | 
| 
39ac1a02c60c
join all theory body forks, notably Toplevel.atom_result (diagnostic commands), before peeking at full status;
 wenzelm parents: 
53375diff
changeset | 184 | val _ = Future.join_tasks (maps Future.group_snapshot (Execution.peek exec_id)); | 
| 
39ac1a02c60c
join all theory body forks, notably Toplevel.atom_result (diagnostic commands), before peeking at full status;
 wenzelm parents: 
53375diff
changeset | 185 | (*fully nested proofs*) | 
| 
39ac1a02c60c
join all theory body forks, notably Toplevel.atom_result (diagnostic commands), before peeking at full status;
 wenzelm parents: 
53375diff
changeset | 186 | val res = Exn.capture Thm.join_theory_proofs theory; | 
| 
39ac1a02c60c
join all theory body forks, notably Toplevel.atom_result (diagnostic commands), before peeking at full status;
 wenzelm parents: 
53375diff
changeset | 187 | in res :: map Exn.Exn (maps Task_Queue.group_status (Execution.peek exec_id)) end; | 
| 43641 
081e009549dc
uniform finish_thy -- always Global_Theory.join_proofs, even with sequential scheduling;
 wenzelm parents: 
43640diff
changeset | 188 | |
| 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: 
37955diff
changeset | 189 | datatype task = | 
| 43641 
081e009549dc
uniform finish_thy -- always Global_Theory.join_proofs, even with sequential scheduling;
 wenzelm parents: 
43640diff
changeset | 190 | Task of string list * (theory list -> result) | | 
| 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: 
37955diff
changeset | 191 | Finished of 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: 
37955diff
changeset | 192 | |
| 
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: 
37955diff
changeset | 193 | fun task_finished (Task _) = false | 
| 
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: 
37955diff
changeset | 194 | | task_finished (Finished _) = true; | 
| 24209 | 195 | |
| 44162 | 196 | fun task_parents deps (parents: string list) = map (the o AList.lookup (op =) deps) parents; | 
| 197 | ||
| 29118 | 198 | local | 
| 199 | ||
| 44162 | 200 | val schedule_seq = | 
| 50862 
5fc8b83322f5
more sensible order of theory nodes (correspondance to Scala version), e.g. relevant to theory progress;
 wenzelm parents: 
50845diff
changeset | 201 | String_Graph.schedule (fn deps => fn (_, task) => | 
| 44162 | 202 | (case task of | 
| 51331 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 203 | Task (parents, body) => | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 204 | let | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 205 | val result = body (task_parents deps parents); | 
| 53190 
5d92649a310e
simplified Goal.forked_proofs: status is determined via group instead of dummy future (see also Pure/PIDE/execution.ML);
 wenzelm parents: 
51661diff
changeset | 206 | val _ = Par_Exn.release_all (join_theory result); | 
| 51331 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 207 | val _ = result_present result (); | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 208 | val _ = result_commit result (); | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 209 | in result_theory result end | 
| 44162 | 210 | | Finished thy => thy)) #> ignore; | 
| 29000 
5e03bc760355
improved future_schedule: more robust handling of failed parents (explicit join), final join of all futures, including Exn.release_all;
 wenzelm parents: 
28980diff
changeset | 211 | |
| 51284 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 212 | val schedule_futures = uninterruptible (fn _ => fn tasks => | 
| 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 213 | let | 
| 51331 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 214 | val futures = tasks | 
| 51284 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 215 | |> String_Graph.schedule (fn deps => fn (name, task) => | 
| 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 216 | (case task of | 
| 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 217 | Task (parents, body) => | 
| 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 218 | (singleton o Future.forks) | 
| 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 219 |               {name = "theory:" ^ name, group = NONE,
 | 
| 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 220 | deps = map (Future.task_of o #2) deps, pri = 0, interrupts = true} | 
| 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 221 | (fn () => | 
| 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 222 | (case filter (not o can Future.join o #2) deps of | 
| 51330 | 223 | [] => body (map (result_theory o Future.join) (task_parents deps parents)) | 
| 51284 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 224 | | bad => | 
| 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 225 |                     error (loader_msg ("failed to load " ^ quote name ^
 | 
| 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 226 | " (unresolved " ^ commas_quote (map #1 bad) ^ ")") []))) | 
| 51331 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 227 | | Finished theory => Future.value (theory_result theory))); | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 228 | |
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 229 | val results1 = futures | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 230 | |> maps (fn future => | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 231 | (case Future.join_result future of | 
| 53190 
5d92649a310e
simplified Goal.forked_proofs: status is determined via group instead of dummy future (see also Pure/PIDE/execution.ML);
 wenzelm parents: 
51661diff
changeset | 232 | Exn.Res result => join_theory result | 
| 51331 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 233 | | Exn.Exn exn => [Exn.Exn exn])); | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 234 | |
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 235 | val results2 = futures | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 236 | |> map_filter (Exn.get_res o Future.join_result) | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 237 | |> sort result_ord | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 238 | |> Par_List.map (fn result => Exn.capture (result_present result) ()); | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 239 | |
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 240 | (* FIXME more precise commit order (!?) *) | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 241 | val results3 = futures | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 242 | |> map (fn future => Exn.capture (fn () => result_commit (Future.join future) ()) ()); | 
| 51284 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 243 | |
| 54370 
39ac1a02c60c
join all theory body forks, notably Toplevel.atom_result (diagnostic commands), before peeking at full status;
 wenzelm parents: 
53375diff
changeset | 244 | (* FIXME avoid global Execution.reset (!??) *) | 
| 53192 | 245 | val results4 = map Exn.Exn (maps Task_Queue.group_status (Execution.reset ())); | 
| 51284 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 246 | |
| 51331 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 247 | val _ = Par_Exn.release_all (results1 @ results2 @ results3 @ results4); | 
| 51284 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 248 | in () end); | 
| 29429 
a6c641f08af7
schedule_futures: tuned final consolidation, explicit after_load phase;
 wenzelm parents: 
29421diff
changeset | 249 | |
| 24209 | 250 | in | 
| 251 | ||
| 32794 | 252 | fun schedule_tasks tasks = | 
| 29118 | 253 | if not (Multithreading.enabled ()) then schedule_seq tasks | 
| 254 | else if Multithreading.self_critical () then | |
| 24209 | 255 | (warning (loader_msg "no multithreading within critical section" []); | 
| 256 | schedule_seq tasks) | |
| 29118 | 257 | else schedule_futures tasks; | 
| 24209 | 258 | |
| 259 | end; | |
| 260 | ||
| 261 | ||
| 23967 | 262 | (* require_thy -- checking database entries wrt. the file-system *) | 
| 15065 | 263 | |
| 7211 | 264 | local | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 265 | |
| 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: 
37955diff
changeset | 266 | fun required_by _ [] = "" | 
| 
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: 
37955diff
changeset | 267 | | required_by s initiators = s ^ "(required by " ^ show_path (rev initiators) ^ ")"; | 
| 
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: 
37955diff
changeset | 268 | |
| 54458 
96ccc8972fc7
prefer explicit "document" flag -- eliminated stateful Present.no_document;
 wenzelm parents: 
54446diff
changeset | 269 | fun load_thy document last_timing initiators update_time deps text (name, pos) keywords parents = | 
| 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: 
37955diff
changeset | 270 | let | 
| 
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: 
37955diff
changeset | 271 | val _ = kill_thy name; | 
| 40132 
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
 wenzelm parents: 
39557diff
changeset | 272 |     val _ = Output.urgent_message ("Loading theory " ^ quote name ^ required_by " " initiators);
 | 
| 51661 | 273 | val _ = Output.try_protocol_message (Markup.loading_theory name) ""; | 
| 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: 
37955diff
changeset | 274 | |
| 41548 
bd0bebf93fa6
Thy_Load.begin_theory: maintain source specification of imports;
 wenzelm parents: 
41536diff
changeset | 275 |     val {master = (thy_path, _), imports} = deps;
 | 
| 38134 
3de75ca6f166
load_thy: refer to physical master directory (not accumulated source import directory) and enable loading files relatively to that;
 wenzelm parents: 
37985diff
changeset | 276 | val dir = Path.dir thy_path; | 
| 51294 
0850d43cb355
discontinued obsolete header "files" -- these are loaded explicitly after exploring dependencies;
 wenzelm parents: 
51293diff
changeset | 277 | val header = Thy_Header.make (name, pos) imports keywords; | 
| 48927 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 wenzelm parents: 
48898diff
changeset | 278 | |
| 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 wenzelm parents: 
48898diff
changeset | 279 | val _ = Position.reports (map #2 imports ~~ map Theory.get_markup parents); | 
| 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: 
37955diff
changeset | 280 | |
| 53375 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 281 | val exec_id = Document_ID.make (); | 
| 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 282 | val _ = | 
| 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 283 | Execution.running Document_ID.none exec_id [] orelse | 
| 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 284 |         raise Fail ("Failed to register execution: " ^ Document_ID.print exec_id);
 | 
| 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 285 | |
| 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 286 | val text_pos = Position.put_id (Document_ID.print exec_id) (Path.position thy_path); | 
| 51331 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 287 | val (theory, present, weight) = | 
| 54458 
96ccc8972fc7
prefer explicit "document" flag -- eliminated stateful Present.no_document;
 wenzelm parents: 
54446diff
changeset | 288 | Thy_Load.load_thy document last_timing update_time dir header text_pos text | 
| 48638 | 289 | (if name = Context.PureN then [ML_Context.the_global_context ()] else parents); | 
| 43640 | 290 | fun commit () = update_thy deps theory; | 
| 53375 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 291 | in | 
| 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 292 |     Result {theory = theory, exec_id = exec_id, present = present, commit = commit, weight = weight}
 | 
| 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 293 | end; | 
| 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: 
37955diff
changeset | 294 | |
| 48898 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 wenzelm parents: 
48888diff
changeset | 295 | fun check_deps dir name = | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 296 | (case lookup_deps name of | 
| 51293 
05b1bbae748d
discontinued obsolete 'uses' within theory header;
 wenzelm parents: 
51284diff
changeset | 297 | SOME NONE => (true, NONE, Position.none, get_imports name, []) | 
| 23893 
8babfcaaf129
deps: maintain source specification of parents (prevents repeated ThyLoad.deps_thy);
 wenzelm parents: 
23886diff
changeset | 298 | | NONE => | 
| 51293 
05b1bbae748d
discontinued obsolete 'uses' within theory header;
 wenzelm parents: 
51284diff
changeset | 299 |       let val {master, text, theory_pos, imports, keywords} = Thy_Load.check_thy dir name
 | 
| 
05b1bbae748d
discontinued obsolete 'uses' within theory header;
 wenzelm parents: 
51284diff
changeset | 300 | in (false, SOME (make_deps master imports, text), theory_pos, imports, keywords) end | 
| 42129 
c17508a61f49
present theory content as future, depending on intermediate proof state futures -- potential to reduce memory requirements and improve parallelization;
 wenzelm parents: 
42003diff
changeset | 301 |   | SOME (SOME {master, ...}) =>
 | 
| 42003 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 wenzelm parents: 
41955diff
changeset | 302 | let | 
| 48928 | 303 |         val {master = master', text = text', theory_pos = theory_pos', imports = imports',
 | 
| 51293 
05b1bbae748d
discontinued obsolete 'uses' within theory header;
 wenzelm parents: 
51284diff
changeset | 304 | keywords = keywords'} = Thy_Load.check_thy dir name; | 
| 42003 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 wenzelm parents: 
41955diff
changeset | 305 | val deps' = SOME (make_deps master' imports', text'); | 
| 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 wenzelm parents: 
41955diff
changeset | 306 | val current = | 
| 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 wenzelm parents: 
41955diff
changeset | 307 | #2 master = #2 master' andalso | 
| 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 wenzelm parents: 
41955diff
changeset | 308 | (case lookup_theory name of | 
| 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 wenzelm parents: 
41955diff
changeset | 309 | NONE => false | 
| 54446 | 310 | | SOME theory => Thy_Load.loaded_files_current theory); | 
| 51293 
05b1bbae748d
discontinued obsolete 'uses' within theory header;
 wenzelm parents: 
51284diff
changeset | 311 | in (current, deps', theory_pos', imports', keywords') end); | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 312 | |
| 23967 | 313 | in | 
| 314 | ||
| 54458 
96ccc8972fc7
prefer explicit "document" flag -- eliminated stateful Present.no_document;
 wenzelm parents: 
54446diff
changeset | 315 | fun require_thys document last_timing initiators dir strs tasks = | 
| 
96ccc8972fc7
prefer explicit "document" flag -- eliminated stateful Present.no_document;
 wenzelm parents: 
54446diff
changeset | 316 | fold_map (require_thy document last_timing initiators dir) strs tasks |>> forall I | 
| 
96ccc8972fc7
prefer explicit "document" flag -- eliminated stateful Present.no_document;
 wenzelm parents: 
54446diff
changeset | 317 | and require_thy document last_timing initiators dir (str, require_pos) tasks = | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 318 | let | 
| 21858 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
20664diff
changeset | 319 | val path = Path.expand (Path.explode str); | 
| 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
20664diff
changeset | 320 | val name = Path.implode (Path.base path); | 
| 7066 | 321 | in | 
| 50862 
5fc8b83322f5
more sensible order of theory nodes (correspondance to Scala version), e.g. relevant to theory progress;
 wenzelm parents: 
50845diff
changeset | 322 | (case try (String_Graph.get_node tasks) name of | 
| 48898 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 wenzelm parents: 
48888diff
changeset | 323 | SOME task => (task_finished task, tasks) | 
| 23967 | 324 | | NONE => | 
| 325 | let | |
| 43651 
511df47bcadc
some support for theory files within Isabelle/Scala session;
 wenzelm parents: 
43641diff
changeset | 326 | val dir' = Path.append dir (Path.dir path); | 
| 
511df47bcadc
some support for theory files within Isabelle/Scala session;
 wenzelm parents: 
43641diff
changeset | 327 | val _ = member (op =) initiators name andalso error (cycle_msg initiators); | 
| 
511df47bcadc
some support for theory files within Isabelle/Scala session;
 wenzelm parents: 
43641diff
changeset | 328 | |
| 51293 
05b1bbae748d
discontinued obsolete 'uses' within theory header;
 wenzelm parents: 
51284diff
changeset | 329 | val (current, deps, theory_pos, imports, keywords) = check_deps dir' name | 
| 23967 | 330 | handle ERROR msg => cat_error msg | 
| 331 | (loader_msg "the error(s) above occurred while examining theory" [name] ^ | |
| 48992 | 332 | Position.here require_pos ^ required_by "\n" initiators); | 
| 37981 
9a15982f41fe
theory loader: removed obsolete touch/outdate operations (require_thy no longer changes the database implicitly);
 wenzelm parents: 
37980diff
changeset | 333 | |
| 48927 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 wenzelm parents: 
48898diff
changeset | 334 | val parents = map (base_name o #1) imports; | 
| 48898 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 wenzelm parents: 
48888diff
changeset | 335 | val (parents_current, tasks') = | 
| 54458 
96ccc8972fc7
prefer explicit "document" flag -- eliminated stateful Present.no_document;
 wenzelm parents: 
54446diff
changeset | 336 | require_thys document last_timing (name :: initiators) | 
| 48898 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 wenzelm parents: 
48888diff
changeset | 337 | (Path.append dir (master_dir (Option.map #1 deps))) imports tasks; | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 338 | |
| 23967 | 339 | val all_current = current andalso parents_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: 
37955diff
changeset | 340 | val task = | 
| 
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: 
37955diff
changeset | 341 | if all_current then Finished (get_theory name) | 
| 
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: 
37955diff
changeset | 342 | else | 
| 37980 
b8c3d4dc1e3e
avoid repeated File.read for theory text (as before);
 wenzelm parents: 
37979diff
changeset | 343 | (case deps of | 
| 
b8c3d4dc1e3e
avoid repeated File.read for theory text (as before);
 wenzelm parents: 
37979diff
changeset | 344 | NONE => raise Fail "Malformed deps" | 
| 42003 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 wenzelm parents: 
41955diff
changeset | 345 | | SOME (dep, text) => | 
| 48876 
157dd47032e0
more standard Thy_Load.check_thy for Pure.thy, relying on its header;
 wenzelm parents: 
48874diff
changeset | 346 | let | 
| 
157dd47032e0
more standard Thy_Load.check_thy for Pure.thy, relying on its header;
 wenzelm parents: 
48874diff
changeset | 347 | val update_time = serial (); | 
| 48928 | 348 | val load = | 
| 54458 
96ccc8972fc7
prefer explicit "document" flag -- eliminated stateful Present.no_document;
 wenzelm parents: 
54446diff
changeset | 349 | load_thy document last_timing initiators update_time dep | 
| 51293 
05b1bbae748d
discontinued obsolete 'uses' within theory header;
 wenzelm parents: 
51284diff
changeset | 350 | text (name, theory_pos) keywords; | 
| 48876 
157dd47032e0
more standard Thy_Load.check_thy for Pure.thy, relying on its header;
 wenzelm parents: 
48874diff
changeset | 351 | in Task (parents, load) end); | 
| 48874 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48710diff
changeset | 352 | |
| 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48710diff
changeset | 353 | val tasks'' = new_entry name parents task tasks'; | 
| 48898 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 wenzelm parents: 
48888diff
changeset | 354 | in (all_current, tasks'') end) | 
| 7066 | 355 | end; | 
| 6484 
3f098b0ec683
use_thy etc.: may specify path prefix, which is temporarily used as load path;
 wenzelm parents: 
6389diff
changeset | 356 | |
| 23967 | 357 | end; | 
| 358 | ||
| 359 | ||
| 37979 | 360 | (* use_thy *) | 
| 23967 | 361 | |
| 54458 
96ccc8972fc7
prefer explicit "document" flag -- eliminated stateful Present.no_document;
 wenzelm parents: 
54446diff
changeset | 362 | fun use_theories {document, last_timing, master_dir} imports =
 | 
| 
96ccc8972fc7
prefer explicit "document" flag -- eliminated stateful Present.no_document;
 wenzelm parents: 
54446diff
changeset | 363 | schedule_tasks (snd (require_thys document last_timing [] master_dir imports String_Graph.empty)); | 
| 23967 | 364 | |
| 54458 
96ccc8972fc7
prefer explicit "document" flag -- eliminated stateful Present.no_document;
 wenzelm parents: 
54446diff
changeset | 365 | val use_thys = use_theories {document = false, last_timing = K NONE, master_dir = Path.current};
 | 
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37942diff
changeset | 366 | val use_thy = use_thys o single; | 
| 7211 | 367 | |
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 368 | |
| 37979 | 369 | (* toplevel begin theory -- without maintaining database *) | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 370 | |
| 48888 | 371 | fun toplevel_begin_theory master_dir (header: Thy_Header.header) = | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 372 | let | 
| 48927 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 wenzelm parents: 
48898diff
changeset | 373 |     val {name = (name, _), imports, ...} = header;
 | 
| 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: 
37955diff
changeset | 374 | val _ = kill_thy name; | 
| 54458 
96ccc8972fc7
prefer explicit "document" flag -- eliminated stateful Present.no_document;
 wenzelm parents: 
54446diff
changeset | 375 |     val _ = use_theories {document = false, last_timing = K NONE, master_dir = master_dir} imports;
 | 
| 46958 
0ec8f04e753a
define keywords early when processing the theory header, before running the body commands;
 wenzelm parents: 
46938diff
changeset | 376 | val _ = Thy_Header.define_keywords header; | 
| 48927 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 wenzelm parents: 
48898diff
changeset | 377 | val parents = map (get_theory o base_name o fst) imports; | 
| 48888 | 378 | in Thy_Load.begin_theory master_dir header parents end; | 
| 7244 | 379 | |
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 380 | |
| 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: 
37955diff
changeset | 381 | (* register theory *) | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 382 | |
| 37954 
a2e73df0b1e0
simplified/clarified register_thy: more precise treatment of new dependencies, remove descendants;
 wenzelm parents: 
37953diff
changeset | 383 | fun register_thy theory = | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 384 | let | 
| 16454 | 385 | val name = Context.theory_name theory; | 
| 48898 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 wenzelm parents: 
48888diff
changeset | 386 |     val {master, ...} = Thy_Load.check_thy (Thy_Load.master_directory theory) name;
 | 
| 41548 
bd0bebf93fa6
Thy_Load.begin_theory: maintain source specification of imports;
 wenzelm parents: 
41536diff
changeset | 387 | val imports = Thy_Load.imports_of theory; | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 388 | in | 
| 38145 
675827e61eb9
more precise CRITICAL sections, using NAMED_CRITICAL uniformly;
 wenzelm parents: 
38143diff
changeset | 389 | NAMED_CRITICAL "Thy_Info" (fn () => | 
| 38142 | 390 | (kill_thy name; | 
| 40132 
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
 wenzelm parents: 
39557diff
changeset | 391 |       Output.urgent_message ("Registering theory " ^ quote name);
 | 
| 43640 | 392 | update_thy (make_deps master imports) theory)) | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 393 | end; | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 394 | |
| 24080 
8c67d869531b
added register_thy (replaces pretend_use_thy_only and really flag);
 wenzelm parents: 
24071diff
changeset | 395 | |
| 
8c67d869531b
added register_thy (replaces pretend_use_thy_only and really flag);
 wenzelm parents: 
24071diff
changeset | 396 | (* finish all theories *) | 
| 
8c67d869531b
added register_thy (replaces pretend_use_thy_only and really flag);
 wenzelm parents: 
24071diff
changeset | 397 | |
| 50862 
5fc8b83322f5
more sensible order of theory nodes (correspondance to Scala version), e.g. relevant to theory progress;
 wenzelm parents: 
50845diff
changeset | 398 | fun finish () = change_thys (String_Graph.map (fn _ => fn (_, entry) => (NONE, entry))); | 
| 24080 
8c67d869531b
added register_thy (replaces pretend_use_thy_only and really flag);
 wenzelm parents: 
24071diff
changeset | 399 | |
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 400 | end; |