| author | wenzelm | 
| Wed, 26 Mar 2014 14:41:52 +0100 | |
| changeset 56294 | 85911b8a6868 | 
| parent 56286 | 7ede6ca96c61 | 
| child 56333 | 38f1422ef473 | 
| permissions | -rw-r--r-- | 
| 56208 | 1  | 
(* Title: Pure/PIDE/resources.ML  | 
| 42002 | 2  | 
Author: Makarius  | 
| 6168 | 3  | 
|
| 56208 | 4  | 
Resources for theories and auxiliary files.  | 
| 6168 | 5  | 
*)  | 
6  | 
||
| 56208 | 7  | 
signature RESOURCES =  | 
| 6168 | 8  | 
sig  | 
| 
37949
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37943 
diff
changeset
 | 
9  | 
val master_directory: theory -> Path.T  | 
| 
48927
 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 
wenzelm 
parents: 
48924 
diff
changeset
 | 
10  | 
val imports_of: theory -> (string * Position.T) list  | 
| 46737 | 11  | 
val thy_path: Path.T -> Path.T  | 
| 
48898
 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 
wenzelm 
parents: 
48897 
diff
changeset
 | 
12  | 
val check_thy: Path.T -> string ->  | 
| 48928 | 13  | 
   {master: Path.T * SHA1.digest, text: string, theory_pos: Position.T,
 | 
| 
51293
 
05b1bbae748d
discontinued obsolete 'uses' within theory header;
 
wenzelm 
parents: 
51273 
diff
changeset
 | 
14  | 
imports: (string * Position.T) list, keywords: Thy_Header.keywords}  | 
| 54525 | 15  | 
val parse_files: string -> (theory -> Token.file list) parser  | 
| 48906 | 16  | 
val provide: Path.T * SHA1.digest -> theory -> theory  | 
17  | 
val provide_parse_files: string -> (theory -> Token.file list * theory) parser  | 
|
| 
43702
 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 
wenzelm 
parents: 
43700 
diff
changeset
 | 
18  | 
val load_file: theory -> Path.T -> (Path.T * SHA1.digest) * string  | 
| 
37949
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37943 
diff
changeset
 | 
19  | 
val loaded_files: theory -> Path.T list  | 
| 54446 | 20  | 
val loaded_files_current: theory -> bool  | 
| 
46938
 
cda018294515
some support for outer syntax keyword declarations within theory header;
 
wenzelm 
parents: 
46811 
diff
changeset
 | 
21  | 
val begin_theory: Path.T -> Thy_Header.header -> theory list -> theory  | 
| 
54458
 
96ccc8972fc7
prefer explicit "document" flag -- eliminated stateful Present.no_document;
 
wenzelm 
parents: 
54456 
diff
changeset
 | 
22  | 
val load_thy: bool -> (Toplevel.transition -> Time.time option) -> int -> Path.T ->  | 
| 
 
96ccc8972fc7
prefer explicit "document" flag -- eliminated stateful Present.no_document;
 
wenzelm 
parents: 
54456 
diff
changeset
 | 
23  | 
Thy_Header.header -> Position.T -> string -> theory list -> theory * (unit -> unit) * int  | 
| 6168 | 24  | 
end;  | 
25  | 
||
| 56208 | 26  | 
structure Resources: RESOURCES =  | 
| 6168 | 27  | 
struct  | 
28  | 
||
| 
37949
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37943 
diff
changeset
 | 
29  | 
(* manage source files *)  | 
| 
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37943 
diff
changeset
 | 
30  | 
|
| 
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37943 
diff
changeset
 | 
31  | 
type files =  | 
| 
40741
 
17d6293a1e26
moved file identification to thy_load.ML (where it is actually used);
 
wenzelm 
parents: 
40625 
diff
changeset
 | 
32  | 
 {master_dir: Path.T,  (*master directory of theory source*)
 | 
| 
48927
 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 
wenzelm 
parents: 
48924 
diff
changeset
 | 
33  | 
imports: (string * Position.T) list, (*source specification of imports*)  | 
| 
48896
 
bb1f461a7815
simplified Thy_Load.provide: do not store full path;
 
wenzelm 
parents: 
48888 
diff
changeset
 | 
34  | 
provided: (Path.T * SHA1.digest) list}; (*source path, digest*)  | 
| 
37949
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37943 
diff
changeset
 | 
35  | 
|
| 
48886
 
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
 
wenzelm 
parents: 
48881 
diff
changeset
 | 
36  | 
fun make_files (master_dir, imports, provided): files =  | 
| 
 
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
 
wenzelm 
parents: 
48881 
diff
changeset
 | 
37  | 
 {master_dir = master_dir, imports = imports, provided = provided};
 | 
| 
37949
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37943 
diff
changeset
 | 
38  | 
|
| 
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37943 
diff
changeset
 | 
39  | 
structure Files = Theory_Data  | 
| 
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37943 
diff
changeset
 | 
40  | 
(  | 
| 
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37943 
diff
changeset
 | 
41  | 
type T = files;  | 
| 
48886
 
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
 
wenzelm 
parents: 
48881 
diff
changeset
 | 
42  | 
val empty = make_files (Path.current, [], []);  | 
| 
37949
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37943 
diff
changeset
 | 
43  | 
fun extend _ = empty;  | 
| 
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37943 
diff
changeset
 | 
44  | 
fun merge _ = empty;  | 
| 
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37943 
diff
changeset
 | 
45  | 
);  | 
| 
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37943 
diff
changeset
 | 
46  | 
|
| 
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37943 
diff
changeset
 | 
47  | 
fun map_files f =  | 
| 
48886
 
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
 
wenzelm 
parents: 
48881 
diff
changeset
 | 
48  | 
  Files.map (fn {master_dir, imports, provided} =>
 | 
| 
 
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
 
wenzelm 
parents: 
48881 
diff
changeset
 | 
49  | 
make_files (f (master_dir, imports, provided)));  | 
| 
37949
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37943 
diff
changeset
 | 
50  | 
|
| 
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37943 
diff
changeset
 | 
51  | 
|
| 
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37943 
diff
changeset
 | 
52  | 
val master_directory = #master_dir o Files.get;  | 
| 
41548
 
bd0bebf93fa6
Thy_Load.begin_theory: maintain source specification of imports;
 
wenzelm 
parents: 
41414 
diff
changeset
 | 
53  | 
val imports_of = #imports o Files.get;  | 
| 
37949
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37943 
diff
changeset
 | 
54  | 
|
| 48888 | 55  | 
fun put_deps master_dir imports = map_files (fn _ => (master_dir, imports, []));  | 
| 
37949
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37943 
diff
changeset
 | 
56  | 
|
| 
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37943 
diff
changeset
 | 
57  | 
|
| 54525 | 58  | 
(* theory files *)  | 
59  | 
||
60  | 
val thy_path = Path.ext "thy";  | 
|
| 48869 | 61  | 
|
| 48878 | 62  | 
fun check_file dir file = File.check_file (File.full_path dir file);  | 
63  | 
||
| 54525 | 64  | 
fun check_thy dir thy_name =  | 
65  | 
let  | 
|
66  | 
val path = thy_path (Path.basic thy_name);  | 
|
67  | 
val master_file = check_file dir path;  | 
|
68  | 
val text = File.read master_file;  | 
|
69  | 
||
70  | 
    val {name = (name, pos), imports, keywords} =
 | 
|
71  | 
Thy_Header.read (Path.position master_file) text;  | 
|
72  | 
val _ = thy_name <> name andalso  | 
|
73  | 
      error ("Bad file name " ^ Path.print path ^ " for theory " ^ quote name ^ Position.here pos);
 | 
|
74  | 
in  | 
|
75  | 
   {master = (master_file, SHA1.digest text), text = text, theory_pos = pos,
 | 
|
76  | 
imports = imports, keywords = keywords}  | 
|
77  | 
end;  | 
|
78  | 
||
79  | 
||
| 54526 | 80  | 
(* load files *)  | 
| 
48898
 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 
wenzelm 
parents: 
48897 
diff
changeset
 | 
81  | 
|
| 
 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 
wenzelm 
parents: 
48897 
diff
changeset
 | 
82  | 
fun parse_files cmd =  | 
| 
 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 
wenzelm 
parents: 
48897 
diff
changeset
 | 
83  | 
Scan.ahead Parse.not_eof -- Parse.path >> (fn (tok, name) => fn thy =>  | 
| 
 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 
wenzelm 
parents: 
48897 
diff
changeset
 | 
84  | 
(case Token.get_files tok of  | 
| 54526 | 85  | 
[] =>  | 
86  | 
let  | 
|
87  | 
val master_dir = master_directory thy;  | 
|
| 55708 | 88  | 
val pos = Token.pos_of tok;  | 
| 54526 | 89  | 
val src_paths = Keyword.command_files cmd (Path.explode name);  | 
90  | 
in map (Command.read_file master_dir pos) src_paths end  | 
|
| 
54520
 
cee77d2e9582
release file errors at runtime: Command.eval instead of Command.read;
 
wenzelm 
parents: 
54519 
diff
changeset
 | 
91  | 
| files => map Exn.release files));  | 
| 48869 | 92  | 
|
| 48906 | 93  | 
fun provide (src_path, id) =  | 
94  | 
map_files (fn (master_dir, imports, provided) =>  | 
|
95  | 
if AList.defined (op =) provided src_path then  | 
|
96  | 
      error ("Duplicate use of source file: " ^ Path.print src_path)
 | 
|
97  | 
else (master_dir, imports, (src_path, id) :: provided));  | 
|
98  | 
||
99  | 
fun provide_parse_files cmd =  | 
|
100  | 
parse_files cmd >> (fn files => fn thy =>  | 
|
101  | 
let  | 
|
102  | 
val fs = files thy;  | 
|
| 
55788
 
67699e08e969
store blobs / inlined files as separate text lines: smaller values are more healthy for the Poly/ML RTS and allow implicit sharing;
 
wenzelm 
parents: 
55708 
diff
changeset
 | 
103  | 
      val thy' = fold (fn {src_path, digest, ...} => provide (src_path, digest)) fs thy;
 | 
| 48906 | 104  | 
in (fs, thy') end);  | 
105  | 
||
| 
43702
 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 
wenzelm 
parents: 
43700 
diff
changeset
 | 
106  | 
fun load_file thy src_path =  | 
| 
 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 
wenzelm 
parents: 
43700 
diff
changeset
 | 
107  | 
let  | 
| 
 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 
wenzelm 
parents: 
43700 
diff
changeset
 | 
108  | 
val full_path = check_file (master_directory thy) src_path;  | 
| 
 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 
wenzelm 
parents: 
43700 
diff
changeset
 | 
109  | 
val text = File.read full_path;  | 
| 
 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 
wenzelm 
parents: 
43700 
diff
changeset
 | 
110  | 
val id = SHA1.digest text;  | 
| 56286 | 111  | 
    val _ = legacy_feature ("Raw file-system access to load file " ^ Path.print full_path);
 | 
| 
43702
 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 
wenzelm 
parents: 
43700 
diff
changeset
 | 
112  | 
in ((full_path, id), text) end;  | 
| 
 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 
wenzelm 
parents: 
43700 
diff
changeset
 | 
113  | 
|
| 54446 | 114  | 
fun loaded_files_current thy =  | 
| 
48886
 
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
 
wenzelm 
parents: 
48881 
diff
changeset
 | 
115  | 
#provided (Files.get thy) |>  | 
| 
48896
 
bb1f461a7815
simplified Thy_Load.provide: do not store full path;
 
wenzelm 
parents: 
48888 
diff
changeset
 | 
116  | 
forall (fn (src_path, id) =>  | 
| 
43702
 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 
wenzelm 
parents: 
43700 
diff
changeset
 | 
117  | 
(case try (load_file thy) src_path of  | 
| 
37949
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37943 
diff
changeset
 | 
118  | 
NONE => false  | 
| 
48886
 
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
 
wenzelm 
parents: 
48881 
diff
changeset
 | 
119  | 
| SOME ((_, id'), _) => id = id'));  | 
| 37952 | 120  | 
|
| 54454 | 121  | 
(*Proof General legacy*)  | 
122  | 
fun loaded_files thy =  | 
|
123  | 
  let val {master_dir, provided, ...} = Files.get thy
 | 
|
124  | 
in map (File.full_path master_dir o #1) provided end;  | 
|
125  | 
||
| 
37949
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37943 
diff
changeset
 | 
126  | 
|
| 54525 | 127  | 
(* load theory *)  | 
| 
37977
 
3ceccd415145
simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
 
wenzelm 
parents: 
37952 
diff
changeset
 | 
128  | 
|
| 
51294
 
0850d43cb355
discontinued obsolete header "files" -- these are loaded explicitly after exploring dependencies;
 
wenzelm 
parents: 
51293 
diff
changeset
 | 
129  | 
fun begin_theory master_dir {name, imports, 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: 
37952 
diff
changeset
 | 
130  | 
Theory.begin_theory name parents  | 
| 48888 | 131  | 
|> put_deps master_dir imports  | 
| 52788 | 132  | 
|> fold Thy_Header.declare_keyword keywords;  | 
| 
37977
 
3ceccd415145
simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
 
wenzelm 
parents: 
37952 
diff
changeset
 | 
133  | 
|
| 54526 | 134  | 
fun excursion master_dir last_timing init elements =  | 
| 
46959
 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
135  | 
let  | 
| 
51268
 
fcc4b89a600d
simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
 
wenzelm 
parents: 
51228 
diff
changeset
 | 
136  | 
fun prepare_span span =  | 
| 
 
fcc4b89a600d
simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
 
wenzelm 
parents: 
51228 
diff
changeset
 | 
137  | 
Thy_Syntax.span_content span  | 
| 54526 | 138  | 
|> Command.read init master_dir []  | 
| 
51268
 
fcc4b89a600d
simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
 
wenzelm 
parents: 
51228 
diff
changeset
 | 
139  | 
|> (fn tr => Toplevel.put_timing (last_timing tr) tr);  | 
| 
51217
 
65ab2c4f4c32
support for prescient timing information within command transactions;
 
wenzelm 
parents: 
50201 
diff
changeset
 | 
140  | 
|
| 
51268
 
fcc4b89a600d
simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
 
wenzelm 
parents: 
51228 
diff
changeset
 | 
141  | 
fun element_result span_elem (st, _) =  | 
| 
46959
 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
142  | 
let  | 
| 
51273
 
d54ee0cad2ab
clarified Toplevel.element_result: scheduling policies happen here;
 
wenzelm 
parents: 
51268 
diff
changeset
 | 
143  | 
val elem = Thy_Syntax.map_element prepare_span span_elem;  | 
| 
 
d54ee0cad2ab
clarified Toplevel.element_result: scheduling policies happen here;
 
wenzelm 
parents: 
51268 
diff
changeset
 | 
144  | 
val (results, st') = Toplevel.element_result elem st;  | 
| 
 
d54ee0cad2ab
clarified Toplevel.element_result: scheduling policies happen here;
 
wenzelm 
parents: 
51268 
diff
changeset
 | 
145  | 
val pos' = Toplevel.pos_of (Thy_Syntax.last_element elem);  | 
| 
51268
 
fcc4b89a600d
simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
 
wenzelm 
parents: 
51228 
diff
changeset
 | 
146  | 
in (results, (st', pos')) end;  | 
| 
46959
 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
147  | 
|
| 
 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
148  | 
val (results, (end_state, end_pos)) =  | 
| 
 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
149  | 
fold_map element_result elements (Toplevel.toplevel, Position.none);  | 
| 
 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
150  | 
|
| 
 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
151  | 
val thy = Toplevel.end_theory end_pos end_state;  | 
| 
51273
 
d54ee0cad2ab
clarified Toplevel.element_result: scheduling policies happen here;
 
wenzelm 
parents: 
51268 
diff
changeset
 | 
152  | 
in (results, thy) end;  | 
| 
46959
 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
153  | 
|
| 
54458
 
96ccc8972fc7
prefer explicit "document" flag -- eliminated stateful Present.no_document;
 
wenzelm 
parents: 
54456 
diff
changeset
 | 
154  | 
fun load_thy document last_timing update_time master_dir header text_pos text parents =  | 
| 
43712
 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 
wenzelm 
parents: 
43702 
diff
changeset
 | 
155  | 
let  | 
| 
 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 
wenzelm 
parents: 
43702 
diff
changeset
 | 
156  | 
val time = ! Toplevel.timing;  | 
| 
 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 
wenzelm 
parents: 
43702 
diff
changeset
 | 
157  | 
|
| 
51294
 
0850d43cb355
discontinued obsolete header "files" -- these are loaded explicitly after exploring dependencies;
 
wenzelm 
parents: 
51293 
diff
changeset
 | 
158  | 
    val {name = (name, _), ...} = header;
 | 
| 
46958
 
0ec8f04e753a
define keywords early when processing the theory header, before running the body commands;
 
wenzelm 
parents: 
46938 
diff
changeset
 | 
159  | 
val _ = Thy_Header.define_keywords header;  | 
| 
43712
 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 
wenzelm 
parents: 
43702 
diff
changeset
 | 
160  | 
|
| 
46959
 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
161  | 
val lexs = Keyword.get_lexicons ();  | 
| 
48927
 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 
wenzelm 
parents: 
48924 
diff
changeset
 | 
162  | 
val toks = Thy_Syntax.parse_tokens lexs text_pos text;  | 
| 54526 | 163  | 
val spans = Thy_Syntax.parse_spans toks;  | 
| 
46959
 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
164  | 
val elements = Thy_Syntax.parse_elements spans;  | 
| 
43712
 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 
wenzelm 
parents: 
43702 
diff
changeset
 | 
165  | 
|
| 54456 | 166  | 
fun init () =  | 
167  | 
begin_theory master_dir header parents  | 
|
168  | 
|> Present.begin_theory update_time  | 
|
169  | 
(fn () => HTML.html_mode (implode o map Thy_Syntax.present_span) spans);  | 
|
| 
43712
 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 
wenzelm 
parents: 
43702 
diff
changeset
 | 
170  | 
|
| 
 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 
wenzelm 
parents: 
43702 
diff
changeset
 | 
171  | 
    val _ = if time then writeln ("\n**** Starting theory " ^ quote name ^ " ****") else ();
 | 
| 54526 | 172  | 
val (results, thy) =  | 
173  | 
cond_timeit time "" (fn () => excursion master_dir last_timing init elements);  | 
|
| 
43712
 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 
wenzelm 
parents: 
43702 
diff
changeset
 | 
174  | 
    val _ = if time then writeln ("**** Finished theory " ^ quote name ^ " ****\n") else ();
 | 
| 
 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 
wenzelm 
parents: 
43702 
diff
changeset
 | 
175  | 
|
| 
51331
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51326 
diff
changeset
 | 
176  | 
fun present () =  | 
| 
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51326 
diff
changeset
 | 
177  | 
let  | 
| 
51332
 
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
 
wenzelm 
parents: 
51331 
diff
changeset
 | 
178  | 
val res = filter_out (Toplevel.is_ignored o #1) (maps Toplevel.join_results results);  | 
| 
51331
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51326 
diff
changeset
 | 
179  | 
val ((minor, _), outer_syntax) = Outer_Syntax.get_syntax ();  | 
| 
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51326 
diff
changeset
 | 
180  | 
in  | 
| 51556 | 181  | 
if exists (Toplevel.is_skipped_proof o #2) res then  | 
182  | 
          warning ("Cannot present theory with skipped proofs: " ^ quote name)
 | 
|
183  | 
else  | 
|
| 
54458
 
96ccc8972fc7
prefer explicit "document" flag -- eliminated stateful Present.no_document;
 
wenzelm 
parents: 
54456 
diff
changeset
 | 
184  | 
let val tex_source =  | 
| 
 
96ccc8972fc7
prefer explicit "document" flag -- eliminated stateful Present.no_document;
 
wenzelm 
parents: 
54456 
diff
changeset
 | 
185  | 
Thy_Output.present_thy minor Keyword.command_tags  | 
| 
 
96ccc8972fc7
prefer explicit "document" flag -- eliminated stateful Present.no_document;
 
wenzelm 
parents: 
54456 
diff
changeset
 | 
186  | 
(Outer_Syntax.is_markup outer_syntax) res toks  | 
| 
 
96ccc8972fc7
prefer explicit "document" flag -- eliminated stateful Present.no_document;
 
wenzelm 
parents: 
54456 
diff
changeset
 | 
187  | 
|> Buffer.content;  | 
| 
 
96ccc8972fc7
prefer explicit "document" flag -- eliminated stateful Present.no_document;
 
wenzelm 
parents: 
54456 
diff
changeset
 | 
188  | 
in if document then Present.theory_output name tex_source else () end  | 
| 
51331
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51326 
diff
changeset
 | 
189  | 
end;  | 
| 
43712
 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 
wenzelm 
parents: 
43702 
diff
changeset
 | 
190  | 
|
| 
51331
 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 
wenzelm 
parents: 
51326 
diff
changeset
 | 
191  | 
in (thy, present, size text) end;  | 
| 
43712
 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 
wenzelm 
parents: 
43702 
diff
changeset
 | 
192  | 
|
| 42002 | 193  | 
|
| 56135 | 194  | 
(* antiquotations *)  | 
| 
49244
 
fb669aff821e
formal markup for @{file} (for hyperlinks etc.) -- interpret path wrt. master directory as usual;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
195  | 
|
| 56135 | 196  | 
local  | 
197  | 
||
198  | 
fun check_path strict ctxt dir (name, pos) =  | 
|
| 54705 | 199  | 
let  | 
| 
56034
 
1c59b555ac4a
some Markup.language_path to prevent completion of symbols (notably "~") -- always "delimited" for simplicity in contrast to 42ac3cfb89f6;
 
wenzelm 
parents: 
55882 
diff
changeset
 | 
200  | 
val _ = Context_Position.report ctxt pos Markup.language_path;  | 
| 
 
1c59b555ac4a
some Markup.language_path to prevent completion of symbols (notably "~") -- always "delimited" for simplicity in contrast to 42ac3cfb89f6;
 
wenzelm 
parents: 
55882 
diff
changeset
 | 
201  | 
|
| 55882 | 202  | 
val path = Path.append dir (Path.explode name)  | 
203  | 
handle ERROR msg => error (msg ^ Position.here pos);  | 
|
204  | 
||
| 56134 | 205  | 
val _ = Context_Position.report ctxt pos (Markup.path (Path.smart_implode path));  | 
| 54705 | 206  | 
val _ =  | 
| 55882 | 207  | 
if can Path.expand path andalso File.exists path then ()  | 
208  | 
else  | 
|
209  | 
let  | 
|
210  | 
val path' = perhaps (try Path.expand) path;  | 
|
211  | 
val msg = "Bad file: " ^ Path.print path' ^ Position.here pos;  | 
|
212  | 
in  | 
|
213  | 
if strict then error msg  | 
|
| 
56294
 
85911b8a6868
prefer Context_Position where a context is available;
 
wenzelm 
parents: 
56286 
diff
changeset
 | 
214  | 
else if Context_Position.is_visible ctxt then  | 
| 
 
85911b8a6868
prefer Context_Position where a context is available;
 
wenzelm 
parents: 
56286 
diff
changeset
 | 
215  | 
Output.report  | 
| 55882 | 216  | 
(Markup.markup (Markup.bad |> Markup.properties (Position.properties_of pos)) msg)  | 
| 
56294
 
85911b8a6868
prefer Context_Position where a context is available;
 
wenzelm 
parents: 
56286 
diff
changeset
 | 
217  | 
else ()  | 
| 55882 | 218  | 
end;  | 
| 56135 | 219  | 
in path end;  | 
220  | 
||
221  | 
fun file_antiq strict ctxt (name, pos) =  | 
|
222  | 
let  | 
|
223  | 
val dir = master_directory (Proof_Context.theory_of ctxt);  | 
|
224  | 
val _ = check_path strict ctxt dir (name, pos);  | 
|
| 54705 | 225  | 
in  | 
226  | 
space_explode "/" name  | 
|
227  | 
|> map Thy_Output.verb_text  | 
|
228  | 
    |> space_implode (Thy_Output.verb_text "/" ^ "\\discretionary{}{}{}")
 | 
|
229  | 
end;  | 
|
230  | 
||
| 56135 | 231  | 
in  | 
232  | 
||
| 53171 | 233  | 
val _ = Theory.setup  | 
| 56204 | 234  | 
 (Thy_Output.antiquotation @{binding file} (Scan.lift (Parse.position Parse.path))
 | 
| 54705 | 235  | 
(file_antiq true o #context) #>  | 
| 56204 | 236  | 
  Thy_Output.antiquotation @{binding file_unchecked} (Scan.lift (Parse.position Parse.path))
 | 
| 56135 | 237  | 
(file_antiq false o #context) #>  | 
| 56204 | 238  | 
  ML_Antiquotation.value @{binding path}
 | 
| 56135 | 239  | 
(Args.context -- Scan.lift (Parse.position Parse.path) >> (fn (ctxt, arg) =>  | 
240  | 
let val path = check_path true ctxt Path.current arg  | 
|
241  | 
in "Path.explode " ^ ML_Syntax.print_string (Path.implode path) end)));  | 
|
| 
49244
 
fb669aff821e
formal markup for @{file} (for hyperlinks etc.) -- interpret path wrt. master directory as usual;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
242  | 
|
| 6168 | 243  | 
end;  | 
| 56135 | 244  | 
|
245  | 
end;  |