| author | wenzelm | 
| Sat, 22 Apr 2023 21:00:24 +0200 | |
| changeset 77908 | a6bd716a6124 | 
| parent 77889 | 5db014c36f42 | 
| child 78021 | ce6e3bc34343 | 
| permissions | -rw-r--r-- | 
| 3604 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 1 | (* Title: Pure/Thy/thy_info.ML | 
| 72651 | 2 | Author: Makarius | 
| 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 | 
| 72651 | 5 | file dependencies, and presentation of theory documents. | 
| 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 | 
| 68180 
112d9624c020
more general presentation hook, with document preparation as application;
 wenzelm parents: 
68179diff
changeset | 10 | type presentation_context = | 
| 68184 
6c693b2700b3
support for dynamic document output while editing;
 wenzelm parents: 
68182diff
changeset | 11 |     {options: Options.T, file_pos: Position.T, adjust_pos: Position.T -> Position.T,
 | 
| 73761 | 12 | segments: Document_Output.segment list} | 
| 71249 | 13 | val adjust_pos_properties: presentation_context -> Position.T -> Properties.T | 
| 68180 
112d9624c020
more general presentation hook, with document preparation as application;
 wenzelm parents: 
68179diff
changeset | 14 | val apply_presentation: presentation_context -> theory -> unit | 
| 
112d9624c020
more general presentation hook, with document preparation as application;
 wenzelm parents: 
68179diff
changeset | 15 | val add_presentation: (presentation_context -> theory -> unit) -> theory -> theory | 
| 26614 | 16 | val get_names: unit -> string list | 
| 7288 | 17 | val lookup_theory: string -> theory option | 
| 76405 | 18 | val defined_theory: string -> bool | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 19 | val get_theory: string -> theory | 
| 26983 | 20 | 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 | 21 | val remove_thy: string -> unit | 
| 73821 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 22 | val use_theories: Options.T -> string -> (string * Position.T) list -> | 
| 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 23 | (theory * Document_Output.segment list) list | 
| 65444 | 24 | val use_thy: string -> unit | 
| 58856 | 25 | 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 | 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 | |
| 72651 | 33 | (** theory presentation **) | 
| 34 | ||
| 35 | (* hook for consolidated theory *) | |
| 68153 | 36 | |
| 68180 
112d9624c020
more general presentation hook, with document preparation as application;
 wenzelm parents: 
68179diff
changeset | 37 | type presentation_context = | 
| 68184 
6c693b2700b3
support for dynamic document output while editing;
 wenzelm parents: 
68182diff
changeset | 38 |   {options: Options.T, file_pos: Position.T, adjust_pos: Position.T -> Position.T,
 | 
| 73761 | 39 | segments: Document_Output.segment list}; | 
| 68180 
112d9624c020
more general presentation hook, with document preparation as application;
 wenzelm parents: 
68179diff
changeset | 40 | |
| 71249 | 41 | fun adjust_pos_properties (context: presentation_context) pos = | 
| 77770 | 42 | let | 
| 43 | val props = Position.properties_of pos; | |
| 44 | val props' = Position.properties_of (#adjust_pos context pos); | |
| 45 | in | |
| 46 | filter (fn (a, _) => a = Markup.offsetN orelse a = Markup.end_offsetN) props' @ | |
| 47 | filter (fn (a, _) => a = Markup.idN orelse a = Markup.fileN) props | |
| 48 | end; | |
| 71249 | 49 | |
| 68153 | 50 | structure Presentation = Theory_Data | 
| 51 | ( | |
| 68180 
112d9624c020
more general presentation hook, with document preparation as application;
 wenzelm parents: 
68179diff
changeset | 52 | type T = ((presentation_context -> theory -> unit) * stamp) list; | 
| 68153 | 53 | val empty = []; | 
| 54 | fun merge data : T = Library.merge (eq_snd op =) data; | |
| 55 | ); | |
| 56 | ||
| 68180 
112d9624c020
more general presentation hook, with document preparation as application;
 wenzelm parents: 
68179diff
changeset | 57 | fun apply_presentation (context: presentation_context) thy = | 
| 73821 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 58 | ignore (Par_List.map (fn (f, _) => f context thy) (Presentation.get thy)); | 
| 68180 
112d9624c020
more general presentation hook, with document preparation as application;
 wenzelm parents: 
68179diff
changeset | 59 | |
| 68153 | 60 | fun add_presentation f = Presentation.map (cons (f, stamp ())); | 
| 61 | ||
| 68180 
112d9624c020
more general presentation hook, with document preparation as application;
 wenzelm parents: 
68179diff
changeset | 62 | val _ = | 
| 73819 | 63 |   Theory.setup (add_presentation (fn {options, segments, ...} => fn thy =>
 | 
| 68182 | 64 | if exists (Toplevel.is_skipped_proof o #state) segments then () | 
| 68180 
112d9624c020
more general presentation hook, with document preparation as application;
 wenzelm parents: 
68179diff
changeset | 65 | else | 
| 68181 
34592bf86424
export generated document.tex, unless explicit document=false;
 wenzelm parents: 
68180diff
changeset | 66 | let | 
| 76432 | 67 | val keywords = Thy_Header.get_keywords thy; | 
| 77889 
5db014c36f42
clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
 wenzelm parents: 
77770diff
changeset | 68 | val thy_name = Context.theory_base_name thy; | 
| 76432 | 69 | val latex = Document_Output.present_thy options keywords thy_name segments; | 
| 68181 
34592bf86424
export generated document.tex, unless explicit document=false;
 wenzelm parents: 
68180diff
changeset | 70 | in | 
| 72574 
d892f6d66402
build documents in Isabelle/Scala, based on generated tex files as session exports;
 wenzelm parents: 
72511diff
changeset | 71 | if Options.string options "document" = "false" then () | 
| 76432 | 72 | else Export.export thy \<^path_binding>\<open>document/latex\<close> latex | 
| 72710 
6bc199a70bf9
more robust: include reports from Thy_Output.present_thy/output_document;
 wenzelm parents: 
72705diff
changeset | 73 | end)); | 
| 68153 | 74 | |
| 75 | ||
| 76 | ||
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 77 | (** thy database **) | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 78 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 79 | (* messages *) | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 80 | |
| 55797 | 81 | val show_path = space_implode " via " o map quote; | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 82 | |
| 55797 | 83 | fun cycle_msg names = "Cyclic dependency of " ^ show_path names; | 
| 3604 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 84 | |
| 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 85 | |
| 6666 | 86 | (* derived graph operations *) | 
| 3604 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 87 | |
| 50862 
5fc8b83322f5
more sensible order of theory nodes (correspondance to Scala version), e.g. relevant to theory progress;
 wenzelm parents: 
50845diff
changeset | 88 | 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 | 89 | 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 | 90 | |
| 37984 
f26352bd82cf
clarified register_thy: clean slate via kill_thy, more precise CRITICAL section;
 wenzelm parents: 
37981diff
changeset | 91 | 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 | 92 | 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 | 93 | |
| 3604 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 94 | |
| 59178 | 95 | (* global thys *) | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 96 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 97 | type deps = | 
| 41955 
703ea96b13c6
files are identified via SHA1 digests -- discontinued ISABELLE_FILE_IDENT;
 wenzelm parents: 
41678diff
changeset | 98 |  {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 | 99 | imports: (string * Position.T) list}; (*source specification of imports (partially qualified)*) | 
| 23886 
f40fba467384
simplified ThyLoad interfaces: only one additional directory;
 wenzelm parents: 
23871diff
changeset | 100 | |
| 38148 
d2f3c8d4a89f
uniform naming of imports (source specification) vs. parents (thy node names) vs. parent_thys (theory values);
 wenzelm parents: 
38147diff
changeset | 101 | 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 | 102 | |
| 65454 | 103 | fun master_dir_deps (d: deps option) = | 
| 59178 | 104 | the_default Path.current (Option.map (Path.dir o #1 o #master) d); | 
| 105 | ||
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 106 | local | 
| 59178 | 107 | val global_thys = | 
| 108 | Synchronized.var "Thy_Info.thys" | |
| 109 | (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 | 110 | in | 
| 59178 | 111 | fun get_thys () = Synchronized.value global_thys; | 
| 112 | fun change_thys f = Synchronized.change global_thys f; | |
| 3604 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 berghofe parents: diff
changeset | 113 | end; | 
| 5209 | 114 | |
| 50862 
5fc8b83322f5
more sensible order of theory nodes (correspondance to Scala version), e.g. relevant to theory progress;
 wenzelm parents: 
50845diff
changeset | 115 | fun get_names () = String_Graph.topological_order (get_thys ()); | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 116 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 117 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 118 | (* access thy *) | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 119 | |
| 59178 | 120 | fun lookup thys name = try (String_Graph.get_node thys) name; | 
| 121 | fun lookup_thy name = lookup (get_thys ()) name; | |
| 7935 | 122 | |
| 59178 | 123 | fun get thys name = | 
| 124 | (case lookup thys name of | |
| 15531 | 125 | SOME thy => thy | 
| 55797 | 126 |   | 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 | 127 | |
| 59178 | 128 | fun get_thy name = get (get_thys ()) name; | 
| 129 | ||
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 130 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 131 | (* access deps *) | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 132 | |
| 15570 | 133 | val lookup_deps = Option.map #1 o lookup_thy; | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 134 | |
| 65454 | 135 | val master_directory = master_dir_deps o #1 o get_thy; | 
| 7191 | 136 | |
| 6211 
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 | (* access theory *) | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 139 | |
| 7687 | 140 | fun lookup_theory name = | 
| 141 | (case lookup_thy name of | |
| 38142 | 142 | SOME (_, SOME theory) => SOME theory | 
| 15531 | 143 | | _ => NONE); | 
| 7288 | 144 | |
| 76405 | 145 | val defined_theory = is_some o lookup_theory; | 
| 146 | ||
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 147 | fun get_theory name = | 
| 7288 | 148 | (case lookup_theory name of | 
| 23871 | 149 | SOME theory => theory | 
| 55797 | 150 |   | _ => error ("Theory loader: undefined entry for theory " ^ quote name));
 | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 151 | |
| 56208 | 152 | 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 | 153 | |
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 154 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 155 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 156 | (** thy operations **) | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 157 | |
| 59178 | 158 | (* remove *) | 
| 29421 
db532e37cda2
use_thys: perform consolidate_thy on loaded theories, which removes failed nodes in post-hoc fashion;
 wenzelm parents: 
29118diff
changeset | 159 | |
| 59178 | 160 | fun remove name thys = | 
| 161 | (case lookup thys name of | |
| 162 | NONE => thys | |
| 163 |   | SOME (NONE, _) => error ("Cannot update finished theory " ^ quote name)
 | |
| 164 | | SOME _ => | |
| 165 | let | |
| 166 | val succs = String_Graph.all_succs thys [name]; | |
| 167 |         val _ = writeln ("Theory loader: removing " ^ commas_quote succs);
 | |
| 168 | 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 | 169 | |
| 59178 | 170 | val remove_thy = change_thys o remove; | 
| 171 | ||
| 7099 | 172 | |
| 59178 | 173 | (* update *) | 
| 174 | ||
| 175 | fun update deps theory thys = | |
| 43640 | 176 | let | 
| 65459 
da59e8e0a663
clarified theory_long_name (for qualified access to Thy_Info) vs. short theory_name (which is unique within any given theory context);
 wenzelm parents: 
65455diff
changeset | 177 | val name = Context.theory_long_name theory; | 
| 
da59e8e0a663
clarified theory_long_name (for qualified access to Thy_Info) vs. short theory_name (which is unique within any given theory context);
 wenzelm parents: 
65455diff
changeset | 178 | val parents = map Context.theory_long_name (Theory.parents_of theory); | 
| 59178 | 179 | |
| 180 | val thys' = remove name thys; | |
| 181 | val _ = map (get thys') parents; | |
| 182 | in new_entry name parents (SOME deps, SOME theory) thys' end; | |
| 183 | ||
| 184 | fun update_thy deps theory = change_thys (update deps theory); | |
| 43640 | 185 | |
| 7099 | 186 | |
| 24209 | 187 | (* scheduling loader tasks *) | 
| 188 | ||
| 51330 | 189 | datatype result = | 
| 73821 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 190 | Result of | 
| 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 191 |    {theory: theory,
 | 
| 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 192 | exec_id: Document_ID.exec, | 
| 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 193 | present: unit -> presentation_context option, | 
| 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 194 | commit: unit -> unit}; | 
| 51331 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 195 | |
| 53375 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 196 | fun theory_result theory = | 
| 73821 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 197 | Result | 
| 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 198 |    {theory = theory,
 | 
| 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 199 | exec_id = Document_ID.none, | 
| 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 200 | present = K NONE, | 
| 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 201 | commit = I}; | 
| 51330 | 202 | |
| 203 | fun result_theory (Result {theory, ...}) = theory;
 | |
| 51331 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 204 | fun result_commit (Result {commit, ...}) = commit;
 | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 205 | |
| 73825 | 206 | datatype task = | 
| 207 | Task of string list * (theory list -> result) | | |
| 208 | Finished of theory; | |
| 209 | ||
| 210 | local | |
| 211 | ||
| 73817 | 212 | fun consolidate_theory (Exn.Exn exn) = [Exn.Exn exn] | 
| 213 |   | consolidate_theory (Exn.Res (Result {theory, exec_id, ...})) =
 | |
| 214 | let | |
| 215 | val _ = Execution.join [exec_id]; | |
| 216 | val res = Exn.capture Thm.consolidate_theory theory; | |
| 217 | val exns = maps Task_Queue.group_status (Execution.peek exec_id); | |
| 218 | in res :: map Exn.Exn exns end; | |
| 43641 
081e009549dc
uniform finish_thy -- always Global_Theory.join_proofs, even with sequential scheduling;
 wenzelm parents: 
43640diff
changeset | 219 | |
| 73821 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 220 | fun present_theory (Exn.Exn exn) = [Exn.Exn exn] | 
| 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 221 |   | present_theory (Exn.Res (Result {theory, present, ...})) =
 | 
| 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 222 | (case present () of | 
| 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 223 | NONE => [] | 
| 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 224 |       | SOME (context as {segments, ...}) =>
 | 
| 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 225 | [Exn.capture (fn () => (apply_presentation context theory; (theory, segments))) ()]); | 
| 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 226 | |
| 73825 | 227 | in | 
| 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 | 228 | |
| 73818 
d67688992bde
more uniform schedule_theories, notably for "present" and "commit" phase after loading;
 wenzelm parents: 
73817diff
changeset | 229 | val schedule_theories = Thread_Attributes.uninterruptible (fn _ => fn tasks => | 
| 
d67688992bde
more uniform schedule_theories, notably for "present" and "commit" phase after loading;
 wenzelm parents: 
73817diff
changeset | 230 | let | 
| 
d67688992bde
more uniform schedule_theories, notably for "present" and "commit" phase after loading;
 wenzelm parents: 
73817diff
changeset | 231 | fun join_parents deps name parents = | 
| 
d67688992bde
more uniform schedule_theories, notably for "present" and "commit" phase after loading;
 wenzelm parents: 
73817diff
changeset | 232 | (case map #1 (filter (not o can Future.join o #2) deps) of | 
| 
d67688992bde
more uniform schedule_theories, notably for "present" and "commit" phase after loading;
 wenzelm parents: 
73817diff
changeset | 233 | [] => map (result_theory o Future.join o the o AList.lookup (op =) deps) parents | 
| 
d67688992bde
more uniform schedule_theories, notably for "present" and "commit" phase after loading;
 wenzelm parents: 
73817diff
changeset | 234 | | bad => | 
| 
d67688992bde
more uniform schedule_theories, notably for "present" and "commit" phase after loading;
 wenzelm parents: 
73817diff
changeset | 235 |           error ("Failed to load theory " ^ quote name ^ " (unresolved " ^ commas_quote bad ^ ")"));
 | 
| 44162 | 236 | |
| 51331 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 237 | val futures = tasks | 
| 51284 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 238 | |> String_Graph.schedule (fn deps => fn (name, task) => | 
| 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 wenzelm parents: 
51228diff
changeset | 239 | (case task of | 
| 66711 
80fa1401cf76
discontinued extra checks (see ce676a750575 and 60c159d490a2) -- qualified theory names are meant to cover this;
 wenzelm parents: 
66377diff
changeset | 240 | Task (parents, body) => | 
| 73818 
d67688992bde
more uniform schedule_theories, notably for "present" and "commit" phase after loading;
 wenzelm parents: 
73817diff
changeset | 241 | if Multithreading.max_threads () > 1 then | 
| 
d67688992bde
more uniform schedule_theories, notably for "present" and "commit" phase after loading;
 wenzelm parents: 
73817diff
changeset | 242 | (singleton o Future.forks) | 
| 
d67688992bde
more uniform schedule_theories, notably for "present" and "commit" phase after loading;
 wenzelm parents: 
73817diff
changeset | 243 |                 {name = "theory:" ^ name, group = NONE,
 | 
| 
d67688992bde
more uniform schedule_theories, notably for "present" and "commit" phase after loading;
 wenzelm parents: 
73817diff
changeset | 244 | deps = map (Future.task_of o #2) deps, pri = 0, interrupts = true} | 
| 
d67688992bde
more uniform schedule_theories, notably for "present" and "commit" phase after loading;
 wenzelm parents: 
73817diff
changeset | 245 | (fn () => body (join_parents deps name parents)) | 
| 
d67688992bde
more uniform schedule_theories, notably for "present" and "commit" phase after loading;
 wenzelm parents: 
73817diff
changeset | 246 | else | 
| 
d67688992bde
more uniform schedule_theories, notably for "present" and "commit" phase after loading;
 wenzelm parents: 
73817diff
changeset | 247 | Future.value_result (Exn.capture (fn () => body (join_parents deps name parents)) ()) | 
| 51331 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 248 | | Finished theory => Future.value (theory_result theory))); | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 249 | |
| 73817 | 250 | val results1 = futures |> maps (consolidate_theory o Future.join_result); | 
| 51331 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 251 | |
| 73821 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 252 | val present_results = futures |> maps (present_theory o Future.join_result); | 
| 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 253 | val results2 = (map o Exn.map_res) (K ()) present_results; | 
| 51331 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 254 | |
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 255 | val results3 = futures | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51330diff
changeset | 256 | |> 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 | 257 | |
| 53192 | 258 | 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 | 259 | |
| 73821 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 260 | val _ = Par_Exn.release_all (results1 @ results2 @ results3 @ results4); | 
| 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 261 | in Par_Exn.release_all present_results end); | 
| 29429 
a6c641f08af7
schedule_futures: tuned final consolidation, explicit after_load phase;
 wenzelm parents: 
29421diff
changeset | 262 | |
| 73825 | 263 | end; | 
| 264 | ||
| 24209 | 265 | |
| 65505 | 266 | (* eval theory *) | 
| 267 | ||
| 72705 | 268 | fun eval_thy options master_dir header text_pos text parents = | 
| 65505 | 269 | let | 
| 270 | val (name, _) = #name header; | |
| 271 | val keywords = | |
| 272 | fold (curry Keyword.merge_keywords o Thy_Header.get_keywords) parents | |
| 273 | (Keyword.add_keywords (#keywords header) Keyword.empty_keywords); | |
| 274 | ||
| 68178 | 275 | val spans = Outer_Syntax.parse_spans (Token.explode keywords text_pos text); | 
| 68839 | 276 | val elements = Thy_Element.parse_elements keywords spans; | 
| 65505 | 277 | |
| 73046 | 278 | val text_id = Position.copy_id text_pos Position.none; | 
| 279 | ||
| 72622 
830222403681
HTML presentation in Isabelle/Scala, based on theory html exports from Isabelle/ML;
 wenzelm parents: 
72620diff
changeset | 280 | fun init () = Resources.begin_theory master_dir header parents; | 
| 73046 | 281 | |
| 282 | fun excursion () = | |
| 283 | let | |
| 284 | fun element_result span_elem (st, _) = | |
| 285 | let | |
| 286 | fun prepare span = | |
| 287 | let | |
| 288 | val tr = | |
| 289 | Position.setmp_thread_data text_id | |
| 290 | (fn () => Command.read_span keywords st master_dir init span) (); | |
| 291 | in Toplevel.timing (Resources.last_timing tr) tr end; | |
| 292 | val elem = Thy_Element.map_element prepare span_elem; | |
| 293 | val (results, st') = Toplevel.element_result keywords elem st; | |
| 294 | val pos' = Toplevel.pos_of (Thy_Element.last_element elem); | |
| 295 | in (results, (st', pos')) end; | |
| 296 | ||
| 297 | val (results, (end_state, end_pos)) = | |
| 76415 | 298 | fold_map element_result elements (Toplevel.make_state NONE, Position.none); | 
| 73046 | 299 | |
| 300 | val thy = Toplevel.end_theory end_pos end_state; | |
| 301 | in (results, thy) end; | |
| 302 | ||
| 303 |     val (results, thy) = cond_timeit true ("theory " ^ quote name) excursion;
 | |
| 65505 | 304 | |
| 73821 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 305 | fun present () : presentation_context = | 
| 65505 | 306 | let | 
| 73687 
54fe8cc0e1c6
clarified signature: provide access to previous state;
 wenzelm parents: 
73559diff
changeset | 307 | val segments = | 
| 
54fe8cc0e1c6
clarified signature: provide access to previous state;
 wenzelm parents: 
73559diff
changeset | 308 | (spans ~~ maps Toplevel.join_results results) |> map (fn (span, (st, tr, st')) => | 
| 
54fe8cc0e1c6
clarified signature: provide access to previous state;
 wenzelm parents: 
73559diff
changeset | 309 |             {span = span, prev_state = st, command = tr, state = st'});
 | 
| 73821 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 310 |       in {options = options, file_pos = text_pos, adjust_pos = I, segments = segments} end;
 | 
| 73046 | 311 | |
| 73798 
1ca35197108f
more predictable sequential presentation (2f9877db82a1), without somewhat pointless result_ord (e7fab0b5dbe7);
 wenzelm parents: 
73785diff
changeset | 312 | in (thy, present) end; | 
| 65505 | 313 | |
| 314 | ||
| 23967 | 315 | (* require_thy -- checking database entries wrt. the file-system *) | 
| 15065 | 316 | |
| 7211 | 317 | local | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 318 | |
| 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 | 319 | 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 | 320 | | 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 | 321 | |
| 73045 | 322 | fun load_thy options initiators deps text (name, header_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 | 323 | let | 
| 73045 | 324 |     val {master = (thy_path, _), imports} = deps;
 | 
| 325 | val dir = Path.dir thy_path; | |
| 326 | ||
| 72692 
22aeec526ffd
support for PIDE markup in batch build (inactive due to pide_reports=false);
 wenzelm parents: 
72669diff
changeset | 327 | val exec_id = Document_ID.make (); | 
| 
22aeec526ffd
support for PIDE markup in batch build (inactive due to pide_reports=false);
 wenzelm parents: 
72669diff
changeset | 328 | val id = Document_ID.print exec_id; | 
| 73045 | 329 | val put_id = Position.put_id id; | 
| 72692 
22aeec526ffd
support for PIDE markup in batch build (inactive due to pide_reports=false);
 wenzelm parents: 
72669diff
changeset | 330 | val _ = | 
| 
22aeec526ffd
support for PIDE markup in batch build (inactive due to pide_reports=false);
 wenzelm parents: 
72669diff
changeset | 331 | Execution.running Document_ID.none exec_id [] orelse | 
| 
22aeec526ffd
support for PIDE markup in batch build (inactive due to pide_reports=false);
 wenzelm parents: 
72669diff
changeset | 332 |         raise Fail ("Failed to register execution: " ^ id);
 | 
| 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 | 333 | |
| 76884 | 334 | val text_pos = put_id (Position.file (File.symbolic_path thy_path)); | 
| 72692 
22aeec526ffd
support for PIDE markup in batch build (inactive due to pide_reports=false);
 wenzelm parents: 
72669diff
changeset | 335 | val text_props = Position.properties_of text_pos; | 
| 
22aeec526ffd
support for PIDE markup in batch build (inactive due to pide_reports=false);
 wenzelm parents: 
72669diff
changeset | 336 | |
| 
22aeec526ffd
support for PIDE markup in batch build (inactive due to pide_reports=false);
 wenzelm parents: 
72669diff
changeset | 337 | val _ = remove_thy name; | 
| 
22aeec526ffd
support for PIDE markup in batch build (inactive due to pide_reports=false);
 wenzelm parents: 
72669diff
changeset | 338 |     val _ = writeln ("Loading theory " ^ quote name ^ required_by " " initiators);
 | 
| 73559 | 339 | val _ = Output.try_protocol_message (Markup.loading_theory name @ text_props) [XML.blob [text]]; | 
| 53375 
78693e46a237
Execution.fork formally requires registered Execution.running;
 wenzelm parents: 
53192diff
changeset | 340 | |
| 73045 | 341 | val _ = | 
| 342 | Position.setmp_thread_data (Position.id_only id) (fn () => | |
| 343 | (parents, map #2 imports) |> ListPair.app (fn (thy, pos) => | |
| 344 | Context_Position.reports_global thy [(put_id pos, Theory.get_markup thy)])) (); | |
| 345 | ||
| 66873 
9953ae603a23
provide theory timing information, similar to command timing but always considered relevant;
 wenzelm parents: 
66711diff
changeset | 346 | val timing_start = Timing.start (); | 
| 
9953ae603a23
provide theory timing information, similar to command timing but always considered relevant;
 wenzelm parents: 
66711diff
changeset | 347 | |
| 73045 | 348 | val header = Thy_Header.make (name, put_id header_pos) imports keywords; | 
| 73798 
1ca35197108f
more predictable sequential presentation (2f9877db82a1), without somewhat pointless result_ord (e7fab0b5dbe7);
 wenzelm parents: 
73785diff
changeset | 349 | val (theory, present) = | 
| 72705 | 350 | eval_thy options dir header text_pos text | 
| 62876 | 351 | (if name = Context.PureN then [Context.the_global_context ()] else parents); | 
| 66873 
9953ae603a23
provide theory timing information, similar to command timing but always considered relevant;
 wenzelm parents: 
66711diff
changeset | 352 | |
| 
9953ae603a23
provide theory timing information, similar to command timing but always considered relevant;
 wenzelm parents: 
66711diff
changeset | 353 | val timing_result = Timing.result timing_start; | 
| 
9953ae603a23
provide theory timing information, similar to command timing but always considered relevant;
 wenzelm parents: 
66711diff
changeset | 354 | val timing_props = [Markup.theory_timing, (Markup.nameN, name)]; | 
| 73559 | 355 | val _ = Output.try_protocol_message (timing_props @ Markup.timing_properties timing_result) [] | 
| 66873 
9953ae603a23
provide theory timing information, similar to command timing but always considered relevant;
 wenzelm parents: 
66711diff
changeset | 356 | |
| 43640 | 357 | fun commit () = update_thy deps theory; | 
| 73821 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 358 |   in Result {theory = theory, exec_id = exec_id, present = SOME o present, commit = commit} 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 | 359 | |
| 70711 | 360 | fun check_thy_deps dir name = | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 361 | (case lookup_deps name of | 
| 51293 
05b1bbae748d
discontinued obsolete 'uses' within theory header;
 wenzelm parents: 
51284diff
changeset | 362 | 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 | 363 | | NONE => | 
| 56208 | 364 |       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 | 365 | 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 | 366 |   | SOME (SOME {master, ...}) =>
 | 
| 42003 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 wenzelm parents: 
41955diff
changeset | 367 | let | 
| 48928 | 368 |         val {master = master', text = text', theory_pos = theory_pos', imports = imports',
 | 
| 56208 | 369 | keywords = keywords'} = Resources.check_thy dir name; | 
| 42003 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 wenzelm parents: 
41955diff
changeset | 370 | val deps' = SOME (make_deps master' imports', text'); | 
| 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 wenzelm parents: 
41955diff
changeset | 371 | val current = | 
| 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 wenzelm parents: 
41955diff
changeset | 372 | #2 master = #2 master' andalso | 
| 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 wenzelm parents: 
41955diff
changeset | 373 | (case lookup_theory name of | 
| 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 wenzelm parents: 
41955diff
changeset | 374 | NONE => false | 
| 56208 | 375 | | SOME theory => Resources.loaded_files_current theory); | 
| 51293 
05b1bbae748d
discontinued obsolete 'uses' within theory header;
 wenzelm parents: 
51284diff
changeset | 376 | in (current, deps', theory_pos', imports', keywords') end); | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 377 | |
| 73817 | 378 | fun task_finished (Task _) = false | 
| 379 | | task_finished (Finished _) = true; | |
| 380 | ||
| 23967 | 381 | in | 
| 382 | ||
| 72638 
2a7fc87495e0
refer to command_timings/last_timing via resources;
 wenzelm parents: 
72624diff
changeset | 383 | fun require_thys options initiators qualifier dir strs tasks = | 
| 
2a7fc87495e0
refer to command_timings/last_timing via resources;
 wenzelm parents: 
72624diff
changeset | 384 | fold_map (require_thy options initiators qualifier dir) strs tasks |>> forall I | 
| 
2a7fc87495e0
refer to command_timings/last_timing via resources;
 wenzelm parents: 
72624diff
changeset | 385 | and require_thy options initiators qualifier dir (s, require_pos) tasks = | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 386 | let | 
| 66711 
80fa1401cf76
discontinued extra checks (see ce676a750575 and 60c159d490a2) -- qualified theory names are meant to cover this;
 wenzelm parents: 
66377diff
changeset | 387 |     val {master_dir, theory_name, ...} = Resources.import_name qualifier dir s;
 | 
| 7066 | 388 | in | 
| 65443 
dccbfc715904
provide Resources.import_name in ML, similar to Scala version;
 wenzelm parents: 
65058diff
changeset | 389 | (case try (String_Graph.get_node tasks) theory_name of | 
| 66711 
80fa1401cf76
discontinued extra checks (see ce676a750575 and 60c159d490a2) -- qualified theory names are meant to cover this;
 wenzelm parents: 
66377diff
changeset | 390 | SOME task => (task_finished task, tasks) | 
| 23967 | 391 | | NONE => | 
| 392 | let | |
| 65443 
dccbfc715904
provide Resources.import_name in ML, similar to Scala version;
 wenzelm parents: 
65058diff
changeset | 393 | val _ = member (op =) initiators theory_name andalso error (cycle_msg initiators); | 
| 43651 
511df47bcadc
some support for theory files within Isabelle/Scala session;
 wenzelm parents: 
43641diff
changeset | 394 | |
| 70711 | 395 | val (current, deps, theory_pos, imports, keywords) = check_thy_deps master_dir theory_name | 
| 55797 | 396 | handle ERROR msg => | 
| 397 | cat_error msg | |
| 65443 
dccbfc715904
provide Resources.import_name in ML, similar to Scala version;
 wenzelm parents: 
65058diff
changeset | 398 |                 ("The error(s) above occurred for theory " ^ quote theory_name ^
 | 
| 55797 | 399 | 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 | 400 | |
| 65455 | 401 | val qualifier' = Resources.theory_qualifier theory_name; | 
| 72511 
460d743010bc
clarified signature: overloaded "+" for Path.append;
 wenzelm parents: 
72311diff
changeset | 402 | val dir' = dir + master_dir_deps (Option.map #1 deps); | 
| 65455 | 403 | |
| 404 | val parents = map (#theory_name o Resources.import_name qualifier' dir' 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 | 405 | val (parents_current, tasks') = | 
| 72638 
2a7fc87495e0
refer to command_timings/last_timing via resources;
 wenzelm parents: 
72624diff
changeset | 406 | require_thys options (theory_name :: initiators) qualifier' dir' imports tasks; | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 407 | |
| 23967 | 408 | 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 | 409 | val task = | 
| 65443 
dccbfc715904
provide Resources.import_name in ML, similar to Scala version;
 wenzelm parents: 
65058diff
changeset | 410 | if all_current then Finished (get_theory 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 | 411 | else | 
| 37980 
b8c3d4dc1e3e
avoid repeated File.read for theory text (as before);
 wenzelm parents: 
37979diff
changeset | 412 | (case deps of | 
| 
b8c3d4dc1e3e
avoid repeated File.read for theory text (as before);
 wenzelm parents: 
37979diff
changeset | 413 | NONE => raise Fail "Malformed deps" | 
| 42003 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 wenzelm parents: 
41955diff
changeset | 414 | | SOME (dep, text) => | 
| 72705 | 415 | Task (parents, | 
| 416 | load_thy options initiators dep text (theory_name, theory_pos) keywords)); | |
| 48874 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48710diff
changeset | 417 | |
| 65443 
dccbfc715904
provide Resources.import_name in ML, similar to Scala version;
 wenzelm parents: 
65058diff
changeset | 418 | val tasks'' = new_entry theory_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 | 419 | in (all_current, tasks'') end) | 
| 7066 | 420 | end; | 
| 6484 
3f098b0ec683
use_thy etc.: may specify path prefix, which is temporarily used as load path;
 wenzelm parents: 
6389diff
changeset | 421 | |
| 23967 | 422 | end; | 
| 423 | ||
| 424 | ||
| 65444 | 425 | (* use theories *) | 
| 23967 | 426 | |
| 72638 
2a7fc87495e0
refer to command_timings/last_timing via resources;
 wenzelm parents: 
72624diff
changeset | 427 | fun use_theories options qualifier imports = | 
| 73818 
d67688992bde
more uniform schedule_theories, notably for "present" and "commit" phase after loading;
 wenzelm parents: 
73817diff
changeset | 428 | schedule_theories (#2 (require_thys options [] qualifier Path.current imports String_Graph.empty)); | 
| 23967 | 429 | |
| 65444 | 430 | fun use_thy name = | 
| 73821 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73819diff
changeset | 431 | ignore (use_theories (Options.default ()) Resources.default_qualifier [(name, Position.none)]); | 
| 7211 | 432 | |
| 59364 | 433 | |
| 57626 | 434 | (* toplevel scripting -- without maintaining database *) | 
| 435 | ||
| 58856 | 436 | fun script_thy pos txt thy = | 
| 57626 | 437 | let | 
| 69891 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69886diff
changeset | 438 | val trs = Outer_Syntax.parse_text thy (K thy) pos txt; | 
| 57626 | 439 | val end_pos = if null trs then pos else Toplevel.pos_of (List.last trs); | 
| 76415 | 440 | val end_state = fold (Toplevel.command_exception true) trs (Toplevel.make_state NONE); | 
| 57626 | 441 | in Toplevel.end_theory end_pos end_state end; | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 442 | |
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 443 | |
| 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 | 444 | (* register theory *) | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 445 | |
| 37954 
a2e73df0b1e0
simplified/clarified register_thy: more precise treatment of new dependencies, remove descendants;
 wenzelm parents: 
37953diff
changeset | 446 | fun register_thy theory = | 
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 447 | let | 
| 65459 
da59e8e0a663
clarified theory_long_name (for qualified access to Thy_Info) vs. short theory_name (which is unique within any given theory context);
 wenzelm parents: 
65455diff
changeset | 448 | val name = Context.theory_long_name theory; | 
| 56208 | 449 |     val {master, ...} = Resources.check_thy (Resources.master_directory theory) name;
 | 
| 450 | val imports = Resources.imports_of theory; | |
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 451 | in | 
| 59178 | 452 | change_thys (fn thys => | 
| 453 | let | |
| 454 | val thys' = remove name thys; | |
| 455 |         val _ = writeln ("Registering theory " ^ quote name);
 | |
| 456 | in update (make_deps master imports) theory thys' end) | |
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 457 | end; | 
| 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 458 | |
| 24080 
8c67d869531b
added register_thy (replaces pretend_use_thy_only and really flag);
 wenzelm parents: 
24071diff
changeset | 459 | |
| 
8c67d869531b
added register_thy (replaces pretend_use_thy_only and really flag);
 wenzelm parents: 
24071diff
changeset | 460 | (* finish all theories *) | 
| 
8c67d869531b
added register_thy (replaces pretend_use_thy_only and really flag);
 wenzelm parents: 
24071diff
changeset | 461 | |
| 50862 
5fc8b83322f5
more sensible order of theory nodes (correspondance to Scala version), e.g. relevant to theory progress;
 wenzelm parents: 
50845diff
changeset | 462 | 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 | 463 | |
| 6211 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 wenzelm parents: 
5211diff
changeset | 464 | end; | 
| 62847 | 465 | |
| 65444 | 466 | fun use_thy name = Runtime.toplevel_program (fn () => Thy_Info.use_thy name); |