| author | wenzelm | 
| Thu, 06 Apr 2017 14:32:56 +0200 | |
| changeset 65404 | 2b819faf45e9 | 
| parent 65058 | 3e9f382fb67e | 
| child 65443 | dccbfc715904 | 
| 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 | 
| 26614 | 10 | val get_names: unit -> string list | 
| 7288 | 11 | val lookup_theory: string -> theory option | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 12 | val get_theory: string -> theory | 
| 60937 | 13 | val pure_theory: unit -> theory | 
| 26983 | 14 | val master_directory: string -> Path.T | 
| 29421 
db532e37cda2
use_thys: perform consolidate_thy on loaded theories, which removes failed nodes in post-hoc fashion;
 wenzelm parents: 
29118diff
changeset | 15 | val remove_thy: string -> unit | 
| 59366 
e94df7f6b608
clarified build_theories: proper protocol handler;
 wenzelm parents: 
59364diff
changeset | 16 | val use_theories: | 
| 61381 | 17 |     {document: bool,
 | 
| 18 | symbols: HTML.symbols, | |
| 65058 
3e9f382fb67e
absent timing information means zero, according to 0070053570c4, f235646b1b73;
 wenzelm parents: 
64574diff
changeset | 19 | last_timing: Toplevel.transition -> Time.time, | 
| 61381 | 20 | master_dir: Path.T} -> (string * Position.T) list -> unit | 
| 48927 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 wenzelm parents: 
48898diff
changeset | 21 | val use_thys: (string * Position.T) list -> unit | 
| 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 wenzelm parents: 
48898diff
changeset | 22 | val use_thy: string * Position.T -> unit | 
| 58856 | 23 | val script_thy: Position.T -> string -> theory -> theory | 
| 37954 
a2e73df0b1e0
simplified/clarified register_thy: more precise treatment of new dependencies, remove descendants;
 wenzelm parents: 
37953diff
changeset | 24 | val register_thy: theory -> unit | 
| 24080 
8c67d869531b
added register_thy (replaces pretend_use_thy_only and really flag);
 wenzelm parents: 
24071diff
changeset | 25 | val finish: unit -> unit | 
| 3604 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 26 | end; | 
| 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 27 | |
| 37216 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 wenzelm parents: 
36953diff
changeset | 28 | structure Thy_Info: THY_INFO = | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 29 | struct | 
| 3604 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 30 | |
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 31 | (** thy database **) | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 32 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 33 | (* messages *) | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 34 | |
| 55797 | 35 | val show_path = space_implode " via " o map quote; | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 36 | |
| 55797 | 37 | fun cycle_msg names = "Cyclic dependency of " ^ show_path names; | 
| 3604 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 38 | |
| 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 39 | |
| 6666 | 40 | (* derived graph operations *) | 
| 3604 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 41 | |
| 50862 
5fc8b83322f5
more sensible order of theory nodes (correspondance to Scala version), e.g. relevant to theory progress;
 wenzelm parents: 
50845diff
changeset | 42 | 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 | 43 | 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 | 44 | |
| 37984 
f26352bd82cf
clarified register_thy: clean slate via kill_thy, more precise CRITICAL section;
 wenzelm parents: 
37981diff
changeset | 45 | 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 | 46 | 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 | 47 | |
| 3604 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 48 | |
| 59178 | 49 | (* global thys *) | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 50 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 51 | type deps = | 
| 41955 
703ea96b13c6
files are identified via SHA1 digests -- discontinued ISABELLE_FILE_IDENT;
 wenzelm parents: 
41678diff
changeset | 52 |  {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 | 53 | imports: (string * Position.T) list}; (*source specification of imports (partially qualified)*) | 
| 23886 
f40fba467384
simplified ThyLoad interfaces: only one additional directory;
 wenzelm parents: 
23871diff
changeset | 54 | |
| 38148 
d2f3c8d4a89f
uniform naming of imports (source specification) vs. parents (thy node names) vs. parent_thys (theory values);
 wenzelm parents: 
38147diff
changeset | 55 | 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 | 56 | |
| 59178 | 57 | fun master_dir (d: deps option) = | 
| 58 | the_default Path.current (Option.map (Path.dir o #1 o #master) d); | |
| 59 | ||
| 44225 
a8f921e6484f
more robust Thy_Header.base_name, with minimal assumptions about path syntax;
 wenzelm parents: 
44222diff
changeset | 60 | fun base_name s = Path.implode (Path.base (Path.explode s)); | 
| 23967 | 61 | |
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 62 | local | 
| 59178 | 63 | val global_thys = | 
| 64 | Synchronized.var "Thy_Info.thys" | |
| 65 | (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 | 66 | in | 
| 59178 | 67 | fun get_thys () = Synchronized.value global_thys; | 
| 68 | fun change_thys f = Synchronized.change global_thys f; | |
| 3604 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 69 | end; | 
| 5209 | 70 | |
| 50862 
5fc8b83322f5
more sensible order of theory nodes (correspondance to Scala version), e.g. relevant to theory progress;
 wenzelm parents: 
50845diff
changeset | 71 | fun get_names () = String_Graph.topological_order (get_thys ()); | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 72 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 73 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 74 | (* access thy *) | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 75 | |
| 59178 | 76 | fun lookup thys name = try (String_Graph.get_node thys) name; | 
| 77 | fun lookup_thy name = lookup (get_thys ()) name; | |
| 7935 | 78 | |
| 59178 | 79 | fun get thys name = | 
| 80 | (case lookup thys name of | |
| 15531 | 81 | SOME thy => thy | 
| 55797 | 82 |   | NONE => error ("Theory loader: nothing known about theory " ^ quote name));
 | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 83 | |
| 59178 | 84 | fun get_thy name = get (get_thys ()) name; | 
| 85 | ||
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 86 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 87 | (* access deps *) | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 88 | |
| 15570 | 89 | val lookup_deps = Option.map #1 o lookup_thy; | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 90 | |
| 59178 | 91 | val master_directory = master_dir o #1 o get_thy; | 
| 7191 | 92 | |
| 6211 
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 theory *) | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 95 | |
| 7687 | 96 | fun lookup_theory name = | 
| 97 | (case lookup_thy name of | |
| 38142 | 98 | SOME (_, SOME theory) => SOME theory | 
| 15531 | 99 | | _ => NONE); | 
| 7288 | 100 | |
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 101 | fun get_theory name = | 
| 7288 | 102 | (case lookup_theory name of | 
| 23871 | 103 | SOME theory => theory | 
| 55797 | 104 |   | _ => error ("Theory loader: undefined entry for theory " ^ quote name));
 | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 105 | |
| 62942 | 106 | fun pure_theory () = get_theory Context.PureN; | 
| 60937 | 107 | |
| 56208 | 108 | val get_imports = Resources.imports_of o get_theory; | 
| 44337 
d453faed4815
clarified get_imports -- should not rely on accidental order within graph;
 wenzelm parents: 
44302diff
changeset | 109 | |
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 110 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 111 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 112 | (** thy operations **) | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 113 | |
| 59178 | 114 | (* remove *) | 
| 29421 
db532e37cda2
use_thys: perform consolidate_thy on loaded theories, which removes failed nodes in post-hoc fashion;
 wenzelm parents: 
29118diff
changeset | 115 | |
| 59178 | 116 | fun remove name thys = | 
| 117 | (case lookup thys name of | |
| 118 | NONE => thys | |
| 119 |   | SOME (NONE, _) => error ("Cannot update finished theory " ^ quote name)
 | |
| 120 | | SOME _ => | |
| 121 | let | |
| 122 | val succs = String_Graph.all_succs thys [name]; | |
| 123 |         val _ = writeln ("Theory loader: removing " ^ commas_quote succs);
 | |
| 124 | in fold String_Graph.del_node succs thys end); | |
| 29421 
db532e37cda2
use_thys: perform consolidate_thy on loaded theories, which removes failed nodes in post-hoc fashion;
 wenzelm parents: 
29118diff
changeset | 125 | |
| 59178 | 126 | val remove_thy = change_thys o remove; | 
| 127 | ||
| 7099 | 128 | |
| 59178 | 129 | (* update *) | 
| 130 | ||
| 131 | fun update deps theory thys = | |
| 43640 | 132 | let | 
| 133 | val name = Context.theory_name theory; | |
| 134 | val parents = map Context.theory_name (Theory.parents_of theory); | |
| 59178 | 135 | |
| 136 | val thys' = remove name thys; | |
| 137 | val _ = map (get thys') parents; | |
| 138 | in new_entry name parents (SOME deps, SOME theory) thys' end; | |
| 139 | ||
| 140 | fun update_thy deps theory = change_thys (update deps theory); | |
| 43640 | 141 | |
| 7099 | 142 | |
| 24209 | 143 | (* scheduling loader tasks *) | 
| 144 | ||
| 51330 | 145 | datatype result = | 
| 53375 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 146 |   Result of {theory: theory, exec_id: Document_ID.exec,
 | 
| 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 147 | 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 | 148 | |
| 53375 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 149 | fun theory_result theory = | 
| 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 150 |   Result {theory = theory, exec_id = Document_ID.none, present = I, commit = I, weight = 0};
 | 
| 51330 | 151 | |
| 152 | fun result_theory (Result {theory, ...}) = theory;
 | |
| 51331 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 153 | fun result_present (Result {present, ...}) = present;
 | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 154 | fun result_commit (Result {commit, ...}) = commit;
 | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 155 | 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 | 156 | |
| 53375 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 157 | 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 | 158 | let | 
| 
39ac1a02c60c
join all theory body forks, notably Toplevel.atom_result (diagnostic commands), before peeking at full status;
 wenzelm parents: 
53375diff
changeset | 159 | (*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 | 160 | 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 | 161 | (*fully nested proofs*) | 
| 64574 
1134e4d5e5b7
consolidate nested thms with persistent result, for improved performance;
 wenzelm parents: 
62942diff
changeset | 162 | val res = Exn.capture Thm.consolidate_theory theory; | 
| 54370 
39ac1a02c60c
join all theory body forks, notably Toplevel.atom_result (diagnostic commands), before peeking at full status;
 wenzelm parents: 
53375diff
changeset | 163 | 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 | 164 | |
| 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 | 165 | datatype task = | 
| 55461 
ce676a750575
more integrity checks of theory names vs. full node names -- at least for the scope of a single use_thys (or "theories" section in ROOT);
 wenzelm parents: 
54722diff
changeset | 166 | Task of Path.T * 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 | 167 | 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 | 168 | |
| 
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 | 169 | 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 | 170 | | task_finished (Finished _) = true; | 
| 24209 | 171 | |
| 44162 | 172 | fun task_parents deps (parents: string list) = map (the o AList.lookup (op =) deps) parents; | 
| 173 | ||
| 29118 | 174 | local | 
| 175 | ||
| 44162 | 176 | 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 | 177 | String_Graph.schedule (fn deps => fn (_, task) => | 
| 44162 | 178 | (case task of | 
| 55461 
ce676a750575
more integrity checks of theory names vs. full node names -- at least for the scope of a single use_thys (or "theories" section in ROOT);
 wenzelm parents: 
54722diff
changeset | 179 | Task (_, parents, body) => | 
| 51331 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 180 | let | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 181 | 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 | 182 | 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 | 183 | val _ = result_present result (); | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 184 | val _ = result_commit result (); | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 185 | in result_theory result end | 
| 44162 | 186 | | 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 | 187 | |
| 62923 | 188 | val schedule_futures = Thread_Attributes.uninterruptible (fn _ => fn tasks => | 
| 51284 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 189 | let | 
| 51331 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 190 | val futures = tasks | 
| 51284 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 191 | |> String_Graph.schedule (fn deps => fn (name, task) => | 
| 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 192 | (case task of | 
| 55461 
ce676a750575
more integrity checks of theory names vs. full node names -- at least for the scope of a single use_thys (or "theories" section in ROOT);
 wenzelm parents: 
54722diff
changeset | 193 | Task (_, parents, body) => | 
| 51284 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 194 | (singleton o Future.forks) | 
| 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 195 |               {name = "theory:" ^ name, group = NONE,
 | 
| 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 196 | 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 | 197 | (fn () => | 
| 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 198 | (case filter (not o can Future.join o #2) deps of | 
| 51330 | 199 | [] => 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 | 200 | | bad => | 
| 55797 | 201 | error | 
| 202 |                       ("Failed to load theory " ^ quote name ^
 | |
| 203 | " (unresolved " ^ commas_quote (map #1 bad) ^ ")"))) | |
| 51331 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 204 | | Finished theory => Future.value (theory_result theory))); | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 205 | |
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 206 | val results1 = futures | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 207 | |> maps (fn future => | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 208 | (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 | 209 | Exn.Res result => join_theory result | 
| 51331 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 210 | | Exn.Exn exn => [Exn.Exn exn])); | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 211 | |
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 212 | val results2 = futures | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 213 | |> 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 | 214 | |> sort result_ord | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 215 | |> 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 | 216 | |
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 217 | (* FIXME more precise commit order (!?) *) | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 218 | val results3 = futures | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 219 | |> 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 | 220 | |
| 54370 
39ac1a02c60c
join all theory body forks, notably Toplevel.atom_result (diagnostic commands), before peeking at full status;
 wenzelm parents: 
53375diff
changeset | 221 | (* FIXME avoid global Execution.reset (!??) *) | 
| 53192 | 222 | 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 | 223 | |
| 51331 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 224 | val _ = Par_Exn.release_all (results1 @ results2 @ results3 @ results4); | 
| 51284 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 225 | in () end); | 
| 29429 
a6c641f08af7
schedule_futures: tuned final consolidation, explicit after_load phase;
 wenzelm parents: 
29421diff
changeset | 226 | |
| 24209 | 227 | in | 
| 228 | ||
| 32794 | 229 | fun schedule_tasks tasks = | 
| 59180 
c0fa3b3bdabd
discontinued central critical sections: NAMED_CRITICAL / CRITICAL;
 wenzelm parents: 
59178diff
changeset | 230 | if Multithreading.enabled () then schedule_futures tasks else schedule_seq tasks; | 
| 24209 | 231 | |
| 232 | end; | |
| 233 | ||
| 234 | ||
| 23967 | 235 | (* require_thy -- checking database entries wrt. the file-system *) | 
| 15065 | 236 | |
| 7211 | 237 | local | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 238 | |
| 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 | 239 | 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 | 240 | | 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 | 241 | |
| 61381 | 242 | fun load_thy document symbols last_timing | 
| 243 | 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 | 244 | let | 
| 59178 | 245 | val _ = remove_thy name; | 
| 58843 | 246 |     val _ = writeln ("Loading theory " ^ quote name ^ required_by " " initiators);
 | 
| 56333 
38f1422ef473
support bulk messages consisting of small string segments, which are more healthy to the Poly/ML RTS and might prevent spurious GC crashes such as MTGCProcessMarkPointers::ScanAddressesInObject;
 wenzelm parents: 
56208diff
changeset | 247 | 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 | 248 | |
| 41548 
bd0bebf93fa6
Thy_Load.begin_theory: maintain source specification of imports;
 wenzelm parents: 
41536diff
changeset | 249 |     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 | 250 | val dir = Path.dir thy_path; | 
| 51294 
0850d43cb355
discontinued obsolete header "files" -- these are loaded explicitly after exploring dependencies;
 wenzelm parents: 
51293diff
changeset | 251 | val header = Thy_Header.make (name, pos) imports keywords; | 
| 48927 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 wenzelm parents: 
48898diff
changeset | 252 | |
| 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 wenzelm parents: 
48898diff
changeset | 253 | 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 | 254 | |
| 53375 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 255 | val exec_id = Document_ID.make (); | 
| 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 256 | val _ = | 
| 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 257 | Execution.running Document_ID.none exec_id [] orelse | 
| 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 258 |         raise Fail ("Failed to register execution: " ^ Document_ID.print exec_id);
 | 
| 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 259 | |
| 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 260 | 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 | 261 | val (theory, present, weight) = | 
| 61381 | 262 | Resources.load_thy document symbols last_timing update_time dir header text_pos text | 
| 62876 | 263 | (if name = Context.PureN then [Context.the_global_context ()] else parents); | 
| 43640 | 264 | fun commit () = update_thy deps theory; | 
| 53375 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 265 | in | 
| 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 266 |     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 | 267 | 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 | 268 | |
| 48898 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 wenzelm parents: 
48888diff
changeset | 269 | fun check_deps dir name = | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 270 | (case lookup_deps name of | 
| 51293 
05b1bbae748d
discontinued obsolete 'uses' within theory header;
 wenzelm parents: 
51284diff
changeset | 271 | 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 | 272 | | NONE => | 
| 56208 | 273 |       let val {master, text, theory_pos, imports, keywords} = Resources.check_thy dir name
 | 
| 51293 
05b1bbae748d
discontinued obsolete 'uses' within theory header;
 wenzelm parents: 
51284diff
changeset | 274 | 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 | 275 |   | SOME (SOME {master, ...}) =>
 | 
| 42003 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 wenzelm parents: 
41955diff
changeset | 276 | let | 
| 48928 | 277 |         val {master = master', text = text', theory_pos = theory_pos', imports = imports',
 | 
| 56208 | 278 | keywords = keywords'} = Resources.check_thy dir name; | 
| 42003 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 wenzelm parents: 
41955diff
changeset | 279 | val deps' = SOME (make_deps master' imports', text'); | 
| 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 wenzelm parents: 
41955diff
changeset | 280 | val current = | 
| 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 wenzelm parents: 
41955diff
changeset | 281 | #2 master = #2 master' andalso | 
| 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 wenzelm parents: 
41955diff
changeset | 282 | (case lookup_theory name of | 
| 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 wenzelm parents: 
41955diff
changeset | 283 | NONE => false | 
| 56208 | 284 | | SOME theory => Resources.loaded_files_current theory); | 
| 51293 
05b1bbae748d
discontinued obsolete 'uses' within theory header;
 wenzelm parents: 
51284diff
changeset | 285 | in (current, deps', theory_pos', imports', keywords') end); | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 286 | |
| 23967 | 287 | in | 
| 288 | ||
| 61381 | 289 | fun require_thys document symbols last_timing initiators dir strs tasks = | 
| 290 | fold_map (require_thy document symbols last_timing initiators dir) strs tasks |>> forall I | |
| 291 | and require_thy document symbols last_timing initiators dir (str, require_pos) tasks = | |
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 292 | let | 
| 21858 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
20664diff
changeset | 293 | 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 | 294 | val name = Path.implode (Path.base path); | 
| 56208 | 295 | val node_name = File.full_path dir (Resources.thy_path path); | 
| 55461 
ce676a750575
more integrity checks of theory names vs. full node names -- at least for the scope of a single use_thys (or "theories" section in ROOT);
 wenzelm parents: 
54722diff
changeset | 296 | fun check_entry (Task (node_name', _, _)) = | 
| 60975 
5f3d6e16ea78
avoid ambiguities on native Windows, such as / vs. /cygdrive/c/cygwin;
 wenzelm parents: 
60937diff
changeset | 297 | if op = (apply2 File.platform_path (node_name, node_name')) | 
| 
5f3d6e16ea78
avoid ambiguities on native Windows, such as / vs. /cygdrive/c/cygwin;
 wenzelm parents: 
60937diff
changeset | 298 | then () | 
| 55461 
ce676a750575
more integrity checks of theory names vs. full node names -- at least for the scope of a single use_thys (or "theories" section in ROOT);
 wenzelm parents: 
54722diff
changeset | 299 | else | 
| 55797 | 300 |             error ("Incoherent imports for theory " ^ quote name ^
 | 
| 55461 
ce676a750575
more integrity checks of theory names vs. full node names -- at least for the scope of a single use_thys (or "theories" section in ROOT);
 wenzelm parents: 
54722diff
changeset | 301 | Position.here require_pos ^ ":\n" ^ | 
| 55488 
60c159d490a2
more integrity checks of theory names vs. full node names;
 wenzelm parents: 
55461diff
changeset | 302 | " " ^ Path.print node_name ^ "\n" ^ | 
| 
60c159d490a2
more integrity checks of theory names vs. full node names;
 wenzelm parents: 
55461diff
changeset | 303 | " " ^ Path.print node_name') | 
| 55461 
ce676a750575
more integrity checks of theory names vs. full node names -- at least for the scope of a single use_thys (or "theories" section in ROOT);
 wenzelm parents: 
54722diff
changeset | 304 | | check_entry _ = (); | 
| 7066 | 305 | in | 
| 50862 
5fc8b83322f5
more sensible order of theory nodes (correspondance to Scala version), e.g. relevant to theory progress;
 wenzelm parents: 
50845diff
changeset | 306 | (case try (String_Graph.get_node tasks) name of | 
| 55461 
ce676a750575
more integrity checks of theory names vs. full node names -- at least for the scope of a single use_thys (or "theories" section in ROOT);
 wenzelm parents: 
54722diff
changeset | 307 | SOME task => (check_entry task; (task_finished task, tasks)) | 
| 23967 | 308 | | NONE => | 
| 309 | let | |
| 43651 
511df47bcadc
some support for theory files within Isabelle/Scala session;
 wenzelm parents: 
43641diff
changeset | 310 | val dir' = Path.append dir (Path.dir path); | 
| 
511df47bcadc
some support for theory files within Isabelle/Scala session;
 wenzelm parents: 
43641diff
changeset | 311 | val _ = member (op =) initiators name andalso error (cycle_msg initiators); | 
| 
511df47bcadc
some support for theory files within Isabelle/Scala session;
 wenzelm parents: 
43641diff
changeset | 312 | |
| 51293 
05b1bbae748d
discontinued obsolete 'uses' within theory header;
 wenzelm parents: 
51284diff
changeset | 313 | val (current, deps, theory_pos, imports, keywords) = check_deps dir' name | 
| 55797 | 314 | handle ERROR msg => | 
| 315 | cat_error msg | |
| 316 |                 ("The error(s) above occurred for theory " ^ quote name ^
 | |
| 317 | 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 | 318 | |
| 48927 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 wenzelm parents: 
48898diff
changeset | 319 | 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 | 320 | val (parents_current, tasks') = | 
| 61381 | 321 | require_thys document symbols 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 | 322 | (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 | 323 | |
| 23967 | 324 | 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 | 325 | 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 | 326 | 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 | 327 | else | 
| 37980 
b8c3d4dc1e3e
avoid repeated File.read for theory text (as before);
 wenzelm parents: 
37979diff
changeset | 328 | (case deps of | 
| 
b8c3d4dc1e3e
avoid repeated File.read for theory text (as before);
 wenzelm parents: 
37979diff
changeset | 329 | NONE => raise Fail "Malformed deps" | 
| 42003 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 wenzelm parents: 
41955diff
changeset | 330 | | SOME (dep, text) => | 
| 48876 
157dd47032e0
more standard Thy_Load.check_thy for Pure.thy, relying on its header;
 wenzelm parents: 
48874diff
changeset | 331 | let | 
| 
157dd47032e0
more standard Thy_Load.check_thy for Pure.thy, relying on its header;
 wenzelm parents: 
48874diff
changeset | 332 | val update_time = serial (); | 
| 48928 | 333 | val load = | 
| 61381 | 334 | load_thy document symbols last_timing initiators update_time dep | 
| 51293 
05b1bbae748d
discontinued obsolete 'uses' within theory header;
 wenzelm parents: 
51284diff
changeset | 335 | text (name, theory_pos) keywords; | 
| 55461 
ce676a750575
more integrity checks of theory names vs. full node names -- at least for the scope of a single use_thys (or "theories" section in ROOT);
 wenzelm parents: 
54722diff
changeset | 336 | in Task (node_name, parents, load) end); | 
| 48874 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48710diff
changeset | 337 | |
| 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48710diff
changeset | 338 | 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 | 339 | in (all_current, tasks'') end) | 
| 7066 | 340 | end; | 
| 6484 
3f098b0ec683
use_thy etc.: may specify path prefix, which is temporarily used as load path;
 wenzelm parents: 
6389diff
changeset | 341 | |
| 23967 | 342 | end; | 
| 343 | ||
| 344 | ||
| 37979 | 345 | (* use_thy *) | 
| 23967 | 346 | |
| 61381 | 347 | fun use_theories {document, symbols, last_timing, master_dir} imports =
 | 
| 348 | schedule_tasks | |
| 349 | (snd (require_thys document symbols last_timing [] master_dir imports String_Graph.empty)); | |
| 23967 | 350 | |
| 61381 | 351 | val use_thys = | 
| 352 | use_theories | |
| 65058 
3e9f382fb67e
absent timing information means zero, according to 0070053570c4, f235646b1b73;
 wenzelm parents: 
64574diff
changeset | 353 |     {document = false, symbols = HTML.no_symbols, last_timing = K Time.zeroTime,
 | 
| 
3e9f382fb67e
absent timing information means zero, according to 0070053570c4, f235646b1b73;
 wenzelm parents: 
64574diff
changeset | 354 | master_dir = Path.current}; | 
| 61381 | 355 | |
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37942diff
changeset | 356 | val use_thy = use_thys o single; | 
| 7211 | 357 | |
| 59364 | 358 | |
| 57626 | 359 | (* toplevel scripting -- without maintaining database *) | 
| 360 | ||
| 58856 | 361 | fun script_thy pos txt thy = | 
| 57626 | 362 | let | 
| 58856 | 363 | val trs = | 
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58856diff
changeset | 364 | Outer_Syntax.parse thy pos txt | 
| 58856 | 365 | |> map (Toplevel.modify_init (K thy)); | 
| 57626 | 366 | val end_pos = if null trs then pos else Toplevel.pos_of (List.last trs); | 
| 367 | val end_state = fold (Toplevel.command_exception true) trs Toplevel.toplevel; | |
| 368 | in Toplevel.end_theory end_pos end_state end; | |
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 369 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 370 | |
| 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 | 371 | (* register theory *) | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 372 | |
| 37954 
a2e73df0b1e0
simplified/clarified register_thy: more precise treatment of new dependencies, remove descendants;
 wenzelm parents: 
37953diff
changeset | 373 | fun register_thy theory = | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 374 | let | 
| 16454 | 375 | val name = Context.theory_name theory; | 
| 56208 | 376 |     val {master, ...} = Resources.check_thy (Resources.master_directory theory) name;
 | 
| 377 | val imports = Resources.imports_of theory; | |
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 378 | in | 
| 59178 | 379 | change_thys (fn thys => | 
| 380 | let | |
| 381 | val thys' = remove name thys; | |
| 382 |         val _ = writeln ("Registering theory " ^ quote name);
 | |
| 383 | in update (make_deps master imports) theory thys' end) | |
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 384 | end; | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 385 | |
| 24080 
8c67d869531b
added register_thy (replaces pretend_use_thy_only and really flag);
 wenzelm parents: 
24071diff
changeset | 386 | |
| 
8c67d869531b
added register_thy (replaces pretend_use_thy_only and really flag);
 wenzelm parents: 
24071diff
changeset | 387 | (* finish all theories *) | 
| 
8c67d869531b
added register_thy (replaces pretend_use_thy_only and really flag);
 wenzelm parents: 
24071diff
changeset | 388 | |
| 50862 
5fc8b83322f5
more sensible order of theory nodes (correspondance to Scala version), e.g. relevant to theory progress;
 wenzelm parents: 
50845diff
changeset | 389 | 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 | 390 | |
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 391 | end; | 
| 62847 | 392 | |
| 393 | fun use_thy name = Runtime.toplevel_program (fn () => Thy_Info.use_thy (name, Position.none)); |