| author | wenzelm | 
| Fri, 09 Mar 2018 14:30:13 +0100 | |
| changeset 67795 | f8037c7e4659 | 
| parent 67380 | 8bef51521f21 | 
| child 68130 | 6fb85346cb79 | 
| 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: 
5211 
diff
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: 
37977 
diff
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: 
5211 
diff
changeset
 | 
12  | 
val get_theory: string -> theory  | 
| 26983 | 13  | 
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: 
29118 
diff
changeset
 | 
14  | 
val remove_thy: string -> unit  | 
| 
59366
 
e94df7f6b608
clarified build_theories: proper protocol handler;
 
wenzelm 
parents: 
59364 
diff
changeset
 | 
15  | 
val use_theories:  | 
| 61381 | 16  | 
    {document: bool,
 | 
17  | 
symbols: HTML.symbols,  | 
|
| 
67297
 
86a099f896fc
formal check of @{cite} bibtex entries -- only in batch-mode session builds;
 
wenzelm 
parents: 
67194 
diff
changeset
 | 
18  | 
bibtex_entries: string list,  | 
| 
65058
 
3e9f382fb67e
absent timing information means zero, according to 0070053570c4, f235646b1b73;
 
wenzelm 
parents: 
64574 
diff
changeset
 | 
19  | 
last_timing: Toplevel.transition -> Time.time,  | 
| 
65445
 
e9e7f5f5794c
more qualifier treatment, but in the end it is still ignored;
 
wenzelm 
parents: 
65444 
diff
changeset
 | 
20  | 
qualifier: string,  | 
| 61381 | 21  | 
master_dir: Path.T} -> (string * Position.T) list -> unit  | 
| 65444 | 22  | 
val use_thy: string -> 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: 
37953 
diff
changeset
 | 
24  | 
val register_thy: theory -> unit  | 
| 
24080
 
8c67d869531b
added register_thy (replaces pretend_use_thy_only and really flag);
 
wenzelm 
parents: 
24071 
diff
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: 
36953 
diff
changeset
 | 
28  | 
structure Thy_Info: THY_INFO =  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
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: 
5211 
diff
changeset
 | 
31  | 
(** thy database **)  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
32  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
33  | 
(* messages *)  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
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: 
5211 
diff
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: 
50845 
diff
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: 
50845 
diff
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: 
37981 
diff
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: 
50845 
diff
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: 
37955 
diff
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: 
5211 
diff
changeset
 | 
50  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
51  | 
type deps =  | 
| 
41955
 
703ea96b13c6
files are identified via SHA1 digests -- discontinued ISABELLE_FILE_IDENT;
 
wenzelm 
parents: 
41678 
diff
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: 
48898 
diff
changeset
 | 
53  | 
imports: (string * Position.T) list}; (*source specification of imports (partially qualified)*)  | 
| 
23886
 
f40fba467384
simplified ThyLoad interfaces: only one additional directory;
 
wenzelm 
parents: 
23871 
diff
changeset
 | 
54  | 
|
| 
38148
 
d2f3c8d4a89f
uniform naming of imports (source specification) vs. parents (thy node names) vs. parent_thys (theory values);
 
wenzelm 
parents: 
38147 
diff
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: 
5211 
diff
changeset
 | 
56  | 
|
| 65454 | 57  | 
fun master_dir_deps (d: deps option) =  | 
| 59178 | 58  | 
the_default Path.current (Option.map (Path.dir o #1 o #master) d);  | 
59  | 
||
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
60  | 
local  | 
| 59178 | 61  | 
val global_thys =  | 
62  | 
Synchronized.var "Thy_Info.thys"  | 
|
63  | 
(String_Graph.empty: (deps option * theory option) String_Graph.T);  | 
|
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
64  | 
in  | 
| 59178 | 65  | 
fun get_thys () = Synchronized.value global_thys;  | 
66  | 
fun change_thys f = Synchronized.change global_thys f;  | 
|
| 
3604
 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 
berghofe 
parents:  
diff
changeset
 | 
67  | 
end;  | 
| 5209 | 68  | 
|
| 
50862
 
5fc8b83322f5
more sensible order of theory nodes (correspondance to Scala version), e.g. relevant to theory progress;
 
wenzelm 
parents: 
50845 
diff
changeset
 | 
69  | 
fun get_names () = String_Graph.topological_order (get_thys ());  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
70  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
71  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
72  | 
(* access thy *)  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
73  | 
|
| 59178 | 74  | 
fun lookup thys name = try (String_Graph.get_node thys) name;  | 
75  | 
fun lookup_thy name = lookup (get_thys ()) name;  | 
|
| 7935 | 76  | 
|
| 59178 | 77  | 
fun get thys name =  | 
78  | 
(case lookup thys name of  | 
|
| 15531 | 79  | 
SOME thy => thy  | 
| 55797 | 80  | 
  | NONE => error ("Theory loader: nothing known about theory " ^ quote name));
 | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
81  | 
|
| 59178 | 82  | 
fun get_thy name = get (get_thys ()) name;  | 
83  | 
||
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
84  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
85  | 
(* access deps *)  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
86  | 
|
| 15570 | 87  | 
val lookup_deps = Option.map #1 o lookup_thy;  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
88  | 
|
| 65454 | 89  | 
val master_directory = master_dir_deps o #1 o get_thy;  | 
| 7191 | 90  | 
|
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
91  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
92  | 
(* access theory *)  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
93  | 
|
| 7687 | 94  | 
fun lookup_theory name =  | 
95  | 
(case lookup_thy name of  | 
|
| 38142 | 96  | 
SOME (_, SOME theory) => SOME theory  | 
| 15531 | 97  | 
| _ => NONE);  | 
| 7288 | 98  | 
|
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
99  | 
fun get_theory name =  | 
| 7288 | 100  | 
(case lookup_theory name of  | 
| 23871 | 101  | 
SOME theory => theory  | 
| 55797 | 102  | 
  | _ => error ("Theory loader: undefined entry for theory " ^ quote name));
 | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
103  | 
|
| 56208 | 104  | 
val get_imports = Resources.imports_of o get_theory;  | 
| 
44337
 
d453faed4815
clarified get_imports -- should not rely on accidental order within graph;
 
wenzelm 
parents: 
44302 
diff
changeset
 | 
105  | 
|
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
106  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
107  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
108  | 
(** thy operations **)  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
109  | 
|
| 59178 | 110  | 
(* remove *)  | 
| 
29421
 
db532e37cda2
use_thys: perform consolidate_thy on loaded theories, which removes failed nodes in post-hoc fashion;
 
wenzelm 
parents: 
29118 
diff
changeset
 | 
111  | 
|
| 59178 | 112  | 
fun remove name thys =  | 
113  | 
(case lookup thys name of  | 
|
114  | 
NONE => thys  | 
|
115  | 
  | SOME (NONE, _) => error ("Cannot update finished theory " ^ quote name)
 | 
|
116  | 
| SOME _ =>  | 
|
117  | 
let  | 
|
118  | 
val succs = String_Graph.all_succs thys [name];  | 
|
119  | 
        val _ = writeln ("Theory loader: removing " ^ commas_quote succs);
 | 
|
120  | 
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: 
29118 
diff
changeset
 | 
121  | 
|
| 59178 | 122  | 
val remove_thy = change_thys o remove;  | 
123  | 
||
| 7099 | 124  | 
|
| 59178 | 125  | 
(* update *)  | 
126  | 
||
127  | 
fun update deps theory thys =  | 
|
| 43640 | 128  | 
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: 
65455 
diff
changeset
 | 
129  | 
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: 
65455 
diff
changeset
 | 
130  | 
val parents = map Context.theory_long_name (Theory.parents_of theory);  | 
| 59178 | 131  | 
|
132  | 
val thys' = remove name thys;  | 
|
133  | 
val _ = map (get thys') parents;  | 
|
134  | 
in new_entry name parents (SOME deps, SOME theory) thys' end;  | 
|
135  | 
||
136  | 
fun update_thy deps theory = change_thys (update deps theory);  | 
|
| 43640 | 137  | 
|
| 7099 | 138  | 
|
| 24209 | 139  | 
(* scheduling loader tasks *)  | 
140  | 
||
| 51330 | 141  | 
datatype result =  | 
| 
53375
 
78693e46a237
Execution.fork formally requires registered Execution.running;
 
wenzelm 
parents: 
53192 
diff
changeset
 | 
142  | 
  Result of {theory: theory, exec_id: Document_ID.exec,
 | 
| 
 
78693e46a237
Execution.fork formally requires registered Execution.running;
 
wenzelm 
parents: 
53192 
diff
changeset
 | 
143  | 
present: unit -> unit, commit: unit -> unit, weight: int};  | 
| 
51331
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
144  | 
|
| 
53375
 
78693e46a237
Execution.fork formally requires registered Execution.running;
 
wenzelm 
parents: 
53192 
diff
changeset
 | 
145  | 
fun theory_result theory =  | 
| 
 
78693e46a237
Execution.fork formally requires registered Execution.running;
 
wenzelm 
parents: 
53192 
diff
changeset
 | 
146  | 
  Result {theory = theory, exec_id = Document_ID.none, present = I, commit = I, weight = 0};
 | 
| 51330 | 147  | 
|
148  | 
fun result_theory (Result {theory, ...}) = theory;
 | 
|
| 
51331
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
149  | 
fun result_present (Result {present, ...}) = present;
 | 
| 
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
150  | 
fun result_commit (Result {commit, ...}) = commit;
 | 
| 
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
151  | 
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: 
51330 
diff
changeset
 | 
152  | 
|
| 
53375
 
78693e46a237
Execution.fork formally requires registered Execution.running;
 
wenzelm 
parents: 
53192 
diff
changeset
 | 
153  | 
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: 
53375 
diff
changeset
 | 
154  | 
let  | 
| 
66371
 
6ce1afc01040
more thorough Execution.join, under the assumption that nested Execution.fork only happens from given exed_ids;
 
wenzelm 
parents: 
66370 
diff
changeset
 | 
155  | 
val _ = Execution.join [exec_id];  | 
| 
64574
 
1134e4d5e5b7
consolidate nested thms with persistent result, for improved performance;
 
wenzelm 
parents: 
62942 
diff
changeset
 | 
156  | 
val res = Exn.capture Thm.consolidate_theory theory;  | 
| 66377 | 157  | 
val exns = maps Task_Queue.group_status (Execution.peek exec_id);  | 
158  | 
in res :: map Exn.Exn exns end;  | 
|
| 
43641
 
081e009549dc
uniform finish_thy -- always Global_Theory.join_proofs, even with sequential scheduling;
 
wenzelm 
parents: 
43640 
diff
changeset
 | 
159  | 
|
| 
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: 
37955 
diff
changeset
 | 
160  | 
datatype task =  | 
| 
66711
 
80fa1401cf76
discontinued extra checks (see ce676a750575 and 60c159d490a2) -- qualified theory names are meant to cover this;
 
wenzelm 
parents: 
66377 
diff
changeset
 | 
161  | 
Task of string list * (theory list -> result) |  | 
| 
37977
 
3ceccd415145
simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
 
wenzelm 
parents: 
37955 
diff
changeset
 | 
162  | 
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: 
37955 
diff
changeset
 | 
163  | 
|
| 
 
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: 
37955 
diff
changeset
 | 
164  | 
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: 
37955 
diff
changeset
 | 
165  | 
| task_finished (Finished _) = true;  | 
| 24209 | 166  | 
|
| 44162 | 167  | 
fun task_parents deps (parents: string list) = map (the o AList.lookup (op =) deps) parents;  | 
168  | 
||
169  | 
val schedule_seq =  | 
|
| 
50862
 
5fc8b83322f5
more sensible order of theory nodes (correspondance to Scala version), e.g. relevant to theory progress;
 
wenzelm 
parents: 
50845 
diff
changeset
 | 
170  | 
String_Graph.schedule (fn deps => fn (_, task) =>  | 
| 44162 | 171  | 
(case task of  | 
| 
66711
 
80fa1401cf76
discontinued extra checks (see ce676a750575 and 60c159d490a2) -- qualified theory names are meant to cover this;
 
wenzelm 
parents: 
66377 
diff
changeset
 | 
172  | 
Task (parents, body) =>  | 
| 
51331
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
173  | 
let  | 
| 
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
174  | 
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: 
51661 
diff
changeset
 | 
175  | 
val _ = Par_Exn.release_all (join_theory result);  | 
| 
51331
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
176  | 
val _ = result_present result ();  | 
| 
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
177  | 
val _ = result_commit result ();  | 
| 
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
178  | 
in result_theory result end  | 
| 44162 | 179  | 
| 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: 
28980 
diff
changeset
 | 
180  | 
|
| 62923 | 181  | 
val schedule_futures = Thread_Attributes.uninterruptible (fn _ => fn tasks =>  | 
| 
51284
 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 
wenzelm 
parents: 
51228 
diff
changeset
 | 
182  | 
let  | 
| 
51331
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
183  | 
val futures = tasks  | 
| 
51284
 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 
wenzelm 
parents: 
51228 
diff
changeset
 | 
184  | 
|> String_Graph.schedule (fn deps => fn (name, task) =>  | 
| 
 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 
wenzelm 
parents: 
51228 
diff
changeset
 | 
185  | 
(case task of  | 
| 
66711
 
80fa1401cf76
discontinued extra checks (see ce676a750575 and 60c159d490a2) -- qualified theory names are meant to cover this;
 
wenzelm 
parents: 
66377 
diff
changeset
 | 
186  | 
Task (parents, body) =>  | 
| 
51284
 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 
wenzelm 
parents: 
51228 
diff
changeset
 | 
187  | 
(singleton o Future.forks)  | 
| 
 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 
wenzelm 
parents: 
51228 
diff
changeset
 | 
188  | 
              {name = "theory:" ^ name, group = NONE,
 | 
| 
 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 
wenzelm 
parents: 
51228 
diff
changeset
 | 
189  | 
deps = map (Future.task_of o #2) deps, pri = 0, interrupts = true}  | 
| 
 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 
wenzelm 
parents: 
51228 
diff
changeset
 | 
190  | 
(fn () =>  | 
| 
 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 
wenzelm 
parents: 
51228 
diff
changeset
 | 
191  | 
(case filter (not o can Future.join o #2) deps of  | 
| 51330 | 192  | 
[] => body (map (result_theory o Future.join) (task_parents deps parents))  | 
| 
51284
 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 
wenzelm 
parents: 
51228 
diff
changeset
 | 
193  | 
| bad =>  | 
| 55797 | 194  | 
error  | 
195  | 
                      ("Failed to load theory " ^ quote name ^
 | 
|
196  | 
" (unresolved " ^ commas_quote (map #1 bad) ^ ")")))  | 
|
| 
51331
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
197  | 
| Finished theory => Future.value (theory_result theory)));  | 
| 
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
198  | 
|
| 
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
199  | 
val results1 = futures  | 
| 
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
200  | 
|> maps (fn future =>  | 
| 
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
201  | 
(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: 
51661 
diff
changeset
 | 
202  | 
Exn.Res result => join_theory result  | 
| 
51331
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
203  | 
| Exn.Exn exn => [Exn.Exn exn]));  | 
| 
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
204  | 
|
| 
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
205  | 
val results2 = futures  | 
| 
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
206  | 
|> map_filter (Exn.get_res o Future.join_result)  | 
| 
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
207  | 
|> sort result_ord  | 
| 
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
208  | 
|> Par_List.map (fn result => Exn.capture (result_present result) ());  | 
| 
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
209  | 
|
| 
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
210  | 
(* FIXME more precise commit order (!?) *)  | 
| 
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
211  | 
val results3 = futures  | 
| 
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
212  | 
|> map (fn future => Exn.capture (fn () => result_commit (Future.join future) ()) ());  | 
| 
51284
 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 
wenzelm 
parents: 
51228 
diff
changeset
 | 
213  | 
|
| 
54370
 
39ac1a02c60c
join all theory body forks, notably Toplevel.atom_result (diagnostic commands), before peeking at full status;
 
wenzelm 
parents: 
53375 
diff
changeset
 | 
214  | 
(* FIXME avoid global Execution.reset (!??) *)  | 
| 53192 | 215  | 
val results4 = map Exn.Exn (maps Task_Queue.group_status (Execution.reset ()));  | 
| 
51284
 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 
wenzelm 
parents: 
51228 
diff
changeset
 | 
216  | 
|
| 
51331
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51330 
diff
changeset
 | 
217  | 
val _ = Par_Exn.release_all (results1 @ results2 @ results3 @ results4);  | 
| 
51284
 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 
wenzelm 
parents: 
51228 
diff
changeset
 | 
218  | 
in () end);  | 
| 
29429
 
a6c641f08af7
schedule_futures: tuned final consolidation, explicit after_load phase;
 
wenzelm 
parents: 
29421 
diff
changeset
 | 
219  | 
|
| 24209 | 220  | 
|
| 65505 | 221  | 
(* eval theory *)  | 
222  | 
||
223  | 
fun excursion keywords master_dir last_timing init elements =  | 
|
224  | 
let  | 
|
225  | 
fun prepare_span st span =  | 
|
226  | 
Command_Span.content span  | 
|
227  | 
|> Command.read keywords (Command.read_thy st) master_dir init ([], ~1)  | 
|
228  | 
|> (fn tr => Toplevel.put_timing (last_timing tr) tr);  | 
|
229  | 
||
230  | 
fun element_result span_elem (st, _) =  | 
|
231  | 
let  | 
|
232  | 
val elem = Thy_Syntax.map_element (prepare_span st) span_elem;  | 
|
233  | 
val (results, st') = Toplevel.element_result keywords elem st;  | 
|
234  | 
val pos' = Toplevel.pos_of (Thy_Syntax.last_element elem);  | 
|
235  | 
in (results, (st', pos')) end;  | 
|
236  | 
||
237  | 
val (results, (end_state, end_pos)) =  | 
|
238  | 
fold_map element_result elements (Toplevel.toplevel, Position.none);  | 
|
239  | 
||
240  | 
val thy = Toplevel.end_theory end_pos end_state;  | 
|
241  | 
in (results, thy) end;  | 
|
242  | 
||
| 
67297
 
86a099f896fc
formal check of @{cite} bibtex entries -- only in batch-mode session builds;
 
wenzelm 
parents: 
67194 
diff
changeset
 | 
243  | 
fun eval_thy {document, symbols, bibtex_entries, last_timing} update_time master_dir header
 | 
| 
 
86a099f896fc
formal check of @{cite} bibtex entries -- only in batch-mode session builds;
 
wenzelm 
parents: 
67194 
diff
changeset
 | 
244  | 
text_pos text parents =  | 
| 65505 | 245  | 
let  | 
246  | 
val (name, _) = #name header;  | 
|
247  | 
val keywords =  | 
|
248  | 
fold (curry Keyword.merge_keywords o Thy_Header.get_keywords) parents  | 
|
249  | 
(Keyword.add_keywords (#keywords header) Keyword.empty_keywords);  | 
|
250  | 
||
251  | 
val toks = Token.explode keywords text_pos text;  | 
|
252  | 
val spans = Outer_Syntax.parse_spans toks;  | 
|
253  | 
val elements = Thy_Syntax.parse_elements keywords spans;  | 
|
254  | 
||
255  | 
fun init () =  | 
|
256  | 
Resources.begin_theory master_dir header parents  | 
|
| 
67297
 
86a099f896fc
formal check of @{cite} bibtex entries -- only in batch-mode session builds;
 
wenzelm 
parents: 
67194 
diff
changeset
 | 
257  | 
|> Present.begin_theory bibtex_entries update_time  | 
| 65505 | 258  | 
(fn () => implode (map (HTML.present_span symbols keywords) spans));  | 
259  | 
||
260  | 
val (results, thy) =  | 
|
261  | 
      cond_timeit true ("theory " ^ quote name)
 | 
|
262  | 
(fn () => excursion keywords master_dir last_timing init elements);  | 
|
263  | 
||
264  | 
fun present () =  | 
|
265  | 
let  | 
|
266  | 
val res = filter_out (Toplevel.is_ignored o #1) (maps Toplevel.join_results results);  | 
|
267  | 
in  | 
|
| 67163 | 268  | 
if exists (Toplevel.is_skipped_proof o #2) res then ()  | 
| 65505 | 269  | 
else  | 
| 
67194
 
1c0a6a957114
positions as postlude: avoid intrusion of odd %-forms into main tex source;
 
wenzelm 
parents: 
67173 
diff
changeset
 | 
270  | 
let val body = Thy_Output.present_thy thy res toks;  | 
| 
 
1c0a6a957114
positions as postlude: avoid intrusion of odd %-forms into main tex source;
 
wenzelm 
parents: 
67173 
diff
changeset
 | 
271  | 
in if document then Present.theory_output text_pos thy body else () end  | 
| 65505 | 272  | 
end;  | 
273  | 
||
274  | 
in (thy, present, size text) end;  | 
|
275  | 
||
276  | 
||
| 23967 | 277  | 
(* require_thy -- checking database entries wrt. the file-system *)  | 
| 15065 | 278  | 
|
| 7211 | 279  | 
local  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
280  | 
|
| 
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: 
37955 
diff
changeset
 | 
281  | 
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: 
37955 
diff
changeset
 | 
282  | 
| 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: 
37955 
diff
changeset
 | 
283  | 
|
| 
67297
 
86a099f896fc
formal check of @{cite} bibtex entries -- only in batch-mode session builds;
 
wenzelm 
parents: 
67194 
diff
changeset
 | 
284  | 
fun load_thy context 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: 
37955 
diff
changeset
 | 
285  | 
let  | 
| 59178 | 286  | 
val _ = remove_thy name;  | 
| 58843 | 287  | 
    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: 
56208 
diff
changeset
 | 
288  | 
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: 
37955 
diff
changeset
 | 
289  | 
|
| 
41548
 
bd0bebf93fa6
Thy_Load.begin_theory: maintain source specification of imports;
 
wenzelm 
parents: 
41536 
diff
changeset
 | 
290  | 
    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: 
37985 
diff
changeset
 | 
291  | 
val dir = Path.dir thy_path;  | 
| 
51294
 
0850d43cb355
discontinued obsolete header "files" -- these are loaded explicitly after exploring dependencies;
 
wenzelm 
parents: 
51293 
diff
changeset
 | 
292  | 
val header = Thy_Header.make (name, pos) imports keywords;  | 
| 
48927
 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 
wenzelm 
parents: 
48898 
diff
changeset
 | 
293  | 
|
| 
 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 
wenzelm 
parents: 
48898 
diff
changeset
 | 
294  | 
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: 
37955 
diff
changeset
 | 
295  | 
|
| 
53375
 
78693e46a237
Execution.fork formally requires registered Execution.running;
 
wenzelm 
parents: 
53192 
diff
changeset
 | 
296  | 
val exec_id = Document_ID.make ();  | 
| 
 
78693e46a237
Execution.fork formally requires registered Execution.running;
 
wenzelm 
parents: 
53192 
diff
changeset
 | 
297  | 
val _ =  | 
| 
 
78693e46a237
Execution.fork formally requires registered Execution.running;
 
wenzelm 
parents: 
53192 
diff
changeset
 | 
298  | 
Execution.running Document_ID.none exec_id [] orelse  | 
| 
 
78693e46a237
Execution.fork formally requires registered Execution.running;
 
wenzelm 
parents: 
53192 
diff
changeset
 | 
299  | 
        raise Fail ("Failed to register execution: " ^ Document_ID.print exec_id);
 | 
| 
 
78693e46a237
Execution.fork formally requires registered Execution.running;
 
wenzelm 
parents: 
53192 
diff
changeset
 | 
300  | 
|
| 
66873
 
9953ae603a23
provide theory timing information, similar to command timing but always considered relevant;
 
wenzelm 
parents: 
66711 
diff
changeset
 | 
301  | 
val timing_start = Timing.start ();  | 
| 
 
9953ae603a23
provide theory timing information, similar to command timing but always considered relevant;
 
wenzelm 
parents: 
66711 
diff
changeset
 | 
302  | 
|
| 
53375
 
78693e46a237
Execution.fork formally requires registered Execution.running;
 
wenzelm 
parents: 
53192 
diff
changeset
 | 
303  | 
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: 
51330 
diff
changeset
 | 
304  | 
val (theory, present, weight) =  | 
| 
67297
 
86a099f896fc
formal check of @{cite} bibtex entries -- only in batch-mode session builds;
 
wenzelm 
parents: 
67194 
diff
changeset
 | 
305  | 
eval_thy context update_time dir header text_pos text  | 
| 62876 | 306  | 
(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: 
66711 
diff
changeset
 | 
307  | 
|
| 
 
9953ae603a23
provide theory timing information, similar to command timing but always considered relevant;
 
wenzelm 
parents: 
66711 
diff
changeset
 | 
308  | 
val timing_result = Timing.result timing_start;  | 
| 
 
9953ae603a23
provide theory timing information, similar to command timing but always considered relevant;
 
wenzelm 
parents: 
66711 
diff
changeset
 | 
309  | 
val timing_props = [Markup.theory_timing, (Markup.nameN, name)];  | 
| 
 
9953ae603a23
provide theory timing information, similar to command timing but always considered relevant;
 
wenzelm 
parents: 
66711 
diff
changeset
 | 
310  | 
val _ = Output.try_protocol_message (timing_props @ Markup.timing_properties timing_result) []  | 
| 
 
9953ae603a23
provide theory timing information, similar to command timing but always considered relevant;
 
wenzelm 
parents: 
66711 
diff
changeset
 | 
311  | 
|
| 43640 | 312  | 
fun commit () = update_thy deps theory;  | 
| 
53375
 
78693e46a237
Execution.fork formally requires registered Execution.running;
 
wenzelm 
parents: 
53192 
diff
changeset
 | 
313  | 
in  | 
| 
 
78693e46a237
Execution.fork formally requires registered Execution.running;
 
wenzelm 
parents: 
53192 
diff
changeset
 | 
314  | 
    Result {theory = theory, exec_id = exec_id, present = present, commit = commit, weight = weight}
 | 
| 
 
78693e46a237
Execution.fork formally requires registered Execution.running;
 
wenzelm 
parents: 
53192 
diff
changeset
 | 
315  | 
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: 
37955 
diff
changeset
 | 
316  | 
|
| 
48898
 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 
wenzelm 
parents: 
48888 
diff
changeset
 | 
317  | 
fun check_deps dir name =  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
318  | 
(case lookup_deps name of  | 
| 
51293
 
05b1bbae748d
discontinued obsolete 'uses' within theory header;
 
wenzelm 
parents: 
51284 
diff
changeset
 | 
319  | 
SOME NONE => (true, NONE, Position.none, get_imports name, [])  | 
| 
23893
 
8babfcaaf129
deps: maintain source specification of parents (prevents repeated ThyLoad.deps_thy);
 
wenzelm 
parents: 
23886 
diff
changeset
 | 
320  | 
| NONE =>  | 
| 56208 | 321  | 
      let val {master, text, theory_pos, imports, keywords} = Resources.check_thy dir name
 | 
| 
51293
 
05b1bbae748d
discontinued obsolete 'uses' within theory header;
 
wenzelm 
parents: 
51284 
diff
changeset
 | 
322  | 
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: 
42003 
diff
changeset
 | 
323  | 
  | SOME (SOME {master, ...}) =>
 | 
| 
42003
 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 
wenzelm 
parents: 
41955 
diff
changeset
 | 
324  | 
let  | 
| 48928 | 325  | 
        val {master = master', text = text', theory_pos = theory_pos', imports = imports',
 | 
| 56208 | 326  | 
keywords = keywords'} = Resources.check_thy dir name;  | 
| 
42003
 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 
wenzelm 
parents: 
41955 
diff
changeset
 | 
327  | 
val deps' = SOME (make_deps master' imports', text');  | 
| 
 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 
wenzelm 
parents: 
41955 
diff
changeset
 | 
328  | 
val current =  | 
| 
 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 
wenzelm 
parents: 
41955 
diff
changeset
 | 
329  | 
#2 master = #2 master' andalso  | 
| 
 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 
wenzelm 
parents: 
41955 
diff
changeset
 | 
330  | 
(case lookup_theory name of  | 
| 
 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 
wenzelm 
parents: 
41955 
diff
changeset
 | 
331  | 
NONE => false  | 
| 56208 | 332  | 
| SOME theory => Resources.loaded_files_current theory);  | 
| 
51293
 
05b1bbae748d
discontinued obsolete 'uses' within theory header;
 
wenzelm 
parents: 
51284 
diff
changeset
 | 
333  | 
in (current, deps', theory_pos', imports', keywords') end);  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
334  | 
|
| 23967 | 335  | 
in  | 
336  | 
||
| 
67297
 
86a099f896fc
formal check of @{cite} bibtex entries -- only in batch-mode session builds;
 
wenzelm 
parents: 
67194 
diff
changeset
 | 
337  | 
fun require_thys context initiators qualifier dir strs tasks =  | 
| 
 
86a099f896fc
formal check of @{cite} bibtex entries -- only in batch-mode session builds;
 
wenzelm 
parents: 
67194 
diff
changeset
 | 
338  | 
fold_map (require_thy context initiators qualifier dir) strs tasks |>> forall I  | 
| 
 
86a099f896fc
formal check of @{cite} bibtex entries -- only in batch-mode session builds;
 
wenzelm 
parents: 
67194 
diff
changeset
 | 
339  | 
and require_thy context initiators qualifier dir (s, require_pos) tasks =  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
340  | 
let  | 
| 
66711
 
80fa1401cf76
discontinued extra checks (see ce676a750575 and 60c159d490a2) -- qualified theory names are meant to cover this;
 
wenzelm 
parents: 
66377 
diff
changeset
 | 
341  | 
    val {master_dir, theory_name, ...} = Resources.import_name qualifier dir s;
 | 
| 7066 | 342  | 
in  | 
| 
65443
 
dccbfc715904
provide Resources.import_name in ML, similar to Scala version;
 
wenzelm 
parents: 
65058 
diff
changeset
 | 
343  | 
(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: 
66377 
diff
changeset
 | 
344  | 
SOME task => (task_finished task, tasks)  | 
| 23967 | 345  | 
| NONE =>  | 
346  | 
let  | 
|
| 
65443
 
dccbfc715904
provide Resources.import_name in ML, similar to Scala version;
 
wenzelm 
parents: 
65058 
diff
changeset
 | 
347  | 
val _ = member (op =) initiators theory_name andalso error (cycle_msg initiators);  | 
| 
43651
 
511df47bcadc
some support for theory files within Isabelle/Scala session;
 
wenzelm 
parents: 
43641 
diff
changeset
 | 
348  | 
|
| 65454 | 349  | 
val (current, deps, theory_pos, imports, keywords) = check_deps master_dir theory_name  | 
| 55797 | 350  | 
handle ERROR msg =>  | 
351  | 
cat_error msg  | 
|
| 
65443
 
dccbfc715904
provide Resources.import_name in ML, similar to Scala version;
 
wenzelm 
parents: 
65058 
diff
changeset
 | 
352  | 
                ("The error(s) above occurred for theory " ^ quote theory_name ^
 | 
| 55797 | 353  | 
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: 
37980 
diff
changeset
 | 
354  | 
|
| 65455 | 355  | 
val qualifier' = Resources.theory_qualifier theory_name;  | 
356  | 
val dir' = Path.append dir (master_dir_deps (Option.map #1 deps));  | 
|
357  | 
||
358  | 
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: 
48888 
diff
changeset
 | 
359  | 
val (parents_current, tasks') =  | 
| 
67297
 
86a099f896fc
formal check of @{cite} bibtex entries -- only in batch-mode session builds;
 
wenzelm 
parents: 
67194 
diff
changeset
 | 
360  | 
require_thys context (theory_name :: initiators) qualifier' dir' imports tasks;  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
361  | 
|
| 23967 | 362  | 
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: 
37955 
diff
changeset
 | 
363  | 
val task =  | 
| 
65443
 
dccbfc715904
provide Resources.import_name in ML, similar to Scala version;
 
wenzelm 
parents: 
65058 
diff
changeset
 | 
364  | 
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: 
37955 
diff
changeset
 | 
365  | 
else  | 
| 
37980
 
b8c3d4dc1e3e
avoid repeated File.read for theory text (as before);
 
wenzelm 
parents: 
37979 
diff
changeset
 | 
366  | 
(case deps of  | 
| 
 
b8c3d4dc1e3e
avoid repeated File.read for theory text (as before);
 
wenzelm 
parents: 
37979 
diff
changeset
 | 
367  | 
NONE => raise Fail "Malformed deps"  | 
| 
42003
 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 
wenzelm 
parents: 
41955 
diff
changeset
 | 
368  | 
| SOME (dep, text) =>  | 
| 
48876
 
157dd47032e0
more standard Thy_Load.check_thy for Pure.thy, relying on its header;
 
wenzelm 
parents: 
48874 
diff
changeset
 | 
369  | 
let  | 
| 
 
157dd47032e0
more standard Thy_Load.check_thy for Pure.thy, relying on its header;
 
wenzelm 
parents: 
48874 
diff
changeset
 | 
370  | 
val update_time = serial ();  | 
| 48928 | 371  | 
val load =  | 
| 
67297
 
86a099f896fc
formal check of @{cite} bibtex entries -- only in batch-mode session builds;
 
wenzelm 
parents: 
67194 
diff
changeset
 | 
372  | 
load_thy context initiators update_time  | 
| 
 
86a099f896fc
formal check of @{cite} bibtex entries -- only in batch-mode session builds;
 
wenzelm 
parents: 
67194 
diff
changeset
 | 
373  | 
dep text (theory_name, theory_pos) keywords;  | 
| 
66711
 
80fa1401cf76
discontinued extra checks (see ce676a750575 and 60c159d490a2) -- qualified theory names are meant to cover this;
 
wenzelm 
parents: 
66377 
diff
changeset
 | 
374  | 
in Task (parents, load) end);  | 
| 
48874
 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 
wenzelm 
parents: 
48710 
diff
changeset
 | 
375  | 
|
| 
65443
 
dccbfc715904
provide Resources.import_name in ML, similar to Scala version;
 
wenzelm 
parents: 
65058 
diff
changeset
 | 
376  | 
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: 
48888 
diff
changeset
 | 
377  | 
in (all_current, tasks'') end)  | 
| 7066 | 378  | 
end;  | 
| 
6484
 
3f098b0ec683
use_thy etc.: may specify path prefix, which is temporarily used as load path;
 
wenzelm 
parents: 
6389 
diff
changeset
 | 
379  | 
|
| 23967 | 380  | 
end;  | 
381  | 
||
382  | 
||
| 65444 | 383  | 
(* use theories *)  | 
| 23967 | 384  | 
|
| 
67297
 
86a099f896fc
formal check of @{cite} bibtex entries -- only in batch-mode session builds;
 
wenzelm 
parents: 
67194 
diff
changeset
 | 
385  | 
fun use_theories {document, symbols, bibtex_entries, last_timing, qualifier, master_dir} imports =
 | 
| 
65445
 
e9e7f5f5794c
more qualifier treatment, but in the end it is still ignored;
 
wenzelm 
parents: 
65444 
diff
changeset
 | 
386  | 
let  | 
| 
67297
 
86a099f896fc
formal check of @{cite} bibtex entries -- only in batch-mode session builds;
 
wenzelm 
parents: 
67194 
diff
changeset
 | 
387  | 
val context =  | 
| 
 
86a099f896fc
formal check of @{cite} bibtex entries -- only in batch-mode session builds;
 
wenzelm 
parents: 
67194 
diff
changeset
 | 
388  | 
      {document = document, symbols = symbols, bibtex_entries = bibtex_entries,
 | 
| 
 
86a099f896fc
formal check of @{cite} bibtex entries -- only in batch-mode session builds;
 
wenzelm 
parents: 
67194 
diff
changeset
 | 
389  | 
last_timing = last_timing};  | 
| 
 
86a099f896fc
formal check of @{cite} bibtex entries -- only in batch-mode session builds;
 
wenzelm 
parents: 
67194 
diff
changeset
 | 
390  | 
val (_, tasks) = require_thys context [] qualifier master_dir imports String_Graph.empty;  | 
| 66368 | 391  | 
in if Multithreading.enabled () then schedule_futures tasks else schedule_seq tasks end;  | 
| 23967 | 392  | 
|
| 65444 | 393  | 
fun use_thy name =  | 
| 61381 | 394  | 
use_theories  | 
| 
67297
 
86a099f896fc
formal check of @{cite} bibtex entries -- only in batch-mode session builds;
 
wenzelm 
parents: 
67194 
diff
changeset
 | 
395  | 
    {document = false, symbols = HTML.no_symbols, bibtex_entries = [],
 | 
| 
 
86a099f896fc
formal check of @{cite} bibtex entries -- only in batch-mode session builds;
 
wenzelm 
parents: 
67194 
diff
changeset
 | 
396  | 
last_timing = K Time.zeroTime, qualifier = Resources.default_qualifier,  | 
| 
 
86a099f896fc
formal check of @{cite} bibtex entries -- only in batch-mode session builds;
 
wenzelm 
parents: 
67194 
diff
changeset
 | 
397  | 
master_dir = Path.current}  | 
| 65444 | 398  | 
[(name, Position.none)];  | 
| 7211 | 399  | 
|
| 59364 | 400  | 
|
| 57626 | 401  | 
(* toplevel scripting -- without maintaining database *)  | 
402  | 
||
| 58856 | 403  | 
fun script_thy pos txt thy =  | 
| 57626 | 404  | 
let  | 
| 58856 | 405  | 
val trs =  | 
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58856 
diff
changeset
 | 
406  | 
Outer_Syntax.parse thy pos txt  | 
| 58856 | 407  | 
|> map (Toplevel.modify_init (K thy));  | 
| 57626 | 408  | 
val end_pos = if null trs then pos else Toplevel.pos_of (List.last trs);  | 
409  | 
val end_state = fold (Toplevel.command_exception true) trs Toplevel.toplevel;  | 
|
410  | 
in Toplevel.end_theory end_pos end_state end;  | 
|
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
411  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
412  | 
|
| 
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: 
37955 
diff
changeset
 | 
413  | 
(* register theory *)  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
414  | 
|
| 
37954
 
a2e73df0b1e0
simplified/clarified register_thy: more precise treatment of new dependencies, remove descendants;
 
wenzelm 
parents: 
37953 
diff
changeset
 | 
415  | 
fun register_thy theory =  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
416  | 
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: 
65455 
diff
changeset
 | 
417  | 
val name = Context.theory_long_name theory;  | 
| 56208 | 418  | 
    val {master, ...} = Resources.check_thy (Resources.master_directory theory) name;
 | 
419  | 
val imports = Resources.imports_of theory;  | 
|
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
420  | 
in  | 
| 59178 | 421  | 
change_thys (fn thys =>  | 
422  | 
let  | 
|
423  | 
val thys' = remove name thys;  | 
|
424  | 
        val _ = writeln ("Registering theory " ^ quote name);
 | 
|
425  | 
in update (make_deps master imports) theory thys' end)  | 
|
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
426  | 
end;  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
427  | 
|
| 
24080
 
8c67d869531b
added register_thy (replaces pretend_use_thy_only and really flag);
 
wenzelm 
parents: 
24071 
diff
changeset
 | 
428  | 
|
| 
 
8c67d869531b
added register_thy (replaces pretend_use_thy_only and really flag);
 
wenzelm 
parents: 
24071 
diff
changeset
 | 
429  | 
(* finish all theories *)  | 
| 
 
8c67d869531b
added register_thy (replaces pretend_use_thy_only and really flag);
 
wenzelm 
parents: 
24071 
diff
changeset
 | 
430  | 
|
| 
50862
 
5fc8b83322f5
more sensible order of theory nodes (correspondance to Scala version), e.g. relevant to theory progress;
 
wenzelm 
parents: 
50845 
diff
changeset
 | 
431  | 
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: 
24071 
diff
changeset
 | 
432  | 
|
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
433  | 
end;  | 
| 62847 | 434  | 
|
| 65444 | 435  | 
fun use_thy name = Runtime.toplevel_program (fn () => Thy_Info.use_thy name);  |