| author | wenzelm | 
| Thu, 01 Jul 1999 21:19:45 +0200 | |
| changeset 6874 | 747f656e04ec | 
| parent 6665 | bf421d724db7 | 
| child 7236 | e077484d50d8 | 
| permissions | -rw-r--r-- | 
| 6330 | 1  | 
(* Title: Pure/Thy/browser_info.ML  | 
| 
3603
 
5eef2ff0eb72
This file now contains all functions for generating html
 
berghofe 
parents:  
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 6330 | 3  | 
Author: Stefan Berghofer and Markus Wenzel, TU Muenchen  | 
4  | 
||
5  | 
Theory browsing information (HTML and graph files).  | 
|
6  | 
*)  | 
|
| 
3603
 
5eef2ff0eb72
This file now contains all functions for generating html
 
berghofe 
parents:  
diff
changeset
 | 
7  | 
|
| 6330 | 8  | 
signature BASIC_BROWSER_INFO =  | 
9  | 
sig  | 
|
10  | 
val section: string -> unit  | 
|
11  | 
end;  | 
|
| 3870 | 12  | 
|
| 6330 | 13  | 
signature BROWSER_INFO =  | 
14  | 
sig  | 
|
15  | 
include BASIC_BROWSER_INFO  | 
|
| 6649 | 16  | 
val init: bool -> string list -> string -> Url.T option -> unit  | 
| 6330 | 17  | 
val finish: unit -> unit  | 
18  | 
val theory_source: string -> (string, 'a) Source.source -> unit  | 
|
| 6649 | 19  | 
val begin_theory: string -> string list -> (Path.T * bool) list -> theory -> theory  | 
| 6330 | 20  | 
val end_theory: string -> unit  | 
21  | 
val theorem: string -> thm -> unit  | 
|
| 6649 | 22  | 
val setup: (theory -> theory) list  | 
| 6330 | 23  | 
end;  | 
24  | 
||
25  | 
structure BrowserInfo: BROWSER_INFO =  | 
|
26  | 
struct  | 
|
| 
3603
 
5eef2ff0eb72
This file now contains all functions for generating html
 
berghofe 
parents:  
diff
changeset
 | 
27  | 
|
| 6665 | 28  | 
|
| 6649 | 29  | 
(** paths **)  | 
| 6330 | 30  | 
|
31  | 
val output_path = Path.variable "ISABELLE_BROWSER_INFO";  | 
|
32  | 
||
| 6649 | 33  | 
fun top_path sess_path offset = Path.append (Path.appends (replicate (offset + length sess_path) Path.parent));  | 
| 6330 | 34  | 
|
35  | 
val html_ext = Path.ext "html";  | 
|
36  | 
val html_path = html_ext o Path.basic;  | 
|
37  | 
val graph_path = Path.ext "graph" o Path.basic;  | 
|
38  | 
val index_path = Path.basic "index.html";  | 
|
39  | 
||
| 6339 | 40  | 
val session_path = Path.basic ".session";  | 
41  | 
val session_entries_path = Path.unpack ".session/entries";  | 
|
42  | 
val pre_index_path = Path.unpack ".session/pre-index";  | 
|
43  | 
||
| 6330 | 44  | 
|
| 6665 | 45  | 
|
| 6649 | 46  | 
(** additional theory data **)  | 
47  | 
||
48  | 
structure BrowserInfoArgs =  | 
|
49  | 
struct  | 
|
50  | 
val name = "Pure/browser_info";  | 
|
51  | 
  type T = {session: string list, is_local: bool};
 | 
|
52  | 
||
53  | 
  val empty = {session = [], is_local = false};
 | 
|
54  | 
val copy = I;  | 
|
55  | 
val prep_ext = I;  | 
|
56  | 
fun merge _ = empty;  | 
|
57  | 
fun print _ _ = ();  | 
|
58  | 
end;  | 
|
59  | 
||
60  | 
structure BrowserInfoData = TheoryDataFun(BrowserInfoArgs);  | 
|
| 6665 | 61  | 
|
62  | 
fun get_info thy =  | 
|
63  | 
if Theory.eq_thy (thy, ProtoPure.thy) then BrowserInfoArgs.empty  | 
|
64  | 
else BrowserInfoData.get thy;  | 
|
65  | 
||
66  | 
||
| 6649 | 67  | 
|
68  | 
(** graphs **)  | 
|
69  | 
||
70  | 
type graph_node =  | 
|
71  | 
  {name: string, ID: string, dir: string, unfold: bool,
 | 
|
72  | 
path: string, parents: string list};  | 
|
73  | 
||
74  | 
fun get_graph path = map (fn (a :: b :: c :: d :: e :: r) =>  | 
|
75  | 
  {name = unenclose a, ID = unenclose b,
 | 
|
76  | 
dir = unenclose c, unfold = (d = "+"),  | 
|
77  | 
path = unenclose (if d = "+" then e else d),  | 
|
78  | 
parents = map unenclose (fst (split_last (if d = "+" then tl r else r)))})  | 
|
79  | 
(map (filter_out (equal "") o space_explode " ") (split_lines (File.read path)));  | 
|
80  | 
||
81  | 
fun put_graph gr path = File.write path (cat_lines (map (fn {name, ID, dir, unfold, path, parents} =>
 | 
|
82  | 
"\"" ^ name ^ "\" \"" ^ ID ^ "\" \"" ^ dir ^ (if unfold then "\" + \"" else "\" \"") ^  | 
|
83  | 
path ^ "\" > " ^ space_implode " " (map quote parents) ^ " ;") gr));  | 
|
84  | 
||
85  | 
fun dir_of sess = space_implode "/" (case sess of  | 
|
86  | 
[] => ["Pure"] | [x] => [x, "base"] | xs => xs);  | 
|
87  | 
||
88  | 
fun ID_of sess s = dir_of sess ^ "/" ^ s;  | 
|
89  | 
||
90  | 
(* retrieve graph data from theory loader database *)  | 
|
91  | 
||
92  | 
fun make_graph remote_path unfold curr_sess = map (fn name =>  | 
|
93  | 
let  | 
|
94  | 
    val {session, is_local} = get_info (ThyInfo.theory name);
 | 
|
95  | 
val path' = Path.append (Path.make session) (html_path name)  | 
|
96  | 
in  | 
|
97  | 
    {name = name, ID = ID_of session name, dir = dir_of session,
 | 
|
98  | 
path = if null session then "" else  | 
|
99  | 
if is_some remote_path andalso not is_local then  | 
|
100  | 
Url.pack (Url.append (the remote_path) (Url.file path'))  | 
|
101  | 
else Path.pack (top_path curr_sess 2 path'),  | 
|
102  | 
unfold = unfold andalso (length session = 1),  | 
|
103  | 
parents = map (fn s => ID_of (#session (get_info (ThyInfo.theory s))) s)  | 
|
104  | 
(ThyInfo.get_preds name)}  | 
|
105  | 
end) (ThyInfo.names ());  | 
|
106  | 
||
107  | 
fun add_graph_entry' (entry as {ID, ...}) (gr : graph_node list) =
 | 
|
108  | 
filter_out (fn entry' => #ID entry' = ID) gr @ [entry];  | 
|
109  | 
||
110  | 
||
| 6665 | 111  | 
|
| 6649 | 112  | 
(** global browser info state **)  | 
113  | 
||
114  | 
(* type theory_info *)  | 
|
115  | 
||
116  | 
type theory_info = {source: Buffer.T, html: Buffer.T};
 | 
|
117  | 
||
118  | 
fun make_theory_info (source, html) =  | 
|
119  | 
  {source = source, html = html}: theory_info;
 | 
|
120  | 
||
121  | 
val empty_theory_info = make_theory_info (Buffer.empty, Buffer.empty);  | 
|
122  | 
fun map_theory_info f {source, html} = make_theory_info (f (source, html));
 | 
|
123  | 
||
124  | 
||
125  | 
(* type browser_info *)  | 
|
126  | 
||
127  | 
type browser_info = {theories: theory_info Symtab.table, index: Buffer.T,
 | 
|
128  | 
graph: graph_node list, all_graph: graph_node list};  | 
|
129  | 
||
130  | 
fun make_browser_info (theories, index, graph, all_graph) =  | 
|
131  | 
  {theories = theories, index = index, graph = graph, all_graph = all_graph}: browser_info;
 | 
|
132  | 
||
133  | 
val empty_browser_info = make_browser_info (Symtab.empty, Buffer.empty, [], []);  | 
|
134  | 
||
135  | 
fun initial_browser_info remote_path graph_path curr_sess = make_browser_info  | 
|
136  | 
(Symtab.empty, Buffer.empty, make_graph remote_path false curr_sess,  | 
|
137  | 
if_none (try get_graph graph_path) (make_graph remote_path true [""]));  | 
|
138  | 
||
139  | 
fun map_browser_info f {theories, index, graph, all_graph} =
 | 
|
140  | 
make_browser_info (f (theories, index, graph, all_graph));  | 
|
141  | 
||
142  | 
||
143  | 
(* state *)  | 
|
144  | 
||
145  | 
val browser_info = ref empty_browser_info;  | 
|
146  | 
||
147  | 
fun change_browser_info f = browser_info := map_browser_info f (! browser_info);  | 
|
148  | 
||
149  | 
fun init_theory_info name info = change_browser_info (fn (theories, index, graph, all_graph) =>  | 
|
150  | 
(Symtab.update ((name, info), theories), index, graph, all_graph));  | 
|
151  | 
||
152  | 
fun change_theory_info name f = change_browser_info (fn (theories, index, graph, all_graph) =>  | 
|
153  | 
(case Symtab.lookup (theories, name) of  | 
|
154  | 
    None => (warning ("Browser info: cannot access theory document " ^ quote name);
 | 
|
155  | 
(theories, index, graph, all_graph))  | 
|
156  | 
| Some info => (Symtab.update ((name, map_theory_info f info), theories), index, graph, all_graph)));  | 
|
157  | 
||
158  | 
||
159  | 
fun add_index txt = change_browser_info (fn (theories, index, graph, all_graph) =>  | 
|
160  | 
(theories, Buffer.add txt index, graph, all_graph));  | 
|
161  | 
||
| 6665 | 162  | 
(*add entry to graph for current session*)  | 
| 6649 | 163  | 
fun add_graph_entry e = change_browser_info (fn (theories, index, graph, all_graph) =>  | 
164  | 
(theories, index, add_graph_entry' e graph, all_graph));  | 
|
165  | 
||
| 6665 | 166  | 
(*add entry to graph for all sessions*)  | 
| 6649 | 167  | 
fun add_all_graph_entry e = change_browser_info (fn (theories, index, graph, all_graph) =>  | 
168  | 
(theories, index, graph, add_graph_entry' e all_graph));  | 
|
169  | 
||
170  | 
fun add_source name txt = change_theory_info name (fn (source, html) =>  | 
|
171  | 
(Buffer.add txt source, html));  | 
|
172  | 
||
173  | 
fun add_html name txt = change_theory_info name (fn (source, html) =>  | 
|
174  | 
(source, Buffer.add txt html));  | 
|
175  | 
||
176  | 
||
| 6665 | 177  | 
|
| 6649 | 178  | 
(** global session state **)  | 
179  | 
||
| 6330 | 180  | 
(* session_info *)  | 
181  | 
||
182  | 
type session_info =  | 
|
| 6342 | 183  | 
  {name: string, parent: string, session: string, path: string list,
 | 
| 6649 | 184  | 
html_prefix: Path.T, graph_path: Path.T, all_graph_path: Path.T, remote_path: Url.T option};  | 
| 6330 | 185  | 
|
| 6649 | 186  | 
fun make_session_info (name, parent, session, path, html_prefix, graph_path, all_graph_path, remote_path) =  | 
| 6342 | 187  | 
  {name = name, parent = parent, session = session, path = path,
 | 
| 6649 | 188  | 
html_prefix = html_prefix, graph_path = graph_path,  | 
189  | 
all_graph_path = all_graph_path, remote_path = remote_path}: session_info;  | 
|
| 6330 | 190  | 
|
| 6348 | 191  | 
|
192  | 
(* state *)  | 
|
193  | 
||
| 6330 | 194  | 
val session_info = ref (None: session_info option);  | 
195  | 
||
| 6649 | 196  | 
fun with_session x f = (case ! session_info of None => x | Some info => f info);  | 
| 6348 | 197  | 
fun with_context f = f (PureThy.get_name (Context.the_context ()));  | 
198  | 
||
| 6330 | 199  | 
|
| 6665 | 200  | 
|
| 6339 | 201  | 
(** HTML output **)  | 
202  | 
||
203  | 
(* maintain index *)  | 
|
204  | 
||
205  | 
val session_entries =  | 
|
| 6649 | 206  | 
HTML.session_entries o map (fn name => (Url.file (Path.append (Path.basic name) index_path), name));  | 
| 6339 | 207  | 
|
208  | 
fun get_entries dir =  | 
|
209  | 
split_lines (File.read (Path.append dir session_entries_path));  | 
|
210  | 
||
211  | 
fun put_entries entries dir =  | 
|
212  | 
File.write (Path.append dir session_entries_path) (cat_lines entries);  | 
|
213  | 
||
214  | 
||
215  | 
fun create_index dir =  | 
|
216  | 
File.read (Path.append dir pre_index_path) ^  | 
|
217  | 
session_entries (get_entries dir) ^ HTML.end_index  | 
|
218  | 
|> File.write (Path.append dir index_path);  | 
|
219  | 
||
220  | 
fun update_index dir name =  | 
|
221  | 
(case try get_entries dir of  | 
|
| 6388 | 222  | 
    None => warning ("Browser info: cannot access session index of " ^ quote (Path.pack dir))
 | 
| 6339 | 223  | 
| Some es => (put_entries ((es \ name) @ [name]) dir; create_index dir));  | 
| 6342 | 224  | 
|
| 6339 | 225  | 
|
| 6330 | 226  | 
(* init session *)  | 
227  | 
||
| 6342 | 228  | 
fun name_of_session elems = space_implode "/" ("Isabelle" :: elems);
 | 
| 6339 | 229  | 
|
| 6348 | 230  | 
fun could_copy inpath outpath =  | 
231  | 
if File.exists inpath then (File.copy inpath outpath; true)  | 
|
232  | 
else false;  | 
|
233  | 
||
| 6649 | 234  | 
fun init false _ _ _ = (browser_info := empty_browser_info; session_info := None)  | 
235  | 
| init true path name remote_path =  | 
|
| 6330 | 236  | 
let  | 
| 6348 | 237  | 
val parent_name = name_of_session (take (length path - 1, path));  | 
| 6342 | 238  | 
val session_name = name_of_session path;  | 
| 6339 | 239  | 
val sess_prefix = Path.make path;  | 
| 6330 | 240  | 
val out_path = Path.expand output_path;  | 
241  | 
val html_prefix = Path.append out_path sess_prefix;  | 
|
242  | 
val graph_prefix = Path.appends [out_path, Path.unpack "graph/data", sess_prefix];  | 
|
| 6649 | 243  | 
val graph_path = Path.append graph_prefix (Path.basic "session.graph");  | 
244  | 
val graph_lnk = Url.file (top_path path 0 (Path.appends  | 
|
245  | 
[Path.unpack "graph/data", sess_prefix, Path.basic "medium.html"]));  | 
|
246  | 
val graph_up_lnk = (Url.file (top_path path 2 (Path.append sess_prefix index_path)), session_name);  | 
|
247  | 
val codebase = Url.file (top_path path 1 Path.current);  | 
|
248  | 
val all_graph_path = Path.appends [out_path, Path.unpack "graph/data",  | 
|
249  | 
Path.basic (hd path), Path.basic "all_sessions.graph"];  | 
|
| 6348 | 250  | 
|
251  | 
val _ = (File.mkdir html_prefix; File.mkdir graph_prefix);  | 
|
252  | 
fun prep_readme readme =  | 
|
253  | 
let val readme_path = Path.basic readme in  | 
|
254  | 
if could_copy readme_path (Path.append html_prefix readme_path) then Some readme_path  | 
|
255  | 
else None  | 
|
256  | 
end;  | 
|
257  | 
val opt_readme =  | 
|
258  | 
(case prep_readme "README.html" of None => prep_readme "README" | some => some);  | 
|
| 6649 | 259  | 
val index_text = HTML.begin_index (Url.file (Path.append Path.parent index_path), parent_name)  | 
260  | 
(Url.file index_path, session_name) (apsome Url.file opt_readme) graph_lnk;  | 
|
| 6330 | 261  | 
in  | 
| 6339 | 262  | 
File.mkdir (Path.append html_prefix session_path);  | 
263  | 
File.write (Path.append html_prefix session_entries_path) "";  | 
|
| 6649 | 264  | 
seq (fn (name, txt) => File.write (Path.append graph_prefix (Path.basic name)) txt)  | 
265  | 
(HTML.applet_pages session_name codebase graph_up_lnk (length path = 1));  | 
|
| 6342 | 266  | 
session_info := Some (make_session_info (name, parent_name, session_name, path,  | 
| 6649 | 267  | 
html_prefix, graph_path, all_graph_path, remote_path));  | 
268  | 
browser_info := initial_browser_info remote_path all_graph_path path;  | 
|
| 6348 | 269  | 
add_index index_text  | 
| 6330 | 270  | 
end;  | 
271  | 
||
272  | 
||
273  | 
(* finish session *)  | 
|
274  | 
||
| 6649 | 275  | 
fun finish () = with_session () (fn {name, html_prefix, graph_path, all_graph_path, path, ...} =>
 | 
| 6339 | 276  | 
let  | 
| 6649 | 277  | 
    val {theories, index, graph, all_graph} = ! browser_info;
 | 
| 6330 | 278  | 
|
| 6649 | 279  | 
    fun finish_node (a, {source = _, html}) =
 | 
280  | 
Buffer.write (Path.append html_prefix (html_path a)) (Buffer.add HTML.end_theory html);  | 
|
| 6339 | 281  | 
in  | 
282  | 
seq finish_node (Symtab.dest theories);  | 
|
283  | 
Buffer.write (Path.append html_prefix pre_index_path) index;  | 
|
| 6649 | 284  | 
put_graph graph graph_path;  | 
285  | 
put_graph all_graph all_graph_path;  | 
|
| 6339 | 286  | 
create_index html_prefix;  | 
287  | 
update_index (Path.append html_prefix Path.parent) name;  | 
|
| 6330 | 288  | 
browser_info := empty_browser_info;  | 
289  | 
session_info := None  | 
|
290  | 
end);  | 
|
291  | 
||
292  | 
||
| 6339 | 293  | 
(* theory elements *)  | 
| 6330 | 294  | 
|
| 6649 | 295  | 
fun theory_source name src = with_session () (fn _ =>  | 
| 6330 | 296  | 
(init_theory_info name empty_theory_info; add_source name (HTML.source src)));  | 
297  | 
||
| 6339 | 298  | 
|
299  | 
(* FIXME clean *)  | 
|
300  | 
||
| 6649 | 301  | 
fun parent_link remote_path curr_session name =  | 
302  | 
let  | 
|
303  | 
    val {session, is_local} = get_info (ThyInfo.theory name);
 | 
|
304  | 
val path = Path.append (Path.make session) (html_path name)  | 
|
305  | 
in (if null session then None else  | 
|
306  | 
Some (if is_some remote_path andalso not is_local then  | 
|
307  | 
Url.append (the remote_path) (Url.file path)  | 
|
308  | 
else Url.file (top_path curr_session 0 path)), name)  | 
|
309  | 
end;  | 
|
| 6361 | 310  | 
|
| 6649 | 311  | 
fun begin_theory name raw_parents orig_files thy =  | 
312  | 
  with_session thy (fn {session, path, html_prefix, remote_path, ...} =>
 | 
|
| 6330 | 313  | 
let  | 
| 6649 | 314  | 
val parents = map (parent_link remote_path path) raw_parents;  | 
| 6330 | 315  | 
val ml_path = ThyLoad.ml_path name;  | 
316  | 
val files = orig_files @ (* FIXME improve!? *)  | 
|
317  | 
(if is_some (ThyLoad.check_file ml_path) then [(ml_path, true)] else []);  | 
|
318  | 
||
| 6348 | 319  | 
fun prep_file (raw_path, loadit) =  | 
| 6330 | 320  | 
(case ThyLoad.check_file raw_path of  | 
321  | 
Some (path, _) =>  | 
|
| 6348 | 322  | 
let  | 
323  | 
val base = Path.base path;  | 
|
324  | 
val base_html = html_ext base;  | 
|
325  | 
in  | 
|
| 6649 | 326  | 
File.write (Path.append html_prefix base_html) (HTML.ml_file (Url.file base) (File.read path));  | 
327  | 
(Some (Url.file base_html), Url.file raw_path, loadit)  | 
|
| 6330 | 328  | 
end  | 
| 6348 | 329  | 
      | None => (warning ("Browser info: expected to find ML file" ^ quote (Path.pack raw_path));
 | 
| 6649 | 330  | 
(None, Url.file raw_path, loadit)));  | 
| 6330 | 331  | 
|
| 6348 | 332  | 
val files_html = map prep_file files;  | 
| 6649 | 333  | 
|
334  | 
fun prep_source (source, html) =  | 
|
| 6330 | 335  | 
let val txt =  | 
| 6649 | 336  | 
HTML.begin_theory (Url.file index_path, session) name parents files_html (Buffer.content source)  | 
337  | 
in (Buffer.empty, Buffer.add txt html) end;  | 
|
338  | 
||
339  | 
fun make_entry unfold all =  | 
|
340  | 
      {name = name, ID = ID_of path name, dir = dir_of path, unfold = unfold,
 | 
|
341  | 
path = Path.pack (top_path (if all then [""] else path) 2  | 
|
342  | 
(Path.append (Path.make path) (html_path name))),  | 
|
343  | 
parents = map (fn s => ID_of (#session (get_info (ThyInfo.theory s))) s) raw_parents};  | 
|
344  | 
||
| 6330 | 345  | 
in  | 
| 6339 | 346  | 
change_theory_info name prep_source;  | 
| 6649 | 347  | 
add_all_graph_entry (make_entry (length path = 1) true);  | 
348  | 
add_graph_entry (make_entry true false);  | 
|
349  | 
add_index (HTML.theory_entry (Url.file (html_path name), name));  | 
|
350  | 
    BrowserInfoData.put {session = path, is_local = is_some remote_path} thy
 | 
|
| 6330 | 351  | 
end);  | 
352  | 
||
| 6339 | 353  | 
fun end_theory _ = ();  | 
| 6330 | 354  | 
|
| 6649 | 355  | 
fun theorem name thm = with_session () (fn _ => with_context add_html (HTML.theorem name thm));  | 
356  | 
fun section s = with_session () (fn _ => with_context add_html (HTML.section s));  | 
|
| 6330 | 357  | 
|
358  | 
||
| 6665 | 359  | 
|
360  | 
(** theory setup **)  | 
|
361  | 
||
362  | 
val setup = [BrowserInfoData.init];  | 
|
363  | 
||
364  | 
||
| 
3603
 
5eef2ff0eb72
This file now contains all functions for generating html
 
berghofe 
parents:  
diff
changeset
 | 
365  | 
end;  |