| author | wenzelm | 
| Sun, 04 Mar 2012 16:02:14 +0100 | |
| changeset 46811 | 03a2dc9e0624 | 
| parent 44986 | 6f27ecf2a951 | 
| child 48445 | cb4136e4cabf | 
| permissions | -rw-r--r-- | 
| 6203 | 1  | 
(* Title: Pure/Thy/present.ML  | 
| 9416 | 2  | 
Author: Markus Wenzel and Stefan Berghofer, TU Muenchen  | 
| 6203 | 3  | 
|
| 9416 | 4  | 
Theory presentation: HTML, graph files, (PDF)LaTeX documents.  | 
| 6203 | 5  | 
*)  | 
6  | 
||
7  | 
signature BASIC_PRESENT =  | 
|
8  | 
sig  | 
|
| 
42008
 
7423e833a880
Present.init/finish/no_document are not thread-safe -- eliminated futile CRITICAL sections;
 
wenzelm 
parents: 
42007 
diff
changeset
 | 
9  | 
  val no_document: ('a -> 'b) -> 'a -> 'b  (*not thread-safe!*)
 | 
| 6203 | 10  | 
end;  | 
11  | 
||
12  | 
signature PRESENT =  | 
|
13  | 
sig  | 
|
| 7727 | 14  | 
include BASIC_PRESENT  | 
| 24561 | 15  | 
val session_name: theory -> string  | 
| 9452 | 16  | 
  val write_graph: {name: string, ID: string, dir: string, unfold: bool,
 | 
17  | 
path: string, parents: string list} list -> Path.T -> unit  | 
|
| 20577 | 18  | 
  val display_graph: {name: string, ID: string, dir: string, unfold: bool,
 | 
19  | 
path: string, parents: string list} list -> unit  | 
|
| 17082 | 20  | 
val init: bool -> bool -> string -> bool -> string list -> string list ->  | 
| 
42008
 
7423e833a880
Present.init/finish/no_document are not thread-safe -- eliminated futile CRITICAL sections;
 
wenzelm 
parents: 
42007 
diff
changeset
 | 
21  | 
string -> (bool * Path.T) option -> Url.T option * bool -> bool ->  | 
| 
 
7423e833a880
Present.init/finish/no_document are not thread-safe -- eliminated futile CRITICAL sections;
 
wenzelm 
parents: 
42007 
diff
changeset
 | 
22  | 
theory list -> unit (*not thread-safe!*)  | 
| 
 
7423e833a880
Present.init/finish/no_document are not thread-safe -- eliminated futile CRITICAL sections;
 
wenzelm 
parents: 
42007 
diff
changeset
 | 
23  | 
val finish: unit -> unit (*not thread-safe!*)  | 
| 7727 | 24  | 
val init_theory: string -> unit  | 
| 
27862
 
8f727f7c8c1d
removed obsolete verbatim_source, results, chapter, section etc.;
 
wenzelm 
parents: 
27491 
diff
changeset
 | 
25  | 
val theory_source: string -> (unit -> HTML.text) -> unit  | 
| 9136 | 26  | 
val theory_output: string -> string -> unit  | 
| 
24150
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
27  | 
val begin_theory: int -> Path.T -> (Path.T * bool) list -> theory -> theory  | 
| 14922 | 28  | 
val drafts: string -> Path.T list -> Path.T  | 
| 6203 | 29  | 
end;  | 
30  | 
||
| 7685 | 31  | 
structure Present: PRESENT =  | 
32  | 
struct  | 
|
33  | 
||
| 7727 | 34  | 
|
35  | 
(** paths **)  | 
|
36  | 
||
37  | 
val output_path = Path.variable "ISABELLE_BROWSER_INFO";  | 
|
38  | 
||
39  | 
val tex_ext = Path.ext "tex";  | 
|
40  | 
val tex_path = tex_ext o Path.basic;  | 
|
41  | 
val html_ext = Path.ext "html";  | 
|
42  | 
val html_path = html_ext o Path.basic;  | 
|
43  | 
val index_path = Path.basic "index.html";  | 
|
| 11856 | 44  | 
val readme_html_path = Path.basic "README.html";  | 
45  | 
val readme_path = Path.basic "README";  | 
|
| 17082 | 46  | 
val documentN = "document";  | 
47  | 
val document_path = Path.basic documentN;  | 
|
| 8196 | 48  | 
val doc_indexN = "session";  | 
| 11856 | 49  | 
val graph_path = Path.basic "session.graph";  | 
50  | 
val graph_pdf_path = Path.basic "session_graph.pdf";  | 
|
51  | 
val graph_eps_path = Path.basic "session_graph.eps";  | 
|
| 7727 | 52  | 
|
53  | 
val session_path = Path.basic ".session";  | 
|
| 
21858
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
20664 
diff
changeset
 | 
54  | 
val session_entries_path = Path.explode ".session/entries";  | 
| 
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
20664 
diff
changeset
 | 
55  | 
val pre_index_path = Path.explode ".session/pre-index";  | 
| 7727 | 56  | 
|
| 9044 | 57  | 
fun mk_rel_path [] ys = Path.make ys  | 
58  | 
| mk_rel_path xs [] = Path.appends (replicate (length xs) Path.parent)  | 
|
| 9416 | 59  | 
| mk_rel_path (ps as x :: xs) (qs as y :: ys) = if x = y then mk_rel_path xs ys else  | 
| 9044 | 60  | 
Path.appends (replicate (length ps) Path.parent @ [Path.make qs]);  | 
| 7727 | 61  | 
|
| 
21858
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
20664 
diff
changeset
 | 
62  | 
fun show_path path = Path.implode (Path.append (File.pwd ()) path);  | 
| 11856 | 63  | 
|
| 7727 | 64  | 
|
| 14922 | 65  | 
|
| 7727 | 66  | 
(** additional theory data **)  | 
67  | 
||
| 
42008
 
7423e833a880
Present.init/finish/no_document are not thread-safe -- eliminated futile CRITICAL sections;
 
wenzelm 
parents: 
42007 
diff
changeset
 | 
68  | 
structure Browser_Info = Theory_Data  | 
| 22846 | 69  | 
(  | 
| 9416 | 70  | 
  type T = {name: string, session: string list, is_local: bool};
 | 
| 27329 | 71  | 
  val empty = {name = "", session = [], is_local = false}: T;
 | 
| 16503 | 72  | 
fun extend _ = empty;  | 
| 33522 | 73  | 
fun merge _ = empty;  | 
| 22846 | 74  | 
);  | 
| 7727 | 75  | 
|
| 
42008
 
7423e833a880
Present.init/finish/no_document are not thread-safe -- eliminated futile CRITICAL sections;
 
wenzelm 
parents: 
42007 
diff
changeset
 | 
76  | 
val put_info = Browser_Info.put;  | 
| 
 
7423e833a880
Present.init/finish/no_document are not thread-safe -- eliminated futile CRITICAL sections;
 
wenzelm 
parents: 
42007 
diff
changeset
 | 
77  | 
val get_info = Browser_Info.get;  | 
| 24561 | 78  | 
val session_name = #name o get_info;  | 
79  | 
||
| 7727 | 80  | 
|
81  | 
||
82  | 
(** graphs **)  | 
|
83  | 
||
84  | 
type graph_node =  | 
|
85  | 
  {name: string, ID: string, dir: string, unfold: bool,
 | 
|
86  | 
path: string, parents: string list};  | 
|
87  | 
||
| 9416 | 88  | 
fun write_graph gr path =  | 
89  | 
  File.write path (cat_lines (map (fn {name, ID, dir, unfold, path, parents} =>
 | 
|
90  | 
"\"" ^ name ^ "\" \"" ^ ID ^ "\" \"" ^ dir ^ (if unfold then "\" + \"" else "\" \"") ^  | 
|
| 28840 | 91  | 
path ^ "\" > " ^ space_implode " " (map quote parents) ^ " ;") gr));  | 
| 7727 | 92  | 
|
| 20577 | 93  | 
fun display_graph gr =  | 
94  | 
let  | 
|
| 
42127
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
42126 
diff
changeset
 | 
95  | 
val path = Isabelle_System.create_tmp_path "graph" "";  | 
| 
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
42126 
diff
changeset
 | 
96  | 
val _ = write_graph gr path;  | 
| 24561 | 97  | 
val _ = writeln "Displaying graph ...";  | 
| 40743 | 98  | 
    val _ = Isabelle_System.isabelle_tool "browser" ("-c " ^ File.shell_path path ^ " &");
 | 
| 20577 | 99  | 
in () end;  | 
100  | 
||
101  | 
||
| 9416 | 102  | 
fun ID_of sess s = space_implode "/" (sess @ [s]);  | 
| 
23899
 
ab37b1f690c7
clarified init/begin_theory: no longer depend on thy_info.ML;
 
wenzelm 
parents: 
23884 
diff
changeset
 | 
103  | 
fun ID_of_thy thy = ID_of (#session (get_info thy)) (Context.theory_name thy);  | 
| 7727 | 104  | 
|
| 
23899
 
ab37b1f690c7
clarified init/begin_theory: no longer depend on thy_info.ML;
 
wenzelm 
parents: 
23884 
diff
changeset
 | 
105  | 
|
| 
 
ab37b1f690c7
clarified init/begin_theory: no longer depend on thy_info.ML;
 
wenzelm 
parents: 
23884 
diff
changeset
 | 
106  | 
(*retrieve graph data from initial collection of theories*)  | 
| 
24150
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
107  | 
fun init_graph remote_path curr_sess = rev o map (fn thy =>  | 
| 7727 | 108  | 
let  | 
| 
23899
 
ab37b1f690c7
clarified init/begin_theory: no longer depend on thy_info.ML;
 
wenzelm 
parents: 
23884 
diff
changeset
 | 
109  | 
val name = Context.theory_name thy;  | 
| 
 
ab37b1f690c7
clarified init/begin_theory: no longer depend on thy_info.ML;
 
wenzelm 
parents: 
23884 
diff
changeset
 | 
110  | 
    val {name = sess_name, session, is_local} = get_info thy;
 | 
| 
24150
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
111  | 
val entry =  | 
| 
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
112  | 
     {name = name, ID = ID_of session name, dir = sess_name,
 | 
| 
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
113  | 
path =  | 
| 
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
114  | 
if null session then "" else  | 
| 
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
115  | 
if is_some remote_path andalso not is_local then  | 
| 
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
116  | 
Url.implode (Url.append (the remote_path) (Url.File  | 
| 
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
117  | 
(Path.append (Path.make session) (html_path name))))  | 
| 
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
118  | 
else Path.implode (Path.append (mk_rel_path curr_sess session) (html_path name)),  | 
| 
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
119  | 
unfold = false,  | 
| 
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
120  | 
parents = map ID_of_thy (Theory.parents_of thy)};  | 
| 
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
121  | 
in (0, entry) end);  | 
| 7727 | 122  | 
|
| 
24150
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
123  | 
fun ins_graph_entry (i, entry as {ID, ...}) (gr: (int * graph_node) list) =
 | 
| 
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
124  | 
(i, entry) :: filter_out (fn (_, entry') => #ID entry' = ID) gr;  | 
| 7727 | 125  | 
|
126  | 
||
127  | 
||
128  | 
(** global browser info state **)  | 
|
129  | 
||
130  | 
(* type theory_info *)  | 
|
131  | 
||
132  | 
type theory_info = {tex_source: Buffer.T, html_source: Buffer.T, html: Buffer.T};
 | 
|
133  | 
||
134  | 
fun make_theory_info (tex_source, html_source, html) =  | 
|
135  | 
  {tex_source = tex_source, html_source = html_source, html = html}: theory_info;
 | 
|
136  | 
||
137  | 
val empty_theory_info = make_theory_info (Buffer.empty, Buffer.empty, Buffer.empty);  | 
|
138  | 
||
139  | 
fun map_theory_info f {tex_source, html_source, html} =
 | 
|
140  | 
make_theory_info (f (tex_source, html_source, html));  | 
|
141  | 
||
142  | 
||
143  | 
(* type browser_info *)  | 
|
144  | 
||
| 11856 | 145  | 
type browser_info = {theories: theory_info Symtab.table, files: (Path.T * string) list,
 | 
| 
24150
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
146  | 
tex_index: (int * string) list, html_index: (int * string) list, graph: (int * graph_node) list};  | 
| 7727 | 147  | 
|
| 11856 | 148  | 
fun make_browser_info (theories, files, tex_index, html_index, graph) =  | 
149  | 
  {theories = theories, files = files, tex_index = tex_index, html_index = html_index,
 | 
|
| 9416 | 150  | 
graph = graph}: browser_info;  | 
| 7727 | 151  | 
|
| 
24150
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
152  | 
val empty_browser_info = make_browser_info (Symtab.empty, [], [], [], []);  | 
| 7727 | 153  | 
|
| 
23899
 
ab37b1f690c7
clarified init/begin_theory: no longer depend on thy_info.ML;
 
wenzelm 
parents: 
23884 
diff
changeset
 | 
154  | 
fun init_browser_info remote_path curr_sess thys = make_browser_info  | 
| 
24150
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
155  | 
(Symtab.empty, [], [], [], init_graph remote_path curr_sess thys);  | 
| 7727 | 156  | 
|
| 11856 | 157  | 
fun map_browser_info f {theories, files, tex_index, html_index, graph} =
 | 
158  | 
make_browser_info (f (theories, files, tex_index, html_index, graph));  | 
|
| 7727 | 159  | 
|
160  | 
||
161  | 
(* state *)  | 
|
162  | 
||
| 32738 | 163  | 
val browser_info = Unsynchronized.ref empty_browser_info;  | 
164  | 
fun change_browser_info f =  | 
|
165  | 
CRITICAL (fn () => Unsynchronized.change browser_info (map_browser_info f));  | 
|
| 7727 | 166  | 
|
| 32738 | 167  | 
val suppress_tex_source = Unsynchronized.ref false;  | 
| 
42008
 
7423e833a880
Present.init/finish/no_document are not thread-safe -- eliminated futile CRITICAL sections;
 
wenzelm 
parents: 
42007 
diff
changeset
 | 
168  | 
fun no_document f x = Unsynchronized.setmp suppress_tex_source true f x;  | 
| 7727 | 169  | 
|
170  | 
fun init_theory_info name info =  | 
|
| 11856 | 171  | 
change_browser_info (fn (theories, files, tex_index, html_index, graph) =>  | 
| 17412 | 172  | 
(Symtab.update (name, info) theories, files, tex_index, html_index, graph));  | 
| 7727 | 173  | 
|
174  | 
fun change_theory_info name f =  | 
|
| 42010 | 175  | 
change_browser_info (fn (theories, files, tex_index, html_index, graph) =>  | 
| 17412 | 176  | 
(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
 | 
177  | 
      NONE => error ("Browser info: cannot access theory document " ^ quote name)
 | 
| 17412 | 178  | 
| SOME info => (Symtab.update (name, map_theory_info f info) theories, files,  | 
| 9416 | 179  | 
tex_index, html_index, graph)));  | 
| 7727 | 180  | 
|
181  | 
||
| 11856 | 182  | 
fun add_file file =  | 
183  | 
change_browser_info (fn (theories, files, tex_index, html_index, graph) =>  | 
|
184  | 
(theories, file :: files, tex_index, html_index, graph));  | 
|
185  | 
||
| 7727 | 186  | 
fun add_tex_index txt =  | 
| 11856 | 187  | 
change_browser_info (fn (theories, files, tex_index, html_index, graph) =>  | 
| 
24150
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
188  | 
(theories, files, txt :: tex_index, html_index, graph));  | 
| 7727 | 189  | 
|
190  | 
fun add_html_index txt =  | 
|
| 11856 | 191  | 
change_browser_info (fn (theories, files, tex_index, html_index, graph) =>  | 
| 
24150
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
192  | 
(theories, files, tex_index, txt :: html_index, graph));  | 
| 7727 | 193  | 
|
| 
24150
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
194  | 
fun add_graph_entry entry =  | 
| 11856 | 195  | 
change_browser_info (fn (theories, files, tex_index, html_index, graph) =>  | 
| 
24150
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
196  | 
(theories, files, tex_index, html_index, ins_graph_entry entry graph));  | 
| 7727 | 197  | 
|
| 11057 | 198  | 
fun add_tex_source name txt =  | 
199  | 
if ! suppress_tex_source then ()  | 
|
200  | 
else change_theory_info name (fn (tex_source, html_source, html) =>  | 
|
201  | 
(Buffer.add txt tex_source, html_source, html));  | 
|
| 7727 | 202  | 
|
203  | 
fun add_html_source name txt = change_theory_info name (fn (tex_source, html_source, html) =>  | 
|
204  | 
(tex_source, Buffer.add txt html_source, html));  | 
|
205  | 
||
206  | 
||
207  | 
||
208  | 
(** global session state **)  | 
|
209  | 
||
210  | 
(* session_info *)  | 
|
211  | 
||
212  | 
type session_info =  | 
|
213  | 
  {name: string, parent: string, session: string, path: string list, html_prefix: Path.T,
 | 
|
| 17082 | 214  | 
info: bool, doc_format: string, doc_graph: bool, documents: (string * string) list,  | 
| 42007 | 215  | 
dump_prefix: (bool * Path.T) option, remote_path: Url.T option, verbose: bool,  | 
216  | 
readme: Path.T option};  | 
|
| 7727 | 217  | 
|
| 9416 | 218  | 
fun make_session_info  | 
| 17082 | 219  | 
(name, parent, session, path, html_prefix, info, doc_format, doc_graph, documents,  | 
| 42007 | 220  | 
dump_prefix, remote_path, verbose, readme) =  | 
| 7802 | 221  | 
  {name = name, parent = parent, session = session, path = path, html_prefix = html_prefix,
 | 
| 17082 | 222  | 
info = info, doc_format = doc_format, doc_graph = doc_graph, documents = documents,  | 
| 42007 | 223  | 
dump_prefix = dump_prefix, remote_path = remote_path, verbose = verbose,  | 
224  | 
readme = readme}: session_info;  | 
|
| 7685 | 225  | 
|
226  | 
||
| 7727 | 227  | 
(* state *)  | 
228  | 
||
| 32738 | 229  | 
val session_info = Unsynchronized.ref (NONE: session_info option);  | 
| 7727 | 230  | 
|
| 42126 | 231  | 
fun session_default x f = (case ! session_info of NONE => x | SOME info => f info);  | 
| 7727 | 232  | 
|
233  | 
||
234  | 
||
| 14922 | 235  | 
(** document preparation **)  | 
| 7727 | 236  | 
|
237  | 
(* maintain index *)  | 
|
238  | 
||
239  | 
val session_entries =  | 
|
240  | 
HTML.session_entries o  | 
|
| 14898 | 241  | 
map (fn name => (Url.File (Path.append (Path.basic name) index_path), name));  | 
| 7727 | 242  | 
|
243  | 
fun get_entries dir =  | 
|
244  | 
split_lines (File.read (Path.append dir session_entries_path));  | 
|
245  | 
||
246  | 
fun put_entries entries dir =  | 
|
247  | 
File.write (Path.append dir session_entries_path) (cat_lines entries);  | 
|
248  | 
||
249  | 
||
250  | 
fun create_index dir =  | 
|
251  | 
File.read (Path.append dir pre_index_path) ^  | 
|
| 
27862
 
8f727f7c8c1d
removed obsolete verbatim_source, results, chapter, section etc.;
 
wenzelm 
parents: 
27491 
diff
changeset
 | 
252  | 
session_entries (get_entries dir) ^ HTML.end_document  | 
| 7727 | 253  | 
|> File.write (Path.append dir index_path);  | 
254  | 
||
| 
42008
 
7423e833a880
Present.init/finish/no_document are not thread-safe -- eliminated futile CRITICAL sections;
 
wenzelm 
parents: 
42007 
diff
changeset
 | 
255  | 
fun update_index dir name =  | 
| 7727 | 256  | 
(case try get_entries dir of  | 
| 
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
 | 
257  | 
    NONE => warning ("Browser info: cannot access session index of " ^ Path.print dir)
 | 
| 
42008
 
7423e833a880
Present.init/finish/no_document are not thread-safe -- eliminated futile CRITICAL sections;
 
wenzelm 
parents: 
42007 
diff
changeset
 | 
258  | 
| SOME es => (put_entries ((remove (op =) name es) @ [name]) dir; create_index dir));  | 
| 7727 | 259  | 
|
260  | 
||
| 42004 | 261  | 
(* document variants *)  | 
| 17082 | 262  | 
|
| 42004 | 263  | 
fun read_variant str =  | 
| 17082 | 264  | 
(case space_explode "=" str of  | 
265  | 
[name] => (name, "")  | 
|
266  | 
| [name, tags] => (name, tags)  | 
|
| 42004 | 267  | 
  | _ => error ("Malformed document variant specification: " ^ quote str));
 | 
| 17082 | 268  | 
|
| 42004 | 269  | 
fun read_variants strs =  | 
270  | 
rev (distinct (eq_fst (op =)) (rev ((documentN, "") :: map read_variant strs)))  | 
|
| 28375 | 271  | 
|> filter_out (fn (_, s) => s = "-");  | 
| 17082 | 272  | 
|
273  | 
||
| 7727 | 274  | 
(* init session *)  | 
275  | 
||
| 12895 | 276  | 
fun name_of_session elems = space_implode "/" ("Isabelle" :: elems);
 | 
277  | 
||
| 42006 | 278  | 
fun init build info doc doc_graph doc_variants path name dump_prefix  | 
| 
42008
 
7423e833a880
Present.init/finish/no_document are not thread-safe -- eliminated futile CRITICAL sections;
 
wenzelm 
parents: 
42007 
diff
changeset
 | 
279  | 
(remote_path, first_time) verbose thys =  | 
| 42006 | 280  | 
if not build andalso not info andalso doc = "" andalso is_none dump_prefix then  | 
| 15531 | 281  | 
(browser_info := empty_browser_info; session_info := NONE)  | 
| 11856 | 282  | 
else  | 
283  | 
let  | 
|
| 33957 | 284  | 
val parent_name = name_of_session (take (length path - 1) path);  | 
| 11856 | 285  | 
val session_name = name_of_session path;  | 
286  | 
val sess_prefix = Path.make path;  | 
|
287  | 
val html_prefix = Path.append (Path.expand output_path) sess_prefix;  | 
|
| 7727 | 288  | 
|
| 42007 | 289  | 
val documents =  | 
290  | 
if doc = "" then []  | 
|
| 42005 | 291  | 
else if 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
 | 
292  | 
(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
 | 
293  | 
else (); [])  | 
| 42007 | 294  | 
else read_variants doc_variants;  | 
| 7727 | 295  | 
|
| 11856 | 296  | 
val parent_index_path = Path.append Path.parent index_path;  | 
| 42007 | 297  | 
val index_up_lnk =  | 
298  | 
if first_time then  | 
|
| 
16263
 
0609fb8df4a7
removed copy, copy_all (superceded by File.copy, File.copy_dir);
 
wenzelm 
parents: 
15801 
diff
changeset
 | 
299  | 
Url.append (the remote_path) (Url.File (Path.append sess_prefix parent_index_path))  | 
| 14898 | 300  | 
else Url.File parent_index_path;  | 
| 11856 | 301  | 
val readme =  | 
| 15531 | 302  | 
if File.exists readme_html_path then SOME readme_html_path  | 
303  | 
else if File.exists readme_path then SOME readme_path  | 
|
304  | 
else NONE;  | 
|
| 7727 | 305  | 
|
| 17082 | 306  | 
val docs =  | 
307  | 
(case readme of NONE => [] | SOME p => [(Url.File p, "README")]) @  | 
|
| 42007 | 308  | 
map (fn (name, _) => (Url.File (Path.ext doc (Path.basic name)), name)) documents;  | 
| 11856 | 309  | 
val index_text = HTML.begin_index (index_up_lnk, parent_name)  | 
| 
21858
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
20664 
diff
changeset
 | 
310  | 
(Url.File index_path, session_name) docs (Url.explode "medium.html");  | 
| 11856 | 311  | 
in  | 
| 42007 | 312  | 
session_info :=  | 
313  | 
SOME (make_session_info (name, parent_name, session_name, path, html_prefix,  | 
|
314  | 
info, doc, doc_graph, documents, dump_prefix, remote_path, verbose, readme));  | 
|
| 
23899
 
ab37b1f690c7
clarified init/begin_theory: no longer depend on thy_info.ML;
 
wenzelm 
parents: 
23884 
diff
changeset
 | 
315  | 
browser_info := init_browser_info remote_path path thys;  | 
| 
24150
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
316  | 
add_html_index (0, index_text)  | 
| 
42008
 
7423e833a880
Present.init/finish/no_document are not thread-safe -- eliminated futile CRITICAL sections;
 
wenzelm 
parents: 
42007 
diff
changeset
 | 
317  | 
end;  | 
| 7727 | 318  | 
|
319  | 
||
| 
28496
 
4cff10648928
renamed isatool to isabelle_tool in programming interfaces;
 
wenzelm 
parents: 
28375 
diff
changeset
 | 
320  | 
(* isabelle tool wrappers *)  | 
| 17082 | 321  | 
|
| 
28496
 
4cff10648928
renamed isatool to isabelle_tool in programming interfaces;
 
wenzelm 
parents: 
28375 
diff
changeset
 | 
322  | 
fun isabelle_document verbose format name tags path result_path =  | 
| 17082 | 323  | 
let  | 
| 28500 | 324  | 
val s = "\"$ISABELLE_TOOL\" document -c -o '" ^ format ^ "' \  | 
| 
33682
 
0c5d1485dea7
isabelle_document: more explicit error output, notably for drafts;
 
wenzelm 
parents: 
33522 
diff
changeset
 | 
325  | 
\-n '" ^ name ^ "' -t '" ^ tags ^ "' " ^ File.shell_path path ^ " 2>&1";  | 
| 17082 | 326  | 
val doc_path = Path.append result_path (Path.ext format (Path.basic name));  | 
| 
33682
 
0c5d1485dea7
isabelle_document: more explicit error output, notably for drafts;
 
wenzelm 
parents: 
33522 
diff
changeset
 | 
327  | 
val _ = if verbose then writeln s else ();  | 
| 
43850
 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 
wenzelm 
parents: 
43712 
diff
changeset
 | 
328  | 
val (out, rc) = Isabelle_System.bash_output s;  | 
| 
33682
 
0c5d1485dea7
isabelle_document: more explicit error output, notably for drafts;
 
wenzelm 
parents: 
33522 
diff
changeset
 | 
329  | 
val _ =  | 
| 
 
0c5d1485dea7
isabelle_document: more explicit error output, notably for drafts;
 
wenzelm 
parents: 
33522 
diff
changeset
 | 
330  | 
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
 | 
331  | 
        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
 | 
332  | 
else if verbose then writeln out  | 
| 
 
0c5d1485dea7
isabelle_document: more explicit error output, notably for drafts;
 
wenzelm 
parents: 
33522 
diff
changeset
 | 
333  | 
else ();  | 
| 
 
0c5d1485dea7
isabelle_document: more explicit error output, notably for drafts;
 
wenzelm 
parents: 
33522 
diff
changeset
 | 
334  | 
in doc_path end;  | 
| 17082 | 335  | 
|
| 
42127
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
42126 
diff
changeset
 | 
336  | 
fun isabelle_browser graph = Isabelle_System.with_tmp_dir "browser" (fn dir =>  | 
| 17082 | 337  | 
let  | 
| 
42127
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
42126 
diff
changeset
 | 
338  | 
val pdf_path = Path.append dir graph_pdf_path;  | 
| 
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
42126 
diff
changeset
 | 
339  | 
val eps_path = Path.append dir graph_eps_path;  | 
| 
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
42126 
diff
changeset
 | 
340  | 
val graph_path = Path.append dir graph_path;  | 
| 
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
42126 
diff
changeset
 | 
341  | 
val _ = write_graph graph graph_path;  | 
| 
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
42126 
diff
changeset
 | 
342  | 
val args = "-o " ^ File.shell_path pdf_path ^ " " ^ File.shell_path graph_path;  | 
| 17082 | 343  | 
in  | 
| 
42127
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
42126 
diff
changeset
 | 
344  | 
if Isabelle_System.isabelle_tool "browser" args = 0 andalso  | 
| 
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
42126 
diff
changeset
 | 
345  | 
File.exists pdf_path andalso File.exists eps_path  | 
| 
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
42126 
diff
changeset
 | 
346  | 
then (File.read pdf_path, File.read eps_path)  | 
| 
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
42126 
diff
changeset
 | 
347  | 
else error "Failed to prepare dependency graph"  | 
| 
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
42126 
diff
changeset
 | 
348  | 
end);  | 
| 17082 | 349  | 
|
350  | 
||
| 11856 | 351  | 
(* finish session -- output all generated text *)  | 
352  | 
||
| 
24150
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
353  | 
fun sorted_index index = map snd (sort (int_ord o pairself fst) (rev index));  | 
| 
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
354  | 
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
 | 
355  | 
|
| 14922 | 356  | 
fun write_tex src name path =  | 
| 28027 | 357  | 
File.write_buffer (Path.append path (tex_path name)) src;  | 
| 14922 | 358  | 
|
359  | 
fun write_tex_index tex_index path =  | 
|
| 
24150
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
360  | 
write_tex (index_buffer tex_index |> Buffer.add Latex.tex_trailer) doc_indexN path;  | 
| 14922 | 361  | 
|
| 7685 | 362  | 
|
| 
42008
 
7423e833a880
Present.init/finish/no_document are not thread-safe -- eliminated futile CRITICAL sections;
 
wenzelm 
parents: 
42007 
diff
changeset
 | 
363  | 
fun finish () =  | 
| 42126 | 364  | 
  session_default () (fn {name, info, html_prefix, doc_format,
 | 
| 
42008
 
7423e833a880
Present.init/finish/no_document are not thread-safe -- eliminated futile CRITICAL sections;
 
wenzelm 
parents: 
42007 
diff
changeset
 | 
365  | 
doc_graph, documents, dump_prefix, path, verbose, readme, ...} =>  | 
| 7727 | 366  | 
let  | 
| 11856 | 367  | 
    val {theories, files, tex_index, html_index, graph} = ! browser_info;
 | 
368  | 
val thys = Symtab.dest theories;  | 
|
| 9416 | 369  | 
val parent_html_prefix = Path.append html_prefix Path.parent;  | 
| 7727 | 370  | 
|
| 11856 | 371  | 
    fun finish_tex path (a, {tex_source, ...}: theory_info) = write_tex tex_source a path;
 | 
372  | 
    fun finish_html (a, {html, ...}: theory_info) =
 | 
|
| 28027 | 373  | 
File.write_buffer (Path.append html_prefix (html_path a)) (Buffer.add HTML.end_document html);  | 
| 11856 | 374  | 
|
| 
24150
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
375  | 
val sorted_graph = sorted_index graph;  | 
| 11856 | 376  | 
val opt_graphs =  | 
| 42007 | 377  | 
if doc_graph andalso (not (null documents) orelse is_some dump_prefix) then  | 
| 
28496
 
4cff10648928
renamed isatool to isabelle_tool in programming interfaces;
 
wenzelm 
parents: 
28375 
diff
changeset
 | 
378  | 
SOME (isabelle_browser sorted_graph)  | 
| 15531 | 379  | 
else NONE;  | 
| 11856 | 380  | 
|
| 17210 | 381  | 
fun prepare_sources cp path =  | 
| 40743 | 382  | 
(Isabelle_System.mkdirs path;  | 
383  | 
if cp then Isabelle_System.copy_dir document_path path else ();  | 
|
384  | 
Isabelle_System.isabelle_tool "latex"  | 
|
| 31819 | 385  | 
        ("-o sty " ^ File.shell_path (Path.append path (Path.basic "root.tex")));
 | 
| 15531 | 386  | 
(case opt_graphs of NONE => () | SOME (pdf, eps) =>  | 
| 11856 | 387  | 
(File.write (Path.append path graph_pdf_path) pdf;  | 
388  | 
File.write (Path.append path graph_eps_path) eps));  | 
|
| 14922 | 389  | 
write_tex_index tex_index path;  | 
| 15570 | 390  | 
List.app (finish_tex path) thys);  | 
| 42007 | 391  | 
val _ =  | 
392  | 
if info then  | 
|
393  | 
(Isabelle_System.mkdirs (Path.append html_prefix session_path);  | 
|
394  | 
File.write_buffer (Path.append html_prefix pre_index_path) (index_buffer html_index);  | 
|
395  | 
File.write (Path.append html_prefix session_entries_path) "";  | 
|
396  | 
create_index html_prefix;  | 
|
397  | 
if length path > 1 then update_index parent_html_prefix name else ();  | 
|
398  | 
(case readme of NONE => () | SOME path => File.copy path html_prefix);  | 
|
399  | 
write_graph sorted_graph (Path.append html_prefix graph_path);  | 
|
400  | 
Isabelle_System.isabelle_tool "browser" "-b";  | 
|
401  | 
File.copy (Path.explode "~~/lib/browser/GraphBrowser.jar") html_prefix;  | 
|
402  | 
List.app (fn (a, txt) => File.write (Path.append html_prefix (Path.basic a)) txt)  | 
|
403  | 
(HTML.applet_pages name (Url.File index_path, name));  | 
|
| 43437 | 404  | 
File.copy (Path.explode "~~/etc/isabelle.css") html_prefix;  | 
| 42007 | 405  | 
List.app finish_html thys;  | 
406  | 
List.app (uncurry File.write) files;  | 
|
| 
44389
 
a3b5fdfb04a3
tuned signature -- contrast physical output primitives versus Output.raw_message;
 
wenzelm 
parents: 
43850 
diff
changeset
 | 
407  | 
        if verbose then Output.physical_stderr ("Browser info at " ^ show_path html_prefix ^ "\n")
 | 
| 42007 | 408  | 
else ())  | 
409  | 
else ();  | 
|
410  | 
||
411  | 
val _ =  | 
|
412  | 
(case dump_prefix of NONE => () | SOME (cp, path) =>  | 
|
413  | 
(prepare_sources cp path;  | 
|
| 
44389
 
a3b5fdfb04a3
tuned signature -- contrast physical output primitives versus Output.raw_message;
 
wenzelm 
parents: 
43850 
diff
changeset
 | 
414  | 
        if verbose then Output.physical_stderr ("Document sources at " ^ show_path path ^ "\n")
 | 
| 42007 | 415  | 
else ()));  | 
416  | 
||
417  | 
val doc_paths =  | 
|
| 
42009
 
b008525c4399
parallel preparation of document variants, within separate directories;
 
wenzelm 
parents: 
42008 
diff
changeset
 | 
418  | 
documents |> Par_List.map (fn (name, tags) =>  | 
| 42007 | 419  | 
let  | 
| 
42009
 
b008525c4399
parallel preparation of document variants, within separate directories;
 
wenzelm 
parents: 
42008 
diff
changeset
 | 
420  | 
val path = Path.append html_prefix (Path.basic name);  | 
| 42007 | 421  | 
val _ = prepare_sources true path;  | 
422  | 
in isabelle_document true doc_format name tags path html_prefix end);  | 
|
423  | 
val _ =  | 
|
424  | 
if verbose then  | 
|
| 
44389
 
a3b5fdfb04a3
tuned signature -- contrast physical output primitives versus Output.raw_message;
 
wenzelm 
parents: 
43850 
diff
changeset
 | 
425  | 
doc_paths  | 
| 
 
a3b5fdfb04a3
tuned signature -- contrast physical output primitives versus Output.raw_message;
 
wenzelm 
parents: 
43850 
diff
changeset
 | 
426  | 
        |> List.app (fn doc => Output.physical_stderr ("Document at " ^ show_path doc ^ "\n"))
 | 
| 42007 | 427  | 
else ();  | 
| 7727 | 428  | 
in  | 
429  | 
browser_info := empty_browser_info;  | 
|
| 15531 | 430  | 
session_info := NONE  | 
| 
42008
 
7423e833a880
Present.init/finish/no_document are not thread-safe -- eliminated futile CRITICAL sections;
 
wenzelm 
parents: 
42007 
diff
changeset
 | 
431  | 
end);  | 
| 7727 | 432  | 
|
433  | 
||
434  | 
(* theory elements *)  | 
|
435  | 
||
| 42126 | 436  | 
fun init_theory name = session_default () (fn _ => init_theory_info name empty_theory_info);  | 
| 7727 | 437  | 
|
| 
27862
 
8f727f7c8c1d
removed obsolete verbatim_source, results, chapter, section etc.;
 
wenzelm 
parents: 
27491 
diff
changeset
 | 
438  | 
fun theory_source name mk_text =  | 
| 42126 | 439  | 
session_default () (fn _ => add_html_source name (HTML.theory_source (mk_text ())));  | 
| 7727 | 440  | 
|
| 9136 | 441  | 
fun theory_output name s =  | 
| 42126 | 442  | 
session_default () (fn _ => add_tex_source name (Latex.isabelle_file name s));  | 
| 7727 | 443  | 
|
444  | 
||
| 
23899
 
ab37b1f690c7
clarified init/begin_theory: no longer depend on thy_info.ML;
 
wenzelm 
parents: 
23884 
diff
changeset
 | 
445  | 
fun parent_link remote_path curr_session thy =  | 
| 
 
ab37b1f690c7
clarified init/begin_theory: no longer depend on thy_info.ML;
 
wenzelm 
parents: 
23884 
diff
changeset
 | 
446  | 
let  | 
| 
 
ab37b1f690c7
clarified init/begin_theory: no longer depend on thy_info.ML;
 
wenzelm 
parents: 
23884 
diff
changeset
 | 
447  | 
    val {name = _, session, is_local} = get_info thy;
 | 
| 
 
ab37b1f690c7
clarified init/begin_theory: no longer depend on thy_info.ML;
 
wenzelm 
parents: 
23884 
diff
changeset
 | 
448  | 
val name = Context.theory_name thy;  | 
| 
 
ab37b1f690c7
clarified init/begin_theory: no longer depend on thy_info.ML;
 
wenzelm 
parents: 
23884 
diff
changeset
 | 
449  | 
val link =  | 
| 
 
ab37b1f690c7
clarified init/begin_theory: no longer depend on thy_info.ML;
 
wenzelm 
parents: 
23884 
diff
changeset
 | 
450  | 
if null session then NONE  | 
| 
 
ab37b1f690c7
clarified init/begin_theory: no longer depend on thy_info.ML;
 
wenzelm 
parents: 
23884 
diff
changeset
 | 
451  | 
else SOME  | 
| 
 
ab37b1f690c7
clarified init/begin_theory: no longer depend on thy_info.ML;
 
wenzelm 
parents: 
23884 
diff
changeset
 | 
452  | 
(if is_some remote_path andalso not is_local then  | 
| 
 
ab37b1f690c7
clarified init/begin_theory: no longer depend on thy_info.ML;
 
wenzelm 
parents: 
23884 
diff
changeset
 | 
453  | 
Url.append (the remote_path) (Url.File (Path.append (Path.make session) (html_path name)))  | 
| 
 
ab37b1f690c7
clarified init/begin_theory: no longer depend on thy_info.ML;
 
wenzelm 
parents: 
23884 
diff
changeset
 | 
454  | 
else Url.File (Path.append (mk_rel_path curr_session session) (html_path name)));  | 
| 
 
ab37b1f690c7
clarified init/begin_theory: no longer depend on thy_info.ML;
 
wenzelm 
parents: 
23884 
diff
changeset
 | 
455  | 
in (link, name) end;  | 
| 7727 | 456  | 
|
| 
26323
 
73efc70edeef
theory loader: discontinued *attached* ML scripts;
 
wenzelm 
parents: 
24829 
diff
changeset
 | 
457  | 
fun begin_theory update_time dir files thy =  | 
| 42126 | 458  | 
    session_default thy (fn {name = sess_name, session, path, html_prefix, remote_path, ...} =>
 | 
| 7727 | 459  | 
let  | 
| 
23899
 
ab37b1f690c7
clarified init/begin_theory: no longer depend on thy_info.ML;
 
wenzelm 
parents: 
23884 
diff
changeset
 | 
460  | 
val name = Context.theory_name thy;  | 
| 
 
ab37b1f690c7
clarified init/begin_theory: no longer depend on thy_info.ML;
 
wenzelm 
parents: 
23884 
diff
changeset
 | 
461  | 
val parents = Theory.parents_of thy;  | 
| 
 
ab37b1f690c7
clarified init/begin_theory: no longer depend on thy_info.ML;
 
wenzelm 
parents: 
23884 
diff
changeset
 | 
462  | 
val parent_specs = map (parent_link remote_path path) parents;  | 
| 7727 | 463  | 
|
| 
37939
 
965537d86fcc
discontinued special treatment of ML files -- no longer complete extensions on demand;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
464  | 
val files_html = files |> map (fn (raw_path, loadit) =>  | 
| 
 
965537d86fcc
discontinued special treatment of ML files -- no longer complete extensions on demand;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
465  | 
let  | 
| 
43712
 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 
wenzelm 
parents: 
43437 
diff
changeset
 | 
466  | 
val path = File.check_file (File.full_path dir raw_path);  | 
| 
37939
 
965537d86fcc
discontinued special treatment of ML files -- no longer complete extensions on demand;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
467  | 
val base = Path.base path;  | 
| 
 
965537d86fcc
discontinued special treatment of ML files -- no longer complete extensions on demand;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
468  | 
val base_html = html_ext base;  | 
| 
 
965537d86fcc
discontinued special treatment of ML files -- no longer complete extensions on demand;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
469  | 
val _ = add_file (Path.append html_prefix base_html,  | 
| 
37940
 
4857eab31298
generic external source files -- nothing special about ML here;
 
wenzelm 
parents: 
37939 
diff
changeset
 | 
470  | 
HTML.external_file (Url.File base) (File.read path));  | 
| 
37939
 
965537d86fcc
discontinued special treatment of ML files -- no longer complete extensions on demand;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
471  | 
in (Url.File base_html, Url.File raw_path, loadit) end);  | 
| 7727 | 472  | 
|
473  | 
fun prep_html_source (tex_source, html_source, html) =  | 
|
474  | 
let  | 
|
| 14898 | 475  | 
val txt = HTML.begin_theory (Url.File index_path, session)  | 
| 
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
 | 
476  | 
name parent_specs files_html (Buffer.content html_source)  | 
| 7727 | 477  | 
in (tex_source, Buffer.empty, Buffer.add txt html) end;  | 
478  | 
||
| 9416 | 479  | 
val entry =  | 
480  | 
     {name = name, ID = ID_of path name, dir = sess_name, unfold = true,
 | 
|
| 
21858
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
20664 
diff
changeset
 | 
481  | 
path = Path.implode (html_path name),  | 
| 
23899
 
ab37b1f690c7
clarified init/begin_theory: no longer depend on thy_info.ML;
 
wenzelm 
parents: 
23884 
diff
changeset
 | 
482  | 
parents = map ID_of_thy parents};  | 
| 7727 | 483  | 
in  | 
484  | 
change_theory_info name prep_html_source;  | 
|
| 
24150
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
485  | 
add_graph_entry (update_time, entry);  | 
| 
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
486  | 
add_html_index (update_time, HTML.theory_entry (Url.File (html_path name), name));  | 
| 
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
487  | 
add_tex_index (update_time, Latex.theory_entry name);  | 
| 27327 | 488  | 
    put_info {name = sess_name, session = path, is_local = is_some remote_path} thy
 | 
| 7727 | 489  | 
end);  | 
490  | 
||
491  | 
||
492  | 
||
| 
42127
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
42126 
diff
changeset
 | 
493  | 
(** draft document output **)  | 
| 14922 | 494  | 
|
| 
42127
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
42126 
diff
changeset
 | 
495  | 
fun drafts doc_format src_paths = Isabelle_System.with_tmp_dir "drafts" (fn dir =>  | 
| 14922 | 496  | 
let  | 
| 
24150
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
497  | 
fun prep_draft path i =  | 
| 14935 | 498  | 
let  | 
499  | 
val base = Path.base path;  | 
|
| 14972 | 500  | 
val name =  | 
| 24829 | 501  | 
(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
 | 
502  | 
"" => "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
 | 
503  | 
| s => s) ^ serial_string ();  | 
| 14935 | 504  | 
in  | 
505  | 
if File.exists path then  | 
|
| 
24150
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
506  | 
(((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
 | 
507  | 
        else error ("Bad file: " ^ Path.print path)
 | 
| 14935 | 508  | 
end;  | 
| 
24150
 
ed724867099a
sort indexes according to symbolic update_time (multithreading-safe);
 
wenzelm 
parents: 
24102 
diff
changeset
 | 
509  | 
val (srcs, tex_index) = split_list (fst (fold_map prep_draft src_paths 0));  | 
| 14935 | 510  | 
|
| 
42127
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
42126 
diff
changeset
 | 
511  | 
val doc_path = Path.append dir document_path;  | 
| 40743 | 512  | 
val _ = Isabelle_System.mkdirs doc_path;  | 
| 14922 | 513  | 
val root_path = Path.append doc_path (Path.basic "root.tex");  | 
| 
21858
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
20664 
diff
changeset
 | 
514  | 
val _ = File.copy (Path.explode "~~/lib/texinputs/draft.tex") root_path;  | 
| 40743 | 515  | 
    val _ = Isabelle_System.isabelle_tool "latex" ("-o sty " ^ File.shell_path root_path);
 | 
516  | 
    val _ = Isabelle_System.isabelle_tool "latex" ("-o syms " ^ File.shell_path root_path);
 | 
|
| 14922 | 517  | 
|
518  | 
fun known name =  | 
|
519  | 
let val ss = split_lines (File.read (Path.append doc_path (Path.basic name)))  | 
|
| 20664 | 520  | 
in member (op =) ss end;  | 
| 14922 | 521  | 
val known_syms = known "syms.lst";  | 
522  | 
val known_ctrls = known "ctrls.lst";  | 
|
523  | 
||
| 15570 | 524  | 
val _ = srcs |> List.app (fn (name, base, txt) =>  | 
| 14935 | 525  | 
Symbol.explode txt  | 
| 
21858
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
20664 
diff
changeset
 | 
526  | 
|> Latex.symbol_source (known_syms, known_ctrls) (Path.implode base)  | 
| 14935 | 527  | 
|> File.write (Path.append doc_path (tex_path name)));  | 
| 14922 | 528  | 
val _ = write_tex_index tex_index doc_path;  | 
529  | 
||
| 
42127
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
42126 
diff
changeset
 | 
530  | 
val result = isabelle_document false doc_format documentN "" doc_path dir;  | 
| 
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
42126 
diff
changeset
 | 
531  | 
val result' = Isabelle_System.create_tmp_path documentN doc_format;  | 
| 
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
42126 
diff
changeset
 | 
532  | 
val _ = File.copy result result';  | 
| 
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
42126 
diff
changeset
 | 
533  | 
in result' end);  | 
| 14922 | 534  | 
|
| 7685 | 535  | 
end;  | 
536  | 
||
| 32738 | 537  | 
structure Basic_Present: BASIC_PRESENT = Present;  | 
538  | 
open Basic_Present;  |