| author | wenzelm | 
| Mon, 21 Mar 2016 20:22:07 +0100 | |
| changeset 62681 | 45b8dd2d3827 | 
| parent 61381 | ddca85598c65 | 
| child 63669 | 256fc20716f2 | 
| 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: 
37943diff
changeset | 9 | val master_directory: theory -> Path.T | 
| 48927 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 wenzelm parents: 
48924diff
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: 
48897diff
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: 
51273diff
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 | |
| 54446 | 18 | val loaded_files_current: theory -> bool | 
| 46938 
cda018294515
some support for outer syntax keyword declarations within theory header;
 wenzelm parents: 
46811diff
changeset | 19 | val begin_theory: Path.T -> Thy_Header.header -> theory list -> theory | 
| 61381 | 20 | val load_thy: bool -> HTML.symbols -> (Toplevel.transition -> Time.time option) -> int -> | 
| 21 | Path.T -> Thy_Header.header -> Position.T -> string -> theory list -> | |
| 22 | theory * (unit -> unit) * int | |
| 6168 | 23 | end; | 
| 24 | ||
| 56208 | 25 | structure Resources: RESOURCES = | 
| 6168 | 26 | struct | 
| 27 | ||
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 28 | (* manage source files *) | 
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 29 | |
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 30 | type files = | 
| 40741 
17d6293a1e26
moved file identification to thy_load.ML (where it is actually used);
 wenzelm parents: 
40625diff
changeset | 31 |  {master_dir: Path.T,  (*master directory of theory source*)
 | 
| 48927 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 wenzelm parents: 
48924diff
changeset | 32 | imports: (string * Position.T) list, (*source specification of imports*) | 
| 48896 
bb1f461a7815
simplified Thy_Load.provide: do not store full path;
 wenzelm parents: 
48888diff
changeset | 33 | 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: 
37943diff
changeset | 34 | |
| 48886 
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
 wenzelm parents: 
48881diff
changeset | 35 | 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: 
48881diff
changeset | 36 |  {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: 
37943diff
changeset | 37 | |
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 38 | 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: 
37943diff
changeset | 39 | ( | 
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 40 | type T = files; | 
| 48886 
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
 wenzelm parents: 
48881diff
changeset | 41 | 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: 
37943diff
changeset | 42 | fun extend _ = empty; | 
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 43 | fun merge _ = empty; | 
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 44 | ); | 
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 45 | |
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 46 | fun map_files f = | 
| 48886 
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
 wenzelm parents: 
48881diff
changeset | 47 |   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: 
48881diff
changeset | 48 | 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: 
37943diff
changeset | 49 | |
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 50 | |
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 51 | val master_directory = #master_dir o Files.get; | 
| 41548 
bd0bebf93fa6
Thy_Load.begin_theory: maintain source specification of imports;
 wenzelm parents: 
41414diff
changeset | 52 | 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: 
37943diff
changeset | 53 | |
| 48888 | 54 | 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: 
37943diff
changeset | 55 | |
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 56 | |
| 54525 | 57 | (* theory files *) | 
| 58 | ||
| 59 | val thy_path = Path.ext "thy"; | |
| 48869 | 60 | |
| 48878 | 61 | fun check_file dir file = File.check_file (File.full_path dir file); | 
| 62 | ||
| 54525 | 63 | fun check_thy dir thy_name = | 
| 64 | let | |
| 65 | val path = thy_path (Path.basic thy_name); | |
| 66 | val master_file = check_file dir path; | |
| 67 | val text = File.read master_file; | |
| 68 | ||
| 69 |     val {name = (name, pos), imports, keywords} =
 | |
| 70 | Thy_Header.read (Path.position master_file) text; | |
| 71 | val _ = thy_name <> name andalso | |
| 59718 | 72 |       error ("Bad theory name " ^ quote name ^
 | 
| 73 | " for file " ^ Path.print path ^ Position.here pos); | |
| 54525 | 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: 
48897diff
changeset | 81 | |
| 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 wenzelm parents: 
48897diff
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: 
48897diff
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: 
48897diff
changeset | 84 | (case Token.get_files tok of | 
| 54526 | 85 | [] => | 
| 86 | let | |
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58923diff
changeset | 87 | val keywords = Thy_Header.get_keywords thy; | 
| 54526 | 88 | val master_dir = master_directory thy; | 
| 55708 | 89 | val pos = Token.pos_of tok; | 
| 58923 | 90 | val src_paths = Keyword.command_files keywords cmd (Path.explode name); | 
| 54526 | 91 | 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: 
54519diff
changeset | 92 | | files => map Exn.release files)); | 
| 48869 | 93 | |
| 48906 | 94 | fun provide (src_path, id) = | 
| 95 | map_files (fn (master_dir, imports, provided) => | |
| 96 | if AList.defined (op =) provided src_path then | |
| 97 |       error ("Duplicate use of source file: " ^ Path.print src_path)
 | |
| 98 | else (master_dir, imports, (src_path, id) :: provided)); | |
| 99 | ||
| 100 | fun provide_parse_files cmd = | |
| 101 | parse_files cmd >> (fn files => fn thy => | |
| 102 | let | |
| 103 | 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: 
55708diff
changeset | 104 |       val thy' = fold (fn {src_path, digest, ...} => provide (src_path, digest)) fs thy;
 | 
| 48906 | 105 | in (fs, thy') end); | 
| 106 | ||
| 43702 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43700diff
changeset | 107 | fun load_file thy src_path = | 
| 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43700diff
changeset | 108 | let | 
| 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43700diff
changeset | 109 | 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: 
43700diff
changeset | 110 | val text = File.read full_path; | 
| 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43700diff
changeset | 111 | val id = SHA1.digest text; | 
| 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43700diff
changeset | 112 | in ((full_path, id), text) end; | 
| 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43700diff
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: 
48881diff
changeset | 115 | #provided (Files.get thy) |> | 
| 48896 
bb1f461a7815
simplified Thy_Load.provide: do not store full path;
 wenzelm parents: 
48888diff
changeset | 116 | forall (fn (src_path, id) => | 
| 43702 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43700diff
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: 
37943diff
changeset | 118 | NONE => false | 
| 48886 
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
 wenzelm parents: 
48881diff
changeset | 119 | | SOME ((_, id'), _) => id = id')); | 
| 37952 | 120 | |
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 121 | |
| 54525 | 122 | (* 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: 
37952diff
changeset | 123 | |
| 51294 
0850d43cb355
discontinued obsolete header "files" -- these are loaded explicitly after exploring dependencies;
 wenzelm parents: 
51293diff
changeset | 124 | 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: 
37952diff
changeset | 125 | Theory.begin_theory name parents | 
| 48888 | 126 | |> put_deps master_dir imports | 
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58923diff
changeset | 127 | |> Thy_Header.add_keywords 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: 
37952diff
changeset | 128 | |
| 58923 | 129 | fun excursion keywords 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: 
46958diff
changeset | 130 | let | 
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58923diff
changeset | 131 | fun prepare_span st span = | 
| 57905 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
56803diff
changeset | 132 | Command_Span.content span | 
| 59689 
7968c57ea240
simplified Command.resolve_files in ML, using blobs_index from Scala;
 wenzelm parents: 
59123diff
changeset | 133 | |> Command.read keywords (Command.read_thy st) master_dir init ([], ~1) | 
| 51268 
fcc4b89a600d
simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
 wenzelm parents: 
51228diff
changeset | 134 | |> (fn tr => Toplevel.put_timing (last_timing tr) tr); | 
| 51217 
65ab2c4f4c32
support for prescient timing information within command transactions;
 wenzelm parents: 
50201diff
changeset | 135 | |
| 51268 
fcc4b89a600d
simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
 wenzelm parents: 
51228diff
changeset | 136 | 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: 
46958diff
changeset | 137 | let | 
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58923diff
changeset | 138 | val elem = Thy_Syntax.map_element (prepare_span st) span_elem; | 
| 58923 | 139 | val (results, st') = Toplevel.element_result keywords elem st; | 
| 51273 
d54ee0cad2ab
clarified Toplevel.element_result: scheduling policies happen here;
 wenzelm parents: 
51268diff
changeset | 140 | val pos' = Toplevel.pos_of (Thy_Syntax.last_element elem); | 
| 51268 
fcc4b89a600d
simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
 wenzelm parents: 
51228diff
changeset | 141 | 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: 
46958diff
changeset | 142 | |
| 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 143 | 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: 
46958diff
changeset | 144 | 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: 
46958diff
changeset | 145 | |
| 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 146 | val thy = Toplevel.end_theory end_pos end_state; | 
| 51273 
d54ee0cad2ab
clarified Toplevel.element_result: scheduling policies happen here;
 wenzelm parents: 
51268diff
changeset | 147 | 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: 
46958diff
changeset | 148 | |
| 61381 | 149 | fun load_thy document symbols 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: 
43702diff
changeset | 150 | let | 
| 58918 | 151 | val (name, _) = #name header; | 
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58923diff
changeset | 152 | val keywords = | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58923diff
changeset | 153 | fold (curry Keyword.merge_keywords o Thy_Header.get_keywords) parents | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58923diff
changeset | 154 | (Keyword.add_keywords (#keywords header) Keyword.empty_keywords); | 
| 43712 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43702diff
changeset | 155 | |
| 59083 | 156 | val toks = Token.explode keywords text_pos text; | 
| 57905 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
56803diff
changeset | 157 | val spans = Outer_Syntax.parse_spans toks; | 
| 58923 | 158 | val elements = Thy_Syntax.parse_elements keywords spans; | 
| 43712 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43702diff
changeset | 159 | |
| 54456 | 160 | fun init () = | 
| 161 | begin_theory master_dir header parents | |
| 162 | |> Present.begin_theory update_time | |
| 61381 | 163 | (fn () => implode (map (HTML.present_span symbols keywords) spans)); | 
| 43712 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43702diff
changeset | 164 | |
| 54526 | 165 | val (results, thy) = | 
| 58849 
ef7700ecce83
discontinued pointless option: timing is always on (overall theory only);
 wenzelm parents: 
58741diff
changeset | 166 |       cond_timeit true ("theory " ^ quote name)
 | 
| 58923 | 167 | (fn () => excursion keywords master_dir last_timing init elements); | 
| 43712 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43702diff
changeset | 168 | |
| 51331 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51326diff
changeset | 169 | fun present () = | 
| 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51326diff
changeset | 170 | let | 
| 51332 
8707df0b0255
refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
 wenzelm parents: 
51331diff
changeset | 171 | 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: 
51326diff
changeset | 172 | in | 
| 51556 | 173 | if exists (Toplevel.is_skipped_proof o #2) res then | 
| 174 |           warning ("Cannot present theory with skipped proofs: " ^ quote name)
 | |
| 175 | else | |
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58923diff
changeset | 176 | let val tex_source = Thy_Output.present_thy thy res toks |> Buffer.content; | 
| 54458 
96ccc8972fc7
prefer explicit "document" flag -- eliminated stateful Present.no_document;
 wenzelm parents: 
54456diff
changeset | 177 | 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: 
51326diff
changeset | 178 | end; | 
| 43712 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43702diff
changeset | 179 | |
| 51331 
e7fab0b5dbe7
join all proofs before scheduling present phase (ordered according to weight);
 wenzelm parents: 
51326diff
changeset | 180 | in (thy, present, size text) end; | 
| 43712 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43702diff
changeset | 181 | |
| 42002 | 182 | |
| 56135 | 183 | (* antiquotations *) | 
| 49244 
fb669aff821e
formal markup for @{file} (for hyperlinks etc.) -- interpret path wrt. master directory as usual;
 wenzelm parents: 
48992diff
changeset | 184 | |
| 56135 | 185 | local | 
| 186 | ||
| 187 | fun check_path strict ctxt dir (name, pos) = | |
| 54705 | 188 | let | 
| 56034 
1c59b555ac4a
some Markup.language_path to prevent completion of symbols (notably "~") -- always "delimited" for simplicity in contrast to 42ac3cfb89f6;
 wenzelm parents: 
55882diff
changeset | 189 | 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: 
55882diff
changeset | 190 | |
| 55882 | 191 | val path = Path.append dir (Path.explode name) | 
| 192 | handle ERROR msg => error (msg ^ Position.here pos); | |
| 193 | ||
| 56134 | 194 | val _ = Context_Position.report ctxt pos (Markup.path (Path.smart_implode path)); | 
| 54705 | 195 | val _ = | 
| 55882 | 196 | if can Path.expand path andalso File.exists path then () | 
| 197 | else | |
| 198 | let | |
| 199 | val path' = perhaps (try Path.expand) path; | |
| 200 | val msg = "Bad file: " ^ Path.print path' ^ Position.here pos; | |
| 201 | in | |
| 202 | if strict then error msg | |
| 56294 
85911b8a6868
prefer Context_Position where a context is available;
 wenzelm parents: 
56286diff
changeset | 203 | else if Context_Position.is_visible ctxt then | 
| 
85911b8a6868
prefer Context_Position where a context is available;
 wenzelm parents: 
56286diff
changeset | 204 | Output.report | 
| 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: 
56294diff
changeset | 205 | [Markup.markup (Markup.bad |> Markup.properties (Position.properties_of pos)) msg] | 
| 56294 
85911b8a6868
prefer Context_Position where a context is available;
 wenzelm parents: 
56286diff
changeset | 206 | else () | 
| 55882 | 207 | end; | 
| 56135 | 208 | in path end; | 
| 209 | ||
| 210 | fun file_antiq strict ctxt (name, pos) = | |
| 211 | let | |
| 212 | val dir = master_directory (Proof_Context.theory_of ctxt); | |
| 213 | val _ = check_path strict ctxt dir (name, pos); | |
| 54705 | 214 | in | 
| 215 | space_explode "/" name | |
| 58716 
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
 wenzelm parents: 
57918diff
changeset | 216 | |> map Latex.output_ascii | 
| 
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
 wenzelm parents: 
57918diff
changeset | 217 |     |> space_implode (Latex.output_ascii "/" ^ "\\discretionary{}{}{}")
 | 
| 58741 
6e7b009e6d94
clarified verbatim line breaks, e.g. relevant for Implementation mldecls;
 wenzelm parents: 
58716diff
changeset | 218 |     |> enclose "\\isatt{" "}"
 | 
| 54705 | 219 | end; | 
| 220 | ||
| 56135 | 221 | in | 
| 222 | ||
| 53171 | 223 | val _ = Theory.setup | 
| 56204 | 224 |  (Thy_Output.antiquotation @{binding file} (Scan.lift (Parse.position Parse.path))
 | 
| 54705 | 225 | (file_antiq true o #context) #> | 
| 56204 | 226 |   Thy_Output.antiquotation @{binding file_unchecked} (Scan.lift (Parse.position Parse.path))
 | 
| 56135 | 227 | (file_antiq false o #context) #> | 
| 56204 | 228 |   ML_Antiquotation.value @{binding path}
 | 
| 56135 | 229 | (Args.context -- Scan.lift (Parse.position Parse.path) >> (fn (ctxt, arg) => | 
| 230 | let val path = check_path true ctxt Path.current arg | |
| 231 | 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: 
48992diff
changeset | 232 | |
| 6168 | 233 | end; | 
| 56135 | 234 | |
| 235 | end; |