| author | wenzelm | 
| Fri, 06 Oct 2017 21:23:21 +0200 | |
| changeset 66772 | a66f11a0b5b1 | 
| parent 66680 | 74a1b722507e | 
| child 66945 | b6f787a17fbe | 
| permissions | -rw-r--r-- | 
| 6203 | 1  | 
(* Title: Pure/Thy/present.ML  | 
| 9416 | 2  | 
Author: Markus Wenzel and Stefan Berghofer, TU Muenchen  | 
| 6203 | 3  | 
|
| 
59448
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
4  | 
Theory presentation: HTML and PDF-LaTeX documents.  | 
| 6203 | 5  | 
*)  | 
6  | 
||
7  | 
signature PRESENT =  | 
|
8  | 
sig  | 
|
| 66037 | 9  | 
val theory_qualifier: theory -> string  | 
| 56614 | 10  | 
val document_enabled: string -> bool  | 
| 56612 | 11  | 
val document_variants: string -> (string * string) list  | 
| 61381 | 12  | 
val init: HTML.symbols -> bool -> bool -> Path.T -> string -> string -> (string * string) list ->  | 
| 
59448
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
13  | 
(Path.T * Path.T) list -> Path.T -> string * string -> bool -> unit  | 
| 59445 | 14  | 
val finish: unit -> unit  | 
| 65505 | 15  | 
val theory_output: theory -> string -> unit  | 
| 54456 | 16  | 
val begin_theory: int -> (unit -> HTML.text) -> theory -> theory  | 
| 52549 | 17  | 
val display_drafts: Path.T list -> int  | 
| 6203 | 18  | 
end;  | 
19  | 
||
| 7685 | 20  | 
structure Present: PRESENT =  | 
21  | 
struct  | 
|
22  | 
||
| 7727 | 23  | 
|
24  | 
(** paths **)  | 
|
25  | 
||
26  | 
val tex_ext = Path.ext "tex";  | 
|
27  | 
val tex_path = tex_ext o Path.basic;  | 
|
28  | 
val html_ext = Path.ext "html";  | 
|
29  | 
val html_path = html_ext o Path.basic;  | 
|
30  | 
val index_path = Path.basic "index.html";  | 
|
| 11856 | 31  | 
val readme_html_path = Path.basic "README.html";  | 
| 17082 | 32  | 
val documentN = "document";  | 
33  | 
val document_path = Path.basic documentN;  | 
|
| 8196 | 34  | 
val doc_indexN = "session";  | 
| 
59448
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
35  | 
val session_graph_path = Path.basic "session_graph.pdf";  | 
| 7727 | 36  | 
|
| 62550 | 37  | 
fun show_path path = Path.implode (Path.expand (File.full_path Path.current path));  | 
| 11856 | 38  | 
|
| 7727 | 39  | 
|
| 14922 | 40  | 
|
| 7727 | 41  | 
(** additional theory data **)  | 
42  | 
||
| 
42008
 
7423e833a880
Present.init/finish/no_document are not thread-safe -- eliminated futile CRITICAL sections;
 
wenzelm 
parents: 
42007 
diff
changeset
 | 
43  | 
structure Browser_Info = Theory_Data  | 
| 22846 | 44  | 
(  | 
| 
51399
 
6ac3c29a300e
discontinued "isabelle usedir" option -r (reset session path);
 
wenzelm 
parents: 
51398 
diff
changeset
 | 
45  | 
  type T = {chapter: string, name: string};
 | 
| 
51567
 
a86c5e02ba58
proper default browser info for interactive mode, notably thy_deps;
 
wenzelm 
parents: 
51419 
diff
changeset
 | 
46  | 
  val empty = {chapter = "Unsorted", name = "Unknown"}: T;
 | 
| 16503 | 47  | 
fun extend _ = empty;  | 
| 33522 | 48  | 
fun merge _ = empty;  | 
| 22846 | 49  | 
);  | 
| 7727 | 50  | 
|
| 53171 | 51  | 
val _ = Theory.setup  | 
52  | 
  (Browser_Info.put {chapter = Context.PureN, name = Context.PureN});
 | 
|
| 
51567
 
a86c5e02ba58
proper default browser info for interactive mode, notably thy_deps;
 
wenzelm 
parents: 
51419 
diff
changeset
 | 
53  | 
|
| 7727 | 54  | 
|
55  | 
||
56  | 
(** global browser info state **)  | 
|
57  | 
||
58  | 
(* type theory_info *)  | 
|
59  | 
||
| 54456 | 60  | 
type theory_info = {tex_source: string, html_source: string};
 | 
| 7727 | 61  | 
|
| 54456 | 62  | 
fun make_theory_info (tex_source, html_source) =  | 
63  | 
  {tex_source = tex_source, html_source = html_source}: theory_info;
 | 
|
| 7727 | 64  | 
|
| 54456 | 65  | 
fun map_theory_info f {tex_source, html_source} =
 | 
66  | 
make_theory_info (f (tex_source, html_source));  | 
|
| 7727 | 67  | 
|
68  | 
||
69  | 
(* type browser_info *)  | 
|
70  | 
||
| 
54455
 
1d977436c1bf
removed remains of HTML presentation of auxiliary files -- inactive since Isabelle2013;
 
wenzelm 
parents: 
53171 
diff
changeset
 | 
71  | 
type browser_info =  | 
| 
 
1d977436c1bf
removed remains of HTML presentation of auxiliary files -- inactive since Isabelle2013;
 
wenzelm 
parents: 
53171 
diff
changeset
 | 
72  | 
 {theories: theory_info Symtab.table,
 | 
| 
 
1d977436c1bf
removed remains of HTML presentation of auxiliary files -- inactive since Isabelle2013;
 
wenzelm 
parents: 
53171 
diff
changeset
 | 
73  | 
tex_index: (int * string) list,  | 
| 
59448
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
74  | 
html_index: (int * string) list};  | 
| 7727 | 75  | 
|
| 
59448
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
76  | 
fun make_browser_info (theories, tex_index, html_index) : browser_info =  | 
| 
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
77  | 
  {theories = theories, tex_index = tex_index, html_index = html_index};
 | 
| 7727 | 78  | 
|
| 
59448
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
79  | 
val empty_browser_info = make_browser_info (Symtab.empty, [], []);  | 
| 7727 | 80  | 
|
| 
59448
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
81  | 
fun map_browser_info f {theories, tex_index, html_index} =
 | 
| 
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
82  | 
make_browser_info (f (theories, tex_index, html_index));  | 
| 7727 | 83  | 
|
84  | 
||
85  | 
(* state *)  | 
|
86  | 
||
| 59173 | 87  | 
val browser_info = Synchronized.var "browser_info" empty_browser_info;  | 
88  | 
fun change_browser_info f = Synchronized.change browser_info (map_browser_info f);  | 
|
| 7727 | 89  | 
|
90  | 
fun init_theory_info name info =  | 
|
| 
59448
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
91  | 
change_browser_info (fn (theories, tex_index, html_index) =>  | 
| 
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
92  | 
(Symtab.update (name, info) theories, tex_index, html_index));  | 
| 7727 | 93  | 
|
94  | 
fun change_theory_info name f =  | 
|
| 
59448
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
95  | 
change_browser_info (fn (theories, tex_index, html_index) =>  | 
| 17412 | 96  | 
(case Symtab.lookup theories name of  | 
| 
27491
 
c8178a6a6480
begin_theory: files_html needs to be produced outside of prep_html_source to make ML files appear!
 
wenzelm 
parents: 
27329 
diff
changeset
 | 
97  | 
      NONE => error ("Browser info: cannot access theory document " ^ quote name)
 | 
| 
59448
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
98  | 
| SOME info => (Symtab.update (name, map_theory_info f info) theories, tex_index, html_index)));  | 
| 7727 | 99  | 
|
100  | 
||
101  | 
fun add_tex_index txt =  | 
|
| 
59448
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
102  | 
change_browser_info (fn (theories, tex_index, html_index) =>  | 
| 
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
103  | 
(theories, txt :: tex_index, html_index));  | 
| 7727 | 104  | 
|
105  | 
fun add_html_index txt =  | 
|
| 
59448
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
106  | 
change_browser_info (fn (theories, tex_index, html_index) =>  | 
| 
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
107  | 
(theories, tex_index, txt :: html_index));  | 
| 7727 | 108  | 
|
109  | 
||
110  | 
||
111  | 
(** global session state **)  | 
|
112  | 
||
113  | 
(* session_info *)  | 
|
114  | 
||
115  | 
type session_info =  | 
|
| 61381 | 116  | 
  {symbols: HTML.symbols, name: string, chapter: string, info_path: Path.T, info: bool,
 | 
117  | 
doc_format: string, doc_output: Path.T option, doc_files: (Path.T * Path.T) list,  | 
|
118  | 
graph_file: Path.T, documents: (string * string) list, verbose: bool, readme: Path.T option};  | 
|
| 7727 | 119  | 
|
| 9416 | 120  | 
fun make_session_info  | 
| 61381 | 121  | 
(symbols, name, chapter, info_path, info, doc_format, doc_output, doc_files,  | 
| 59446 | 122  | 
graph_file, documents, verbose, readme) =  | 
| 61381 | 123  | 
  {symbols = symbols, name = name, chapter = chapter, info_path = info_path, info = info,
 | 
124  | 
doc_format = doc_format, doc_output = doc_output, doc_files = doc_files, graph_file = graph_file,  | 
|
| 59446 | 125  | 
documents = documents, verbose = verbose, readme = readme}: session_info;  | 
| 7685 | 126  | 
|
127  | 
||
| 7727 | 128  | 
(* state *)  | 
129  | 
||
| 32738 | 130  | 
val session_info = Unsynchronized.ref (NONE: session_info option);  | 
| 7727 | 131  | 
|
| 
51399
 
6ac3c29a300e
discontinued "isabelle usedir" option -r (reset session path);
 
wenzelm 
parents: 
51398 
diff
changeset
 | 
132  | 
fun with_session_info x f = (case ! session_info of NONE => x | SOME info => f info);  | 
| 7727 | 133  | 
|
| 66037 | 134  | 
val theory_qualifier = Resources.theory_qualifier o Context.theory_long_name;  | 
135  | 
||
| 65505 | 136  | 
fun is_session_theory thy =  | 
137  | 
(case ! session_info of  | 
|
138  | 
NONE => false  | 
|
| 66037 | 139  | 
  | SOME {name, ...} => name = theory_qualifier thy);
 | 
| 7727 | 140  | 
|
141  | 
||
| 14922 | 142  | 
(** document preparation **)  | 
| 7727 | 143  | 
|
| 56614 | 144  | 
(* options *)  | 
145  | 
||
146  | 
fun document_enabled s = s <> "" andalso s <> "false";  | 
|
| 17082 | 147  | 
|
| 56612 | 148  | 
fun document_variants str =  | 
149  | 
let  | 
|
150  | 
fun read_variant s =  | 
|
151  | 
(case space_explode "=" s of  | 
|
152  | 
[name] => (name, "")  | 
|
153  | 
| [name, tags] => (name, tags)  | 
|
154  | 
      | _ => error ("Malformed document variant specification: " ^ quote s));
 | 
|
155  | 
val variants = map read_variant (space_explode ":" str);  | 
|
156  | 
val _ =  | 
|
157  | 
(case duplicates (op =) (map #1 variants) of  | 
|
158  | 
[] => ()  | 
|
159  | 
      | dups => error ("Duplicate document variants: " ^ commas_quote dups));
 | 
|
160  | 
in variants end;  | 
|
| 17082 | 161  | 
|
162  | 
||
| 7727 | 163  | 
(* init session *)  | 
164  | 
||
| 61381 | 165  | 
fun init symbols build info info_path doc document_output doc_variants doc_files graph_file  | 
| 
59448
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
166  | 
(chapter, name) verbose =  | 
| 
56530
 
5c178501cf78
removed obsolete doc_dump option (see also 892061142ba6);
 
wenzelm 
parents: 
54683 
diff
changeset
 | 
167  | 
if not build andalso not info andalso doc = "" then  | 
| 59173 | 168  | 
(Synchronized.change browser_info (K empty_browser_info); session_info := NONE)  | 
| 11856 | 169  | 
else  | 
170  | 
let  | 
|
| 
51399
 
6ac3c29a300e
discontinued "isabelle usedir" option -r (reset session path);
 
wenzelm 
parents: 
51398 
diff
changeset
 | 
171  | 
val doc_output =  | 
| 
 
6ac3c29a300e
discontinued "isabelle usedir" option -r (reset session path);
 
wenzelm 
parents: 
51398 
diff
changeset
 | 
172  | 
if document_output = "" then NONE else SOME (Path.explode document_output);  | 
| 7727 | 173  | 
|
| 42007 | 174  | 
val documents =  | 
175  | 
if doc = "" then []  | 
|
| 
56533
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56531 
diff
changeset
 | 
176  | 
else if null doc_files andalso not (can File.check_dir document_path) then  | 
| 
44389
 
a3b5fdfb04a3
tuned signature -- contrast physical output primitives versus Output.raw_message;
 
wenzelm 
parents: 
43850 
diff
changeset
 | 
177  | 
(if verbose then Output.physical_stderr "Warning: missing document directory\n"  | 
| 
 
a3b5fdfb04a3
tuned signature -- contrast physical output primitives versus Output.raw_message;
 
wenzelm 
parents: 
43850 
diff
changeset
 | 
178  | 
else (); [])  | 
| 
48804
 
6348e5fca42e
more direct interpretation of document_variants for build (unchanged for usedir);
 
wenzelm 
parents: 
48543 
diff
changeset
 | 
179  | 
else doc_variants;  | 
| 7727 | 180  | 
|
| 
51419
 
5313abe76bd4
include only README.html, not historic README, which tends towards surprises like src/HOL/SPARK/Examples/README;
 
wenzelm 
parents: 
51415 
diff
changeset
 | 
181  | 
val readme = if File.exists readme_html_path then SOME readme_html_path else NONE;  | 
| 7727 | 182  | 
|
| 17082 | 183  | 
val docs =  | 
184  | 
(case readme of NONE => [] | SOME p => [(Url.File p, "README")]) @  | 
|
| 42007 | 185  | 
map (fn (name, _) => (Url.File (Path.ext doc (Path.basic name)), name)) documents;  | 
| 11856 | 186  | 
in  | 
| 42007 | 187  | 
session_info :=  | 
| 61381 | 188  | 
SOME (make_session_info (symbols, name, chapter, info_path, info, doc,  | 
| 59446 | 189  | 
doc_output, doc_files, graph_file, documents, verbose, readme));  | 
| 
59448
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
190  | 
Synchronized.change browser_info (K empty_browser_info);  | 
| 61381 | 191  | 
add_html_index (0, HTML.begin_session_index symbols name (Url.File session_graph_path) docs)  | 
| 
42008
 
7423e833a880
Present.init/finish/no_document are not thread-safe -- eliminated futile CRITICAL sections;
 
wenzelm 
parents: 
42007 
diff
changeset
 | 
192  | 
end;  | 
| 7727 | 193  | 
|
194  | 
||
| 
28496
 
4cff10648928
renamed isatool to isabelle_tool in programming interfaces;
 
wenzelm 
parents: 
28375 
diff
changeset
 | 
195  | 
(* isabelle tool wrappers *)  | 
| 17082 | 196  | 
|
| 
48805
 
c3ea910b3581
explicit document_output directory, without implicit purge of default in ISABELLE_BROWSER_INFO;
 
wenzelm 
parents: 
48804 
diff
changeset
 | 
197  | 
fun isabelle_document {verbose, purge} format name tags dir =
 | 
| 17082 | 198  | 
let  | 
| 
66680
 
74a1b722507e
clarified "purge": retain .aux files etc. before "isabelle document", to allow 'document_files' providing such generated files (see also c3ea910b3581, 38ce936acb99);
 
wenzelm 
parents: 
66037 
diff
changeset
 | 
199  | 
val s = "isabelle document -o '" ^ format ^ "' \  | 
| 
62549
 
9498623b27f0
File.bash_string operations in ML as in Scala -- exclusively for GNU bash, not perl and not user output;
 
wenzelm 
parents: 
61381 
diff
changeset
 | 
200  | 
\-n '" ^ name ^ "' -t '" ^ tags ^ "' " ^ File.bash_path dir ^ " 2>&1";  | 
| 
48805
 
c3ea910b3581
explicit document_output directory, without implicit purge of default in ISABELLE_BROWSER_INFO;
 
wenzelm 
parents: 
48804 
diff
changeset
 | 
201  | 
val doc_path = Path.appends [dir, Path.parent, Path.basic name |> Path.ext format];  | 
| 
33682
 
0c5d1485dea7
isabelle_document: more explicit error output, notably for drafts;
 
wenzelm 
parents: 
33522 
diff
changeset
 | 
202  | 
val _ = if verbose then writeln s else ();  | 
| 
43850
 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 
wenzelm 
parents: 
43712 
diff
changeset
 | 
203  | 
val (out, rc) = Isabelle_System.bash_output s;  | 
| 
33682
 
0c5d1485dea7
isabelle_document: more explicit error output, notably for drafts;
 
wenzelm 
parents: 
33522 
diff
changeset
 | 
204  | 
val _ =  | 
| 
 
0c5d1485dea7
isabelle_document: more explicit error output, notably for drafts;
 
wenzelm 
parents: 
33522 
diff
changeset
 | 
205  | 
if not (File.exists doc_path) orelse rc <> 0 then  | 
| 
 
0c5d1485dea7
isabelle_document: more explicit error output, notably for drafts;
 
wenzelm 
parents: 
33522 
diff
changeset
 | 
206  | 
        cat_error out ("Failed to build document " ^ quote (show_path doc_path))
 | 
| 
 
0c5d1485dea7
isabelle_document: more explicit error output, notably for drafts;
 
wenzelm 
parents: 
33522 
diff
changeset
 | 
207  | 
else if verbose then writeln out  | 
| 
 
0c5d1485dea7
isabelle_document: more explicit error output, notably for drafts;
 
wenzelm 
parents: 
33522 
diff
changeset
 | 
208  | 
else ();  | 
| 
66680
 
74a1b722507e
clarified "purge": retain .aux files etc. before "isabelle document", to allow 'document_files' providing such generated files (see also c3ea910b3581, 38ce936acb99);
 
wenzelm 
parents: 
66037 
diff
changeset
 | 
209  | 
val _ = if purge andalso rc = 0 then Isabelle_System.rm_tree dir else ();  | 
| 
33682
 
0c5d1485dea7
isabelle_document: more explicit error output, notably for drafts;
 
wenzelm 
parents: 
33522 
diff
changeset
 | 
210  | 
in doc_path end;  | 
| 17082 | 211  | 
|
212  | 
||
| 11856 | 213  | 
(* finish session -- output all generated text *)  | 
214  | 
||
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58870 
diff
changeset
 | 
215  | 
fun sorted_index index = map snd (sort (int_ord o apply2 fst) (rev index));  | 
| 
24150
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
216  | 
fun index_buffer index = Buffer.add (implode (sorted_index index)) Buffer.empty;  | 
| 
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
217  | 
|
| 14922 | 218  | 
fun write_tex src name path =  | 
| 28027 | 219  | 
File.write_buffer (Path.append path (tex_path name)) src;  | 
| 14922 | 220  | 
|
221  | 
fun write_tex_index tex_index path =  | 
|
| 
24150
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
222  | 
write_tex (index_buffer tex_index |> Buffer.add Latex.tex_trailer) doc_indexN path;  | 
| 14922 | 223  | 
|
| 
42008
 
7423e833a880
Present.init/finish/no_document are not thread-safe -- eliminated futile CRITICAL sections;
 
wenzelm 
parents: 
42007 
diff
changeset
 | 
224  | 
fun finish () =  | 
| 59446 | 225  | 
  with_session_info () (fn {name, chapter, info, info_path, doc_format,
 | 
| 59445 | 226  | 
doc_output, doc_files, graph_file, documents, verbose, readme, ...} =>  | 
| 7727 | 227  | 
let  | 
| 
59448
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
228  | 
    val {theories, tex_index, html_index} = Synchronized.value browser_info;
 | 
| 11856 | 229  | 
val thys = Symtab.dest theories;  | 
| 
51399
 
6ac3c29a300e
discontinued "isabelle usedir" option -r (reset session path);
 
wenzelm 
parents: 
51398 
diff
changeset
 | 
230  | 
|
| 61372 | 231  | 
val session_prefix =  | 
232  | 
Path.append (Path.append info_path (Path.basic chapter)) (Path.basic name);  | 
|
| 7727 | 233  | 
|
| 54456 | 234  | 
    fun finish_html (a, {html_source, ...}: theory_info) =
 | 
235  | 
File.write (Path.append session_prefix (html_path a)) html_source;  | 
|
| 11856 | 236  | 
|
| 42007 | 237  | 
val _ =  | 
238  | 
if info then  | 
|
| 
51399
 
6ac3c29a300e
discontinued "isabelle usedir" option -r (reset session path);
 
wenzelm 
parents: 
51398 
diff
changeset
 | 
239  | 
(Isabelle_System.mkdirs session_prefix;  | 
| 
 
6ac3c29a300e
discontinued "isabelle usedir" option -r (reset session path);
 
wenzelm 
parents: 
51398 
diff
changeset
 | 
240  | 
File.write_buffer (Path.append session_prefix index_path)  | 
| 
 
6ac3c29a300e
discontinued "isabelle usedir" option -r (reset session path);
 
wenzelm 
parents: 
51398 
diff
changeset
 | 
241  | 
(index_buffer html_index |> Buffer.add HTML.end_document);  | 
| 
56533
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56531 
diff
changeset
 | 
242  | 
(case readme of NONE => () | SOME path => Isabelle_System.copy_file path session_prefix);  | 
| 42007 | 243  | 
List.app finish_html thys;  | 
| 
54455
 
1d977436c1bf
removed remains of HTML presentation of auxiliary files -- inactive since Isabelle2013;
 
wenzelm 
parents: 
53171 
diff
changeset
 | 
244  | 
if verbose  | 
| 
 
1d977436c1bf
removed remains of HTML presentation of auxiliary files -- inactive since Isabelle2013;
 
wenzelm 
parents: 
53171 
diff
changeset
 | 
245  | 
        then Output.physical_stderr ("Browser info at " ^ show_path session_prefix ^ "\n")
 | 
| 42007 | 246  | 
else ())  | 
247  | 
else ();  | 
|
248  | 
||
| 
56533
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56531 
diff
changeset
 | 
249  | 
fun document_job doc_prefix backdrop (doc_name, tags) =  | 
| 
48935
 
4c92a2f310b6
clarified document directories: browser_info as backdrop vs. optional output directory in the foreground;
 
wenzelm 
parents: 
48933 
diff
changeset
 | 
250  | 
let  | 
| 
56533
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56531 
diff
changeset
 | 
251  | 
val doc_dir = Path.append doc_prefix (Path.basic doc_name);  | 
| 56531 | 252  | 
val _ = Isabelle_System.mkdirs doc_dir;  | 
253  | 
val _ =  | 
|
| 
62829
 
4141c2a8458b
clarified Isabelle tool wrapper: bash, Scala, no perl, no ML;
 
wenzelm 
parents: 
62589 
diff
changeset
 | 
254  | 
          Isabelle_System.bash ("isabelle latex -o sty " ^
 | 
| 
 
4141c2a8458b
clarified Isabelle tool wrapper: bash, Scala, no perl, no ML;
 
wenzelm 
parents: 
62589 
diff
changeset
 | 
255  | 
File.bash_path (Path.append doc_dir (Path.basic "root.tex")));  | 
| 57885 | 256  | 
val _ = List.app (fn file => Isabelle_System.copy_file_base file doc_dir) doc_files;  | 
| 
59448
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
257  | 
val _ = Isabelle_System.copy_file graph_file (Path.append doc_dir session_graph_path);  | 
| 56531 | 258  | 
val _ = write_tex_index tex_index doc_dir;  | 
259  | 
val _ =  | 
|
260  | 
          List.app (fn (a, {tex_source, ...}) =>
 | 
|
261  | 
write_tex (Buffer.add tex_source Buffer.empty) a doc_dir) thys;  | 
|
| 
48935
 
4c92a2f310b6
clarified document directories: browser_info as backdrop vs. optional output directory in the foreground;
 
wenzelm 
parents: 
48933 
diff
changeset
 | 
262  | 
in  | 
| 
 
4c92a2f310b6
clarified document directories: browser_info as backdrop vs. optional output directory in the foreground;
 
wenzelm 
parents: 
48933 
diff
changeset
 | 
263  | 
fn () =>  | 
| 
56533
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56531 
diff
changeset
 | 
264  | 
          (isabelle_document {verbose = true, purge = backdrop} doc_format doc_name tags doc_dir,
 | 
| 56531 | 265  | 
fn doc =>  | 
266  | 
if verbose orelse not backdrop then  | 
|
267  | 
                Output.physical_stderr ("Document at " ^ show_path doc ^ "\n")
 | 
|
268  | 
else ())  | 
|
| 
48935
 
4c92a2f310b6
clarified document directories: browser_info as backdrop vs. optional output directory in the foreground;
 
wenzelm 
parents: 
48933 
diff
changeset
 | 
269  | 
end;  | 
| 
 
4c92a2f310b6
clarified document directories: browser_info as backdrop vs. optional output directory in the foreground;
 
wenzelm 
parents: 
48933 
diff
changeset
 | 
270  | 
|
| 
 
4c92a2f310b6
clarified document directories: browser_info as backdrop vs. optional output directory in the foreground;
 
wenzelm 
parents: 
48933 
diff
changeset
 | 
271  | 
val jobs =  | 
| 
 
4c92a2f310b6
clarified document directories: browser_info as backdrop vs. optional output directory in the foreground;
 
wenzelm 
parents: 
48933 
diff
changeset
 | 
272  | 
(if info orelse is_none doc_output then  | 
| 
51399
 
6ac3c29a300e
discontinued "isabelle usedir" option -r (reset session path);
 
wenzelm 
parents: 
51398 
diff
changeset
 | 
273  | 
map (document_job session_prefix true) documents  | 
| 
48935
 
4c92a2f310b6
clarified document directories: browser_info as backdrop vs. optional output directory in the foreground;
 
wenzelm 
parents: 
48933 
diff
changeset
 | 
274  | 
else []) @  | 
| 
 
4c92a2f310b6
clarified document directories: browser_info as backdrop vs. optional output directory in the foreground;
 
wenzelm 
parents: 
48933 
diff
changeset
 | 
275  | 
(case doc_output of  | 
| 
 
4c92a2f310b6
clarified document directories: browser_info as backdrop vs. optional output directory in the foreground;
 
wenzelm 
parents: 
48933 
diff
changeset
 | 
276  | 
NONE => []  | 
| 
 
4c92a2f310b6
clarified document directories: browser_info as backdrop vs. optional output directory in the foreground;
 
wenzelm 
parents: 
48933 
diff
changeset
 | 
277  | 
| SOME path => map (document_job path false) documents);  | 
| 
 
4c92a2f310b6
clarified document directories: browser_info as backdrop vs. optional output directory in the foreground;
 
wenzelm 
parents: 
48933 
diff
changeset
 | 
278  | 
|
| 
 
4c92a2f310b6
clarified document directories: browser_info as backdrop vs. optional output directory in the foreground;
 
wenzelm 
parents: 
48933 
diff
changeset
 | 
279  | 
val _ = jobs |> Par_List.map (fn job => job ()) |> List.app (op |>);  | 
| 7727 | 280  | 
in  | 
| 59173 | 281  | 
Synchronized.change browser_info (K empty_browser_info);  | 
| 15531 | 282  | 
session_info := NONE  | 
| 
42008
 
7423e833a880
Present.init/finish/no_document are not thread-safe -- eliminated futile CRITICAL sections;
 
wenzelm 
parents: 
42007 
diff
changeset
 | 
283  | 
end);  | 
| 7727 | 284  | 
|
285  | 
||
286  | 
(* theory elements *)  | 
|
287  | 
||
| 65505 | 288  | 
fun theory_output thy tex =  | 
| 
54458
 
96ccc8972fc7
prefer explicit "document" flag -- eliminated stateful Present.no_document;
 
wenzelm 
parents: 
54456 
diff
changeset
 | 
289  | 
with_session_info () (fn _ =>  | 
| 65505 | 290  | 
if is_session_theory thy then  | 
291  | 
let val name = Context.theory_name thy;  | 
|
292  | 
in change_theory_info name (fn (_, html) => (Latex.isabelle_theory name tex, html)) end  | 
|
293  | 
else ());  | 
|
| 7727 | 294  | 
|
| 
59448
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
295  | 
fun theory_link (curr_chapter, curr_session) thy =  | 
| 
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
296  | 
let  | 
| 
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
297  | 
    val {chapter, name = session} = Browser_Info.get thy;
 | 
| 
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
298  | 
val link = html_path (Context.theory_name thy);  | 
| 
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
299  | 
in  | 
| 
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
300  | 
if curr_session = session then SOME link  | 
| 
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
301  | 
else if curr_chapter = chapter then  | 
| 
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
302  | 
SOME (Path.appends [Path.parent, Path.basic session, link])  | 
| 
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
303  | 
else if chapter = Context.PureN then NONE  | 
| 
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
304  | 
else SOME (Path.appends [Path.parent, Path.parent, Path.basic chapter, Path.basic session, link])  | 
| 
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
305  | 
end;  | 
| 
 
149d2bc5ddb6
prefer plain session_graph.pdf over GraphBrowser applet;
 
wenzelm 
parents: 
59446 
diff
changeset
 | 
306  | 
|
| 54456 | 307  | 
fun begin_theory update_time mk_text thy =  | 
| 61381 | 308  | 
  with_session_info thy (fn {symbols, name = session_name, chapter, ...} =>
 | 
| 54456 | 309  | 
let  | 
310  | 
val name = Context.theory_name thy;  | 
|
311  | 
||
| 65505 | 312  | 
val parent_specs =  | 
313  | 
Theory.parents_of thy |> map (fn parent =>  | 
|
314  | 
(Option.map Url.File (theory_link (chapter, session_name) parent),  | 
|
315  | 
(Context.theory_name parent)));  | 
|
| 61381 | 316  | 
val html_source = HTML.theory symbols name parent_specs (mk_text ());  | 
| 65505 | 317  | 
      val _ = init_theory_info name (make_theory_info ("", html_source));
 | 
318  | 
||
319  | 
val _ =  | 
|
320  | 
if is_session_theory thy then  | 
|
321  | 
(add_html_index (update_time, HTML.theory_entry symbols (Url.File (html_path name), name));  | 
|
322  | 
add_tex_index (update_time, Latex.theory_entry name))  | 
|
323  | 
else ();  | 
|
324  | 
    in Browser_Info.put {chapter = chapter, name = session_name} thy end);
 | 
|
| 7727 | 325  | 
|
326  | 
||
327  | 
||
| 
42127
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
42126 
diff
changeset
 | 
328  | 
(** draft document output **)  | 
| 14922 | 329  | 
|
| 52549 | 330  | 
fun display_drafts src_paths = Isabelle_System.with_tmp_dir "drafts" (fn dir =>  | 
| 14922 | 331  | 
let  | 
| 
24150
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
332  | 
fun prep_draft path i =  | 
| 14935 | 333  | 
let  | 
334  | 
val base = Path.base path;  | 
|
| 14972 | 335  | 
val name =  | 
| 24829 | 336  | 
(case Path.implode (#1 (Path.split_ext base)) of  | 
| 
44986
 
6f27ecf2a951
unique file names via serial numbers, to allow files like "root" or multiple files with same base name;
 
wenzelm 
parents: 
44389 
diff
changeset
 | 
337  | 
"" => "DUMMY"  | 
| 
 
6f27ecf2a951
unique file names via serial numbers, to allow files like "root" or multiple files with same base name;
 
wenzelm 
parents: 
44389 
diff
changeset
 | 
338  | 
| s => s) ^ serial_string ();  | 
| 14935 | 339  | 
in  | 
340  | 
if File.exists path then  | 
|
| 
24150
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
341  | 
(((name, base, File.read path), (i, Latex.theory_entry name)), i + 1)  | 
| 
41944
 
b97091ae583a
Path.print is the official way to show file-system paths to users -- note that Path.implode often indicates violation of the abstract datatype;
 
wenzelm 
parents: 
41886 
diff
changeset
 | 
342  | 
        else error ("Bad file: " ^ Path.print path)
 | 
| 14935 | 343  | 
end;  | 
| 
24150
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
344  | 
val (srcs, tex_index) = split_list (fst (fold_map prep_draft src_paths 0));  | 
| 14935 | 345  | 
|
| 
42127
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
42126 
diff
changeset
 | 
346  | 
val doc_path = Path.append dir document_path;  | 
| 40743 | 347  | 
val _ = Isabelle_System.mkdirs doc_path;  | 
| 14922 | 348  | 
val root_path = Path.append doc_path (Path.basic "root.tex");  | 
| 
56533
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56531 
diff
changeset
 | 
349  | 
val _ = Isabelle_System.copy_file (Path.explode "~~/lib/texinputs/draft.tex") root_path;  | 
| 
62829
 
4141c2a8458b
clarified Isabelle tool wrapper: bash, Scala, no perl, no ML;
 
wenzelm 
parents: 
62589 
diff
changeset
 | 
350  | 
    val _ = Isabelle_System.bash ("isabelle latex -o sty " ^ File.bash_path root_path);
 | 
| 
 
4141c2a8458b
clarified Isabelle tool wrapper: bash, Scala, no perl, no ML;
 
wenzelm 
parents: 
62589 
diff
changeset
 | 
351  | 
    val _ = Isabelle_System.bash ("isabelle latex -o syms " ^ File.bash_path root_path);
 | 
| 14922 | 352  | 
|
353  | 
fun known name =  | 
|
354  | 
let val ss = split_lines (File.read (Path.append doc_path (Path.basic name)))  | 
|
| 20664 | 355  | 
in member (op =) ss end;  | 
| 14922 | 356  | 
val known_syms = known "syms.lst";  | 
357  | 
val known_ctrls = known "ctrls.lst";  | 
|
358  | 
||
| 15570 | 359  | 
val _ = srcs |> List.app (fn (name, base, txt) =>  | 
| 14935 | 360  | 
Symbol.explode txt  | 
| 
21858
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
20664 
diff
changeset
 | 
361  | 
|> Latex.symbol_source (known_syms, known_ctrls) (Path.implode base)  | 
| 14935 | 362  | 
|> File.write (Path.append doc_path (tex_path name)));  | 
| 14922 | 363  | 
val _ = write_tex_index tex_index doc_path;  | 
364  | 
||
| 
48805
 
c3ea910b3581
explicit document_output directory, without implicit purge of default in ISABELLE_BROWSER_INFO;
 
wenzelm 
parents: 
48804 
diff
changeset
 | 
365  | 
val result =  | 
| 52743 | 366  | 
      isabelle_document {verbose = false, purge = true} "pdf" documentN "" doc_path;
 | 
| 
54683
 
cf48ddc266e5
clarified "isabelle display" and 'display_drafts': re-use file and program instance, open asynchronously via desktop environment;
 
wenzelm 
parents: 
54458 
diff
changeset
 | 
367  | 
|
| 
 
cf48ddc266e5
clarified "isabelle display" and 'display_drafts': re-use file and program instance, open asynchronously via desktop environment;
 
wenzelm 
parents: 
54458 
diff
changeset
 | 
368  | 
val target_dir = Path.explode "$ISABELLE_HOME_USER/tmp";  | 
| 
 
cf48ddc266e5
clarified "isabelle display" and 'display_drafts': re-use file and program instance, open asynchronously via desktop environment;
 
wenzelm 
parents: 
54458 
diff
changeset
 | 
369  | 
val target = Path.explode "$ISABELLE_HOME_USER/tmp/drafts.pdf"  | 
| 
 
cf48ddc266e5
clarified "isabelle display" and 'display_drafts': re-use file and program instance, open asynchronously via desktop environment;
 
wenzelm 
parents: 
54458 
diff
changeset
 | 
370  | 
val _ = Isabelle_System.mkdirs target_dir;  | 
| 
56533
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56531 
diff
changeset
 | 
371  | 
val _ = Isabelle_System.copy_file result target;  | 
| 
62829
 
4141c2a8458b
clarified Isabelle tool wrapper: bash, Scala, no perl, no ML;
 
wenzelm 
parents: 
62589 
diff
changeset
 | 
372  | 
  in Isabelle_System.bash ("isabelle display " ^ File.bash_path target ^ " &") end);
 | 
| 14922 | 373  | 
|
| 7685 | 374  | 
end;  |