wenzelm@6203: (* Title: Pure/Thy/present.ML wenzelm@6203: ID: $Id$ wenzelm@9416: Author: Markus Wenzel and Stefan Berghofer, TU Muenchen wenzelm@6203: wenzelm@9416: Theory presentation: HTML, graph files, (PDF)LaTeX documents. wenzelm@6203: *) wenzelm@6203: wenzelm@6203: signature BASIC_PRESENT = wenzelm@6203: sig wenzelm@11057: val no_document: ('a -> 'b) -> 'a -> 'b wenzelm@7727: val section: string -> unit wenzelm@6203: end; wenzelm@6203: wenzelm@6203: signature PRESENT = wenzelm@6203: sig wenzelm@7727: include BASIC_PRESENT wenzelm@14922: val get_info: theory -> {name: string, session: string list, is_local: bool} wenzelm@9452: val write_graph: {name: string, ID: string, dir: string, unfold: bool, wenzelm@9452: path: string, parents: string list} list -> Path.T -> unit wenzelm@17082: val init: bool -> bool -> string -> bool -> string list -> string list -> wenzelm@17210: string -> (bool * Path.T) option -> Url.T option * bool -> bool -> unit wenzelm@7727: val finish: unit -> unit wenzelm@7727: val init_theory: string -> unit wenzelm@8192: val verbatim_source: string -> (unit -> Symbol.symbol list) -> unit wenzelm@9136: val theory_output: string -> string -> unit berghofe@15159: val begin_theory: Path.T option -> string -> string list -> berghofe@15159: (Path.T * bool) list -> theory -> theory wenzelm@13532: val add_hook: (string -> (string * thm list) list -> unit) -> unit wenzelm@13532: val results: string -> (string * thm list) list -> unit wenzelm@7727: val theorem: string -> thm -> unit wenzelm@7727: val theorems: string -> thm list -> unit wenzelm@8088: val chapter: string -> unit wenzelm@7727: val subsection: string -> unit wenzelm@7727: val subsubsection: string -> unit wenzelm@14922: val drafts: string -> Path.T list -> Path.T wenzelm@6203: end; wenzelm@6203: wenzelm@7685: structure Present: PRESENT = wenzelm@7685: struct wenzelm@7685: wenzelm@7727: wenzelm@7727: (** paths **) wenzelm@7727: wenzelm@7727: val output_path = Path.variable "ISABELLE_BROWSER_INFO"; wenzelm@7727: wenzelm@7727: val tex_ext = Path.ext "tex"; wenzelm@7727: val tex_path = tex_ext o Path.basic; wenzelm@7727: val html_ext = Path.ext "html"; wenzelm@7727: val html_path = html_ext o Path.basic; wenzelm@7727: val index_path = Path.basic "index.html"; wenzelm@11856: val readme_html_path = Path.basic "README.html"; wenzelm@11856: val readme_path = Path.basic "README"; wenzelm@17082: val documentN = "document"; wenzelm@17082: val document_path = Path.basic documentN; wenzelm@8196: val doc_indexN = "session"; wenzelm@11856: val graph_path = Path.basic "session.graph"; wenzelm@11856: val graph_pdf_path = Path.basic "session_graph.pdf"; wenzelm@11856: val graph_eps_path = Path.basic "session_graph.eps"; wenzelm@7727: wenzelm@7727: val session_path = Path.basic ".session"; wenzelm@7727: val session_entries_path = Path.unpack ".session/entries"; wenzelm@7727: val pre_index_path = Path.unpack ".session/pre-index"; wenzelm@7727: berghofe@9044: fun mk_rel_path [] ys = Path.make ys berghofe@9044: | mk_rel_path xs [] = Path.appends (replicate (length xs) Path.parent) wenzelm@9416: | mk_rel_path (ps as x :: xs) (qs as y :: ys) = if x = y then mk_rel_path xs ys else berghofe@9044: Path.appends (replicate (length ps) Path.parent @ [Path.make qs]); wenzelm@7727: wenzelm@11856: fun show_path path = Path.pack (Path.append (File.pwd ()) path); wenzelm@11856: wenzelm@7727: wenzelm@14922: wenzelm@7727: (** additional theory data **) wenzelm@7727: wenzelm@16426: structure BrowserInfoData = TheoryDataFun wenzelm@16426: (struct wenzelm@7727: val name = "Pure/browser_info"; wenzelm@9416: type T = {name: string, session: string list, is_local: bool}; wenzelm@16503: val empty = {name = "", session = [], is_local = false}: T; wenzelm@7727: val copy = I; wenzelm@16503: fun extend _ = empty; wenzelm@16426: fun merge _ _ = empty; wenzelm@7727: fun print _ _ = (); wenzelm@16426: end); wenzelm@7727: wenzelm@18708: val _ = Context.add_setup BrowserInfoData.init; wenzelm@7727: wenzelm@7727: fun get_info thy = wenzelm@16426: if Context.theory_name thy mem_string [Context.ProtoPureN, Context.PureN, Context.CPureN] wenzelm@16503: then {name = Context.PureN, session = [], is_local = false} wenzelm@7727: else BrowserInfoData.get thy; wenzelm@7727: wenzelm@7727: wenzelm@7727: wenzelm@7727: (** graphs **) wenzelm@7727: wenzelm@7727: type graph_node = wenzelm@7727: {name: string, ID: string, dir: string, unfold: bool, wenzelm@7727: path: string, parents: string list}; wenzelm@7727: wenzelm@9416: fun write_graph gr path = wenzelm@9416: File.write path (cat_lines (map (fn {name, ID, dir, unfold, path, parents} => wenzelm@9416: "\"" ^ name ^ "\" \"" ^ ID ^ "\" \"" ^ dir ^ (if unfold then "\" + \"" else "\" \"") ^ berghofe@14598: path ^ "\" > " ^ space_implode " " (map Library.quote parents) ^ " ;") gr)); wenzelm@7727: wenzelm@9416: fun ID_of sess s = space_implode "/" (sess @ [s]); wenzelm@7727: wenzelm@9416: (*retrieve graph data from initial theory loader database*) wenzelm@9416: fun init_graph remote_path curr_sess = map (fn name => wenzelm@7727: let wenzelm@9416: val {name = sess_name, session, is_local} = get_info (ThyInfo.theory name); wenzelm@9416: val path' = Path.append (Path.make session) (html_path name); wenzelm@7727: in wenzelm@9416: {name = name, ID = ID_of session name, dir = sess_name, wenzelm@9416: path = wenzelm@9416: if null session then "" else wenzelm@16263: if is_some remote_path andalso not is_local then wenzelm@16263: Url.pack (Url.append (the remote_path) (Url.File wenzelm@9416: (Path.append (Path.make session) (html_path name)))) wenzelm@9416: else Path.pack (Path.append (mk_rel_path curr_sess session) (html_path name)), wenzelm@9416: unfold = false, wenzelm@9416: parents = wenzelm@9416: map (fn s => ID_of (#session (get_info (ThyInfo.theory s))) s) (ThyInfo.get_preds name)} wenzelm@7727: end) (ThyInfo.names ()); wenzelm@7727: wenzelm@9416: fun ins_graph_entry (entry as {ID, ...}) (gr: graph_node list) = wenzelm@7727: filter_out (fn entry' => #ID entry' = ID) gr @ [entry]; wenzelm@7727: wenzelm@7727: wenzelm@7727: wenzelm@7727: (** global browser info state **) wenzelm@7727: wenzelm@7727: (* type theory_info *) wenzelm@7727: wenzelm@7727: type theory_info = {tex_source: Buffer.T, html_source: Buffer.T, html: Buffer.T}; wenzelm@7727: wenzelm@7727: fun make_theory_info (tex_source, html_source, html) = wenzelm@7727: {tex_source = tex_source, html_source = html_source, html = html}: theory_info; wenzelm@7727: wenzelm@7727: val empty_theory_info = make_theory_info (Buffer.empty, Buffer.empty, Buffer.empty); wenzelm@7727: wenzelm@7727: fun map_theory_info f {tex_source, html_source, html} = wenzelm@7727: make_theory_info (f (tex_source, html_source, html)); wenzelm@7727: wenzelm@7727: wenzelm@7727: (* type browser_info *) wenzelm@7727: wenzelm@11856: type browser_info = {theories: theory_info Symtab.table, files: (Path.T * string) list, wenzelm@11856: tex_index: Buffer.T, html_index: Buffer.T, graph: graph_node list}; wenzelm@7727: wenzelm@11856: fun make_browser_info (theories, files, tex_index, html_index, graph) = wenzelm@11856: {theories = theories, files = files, tex_index = tex_index, html_index = html_index, wenzelm@9416: graph = graph}: browser_info; wenzelm@7727: wenzelm@11856: val empty_browser_info = make_browser_info (Symtab.empty, [], Buffer.empty, Buffer.empty, []); wenzelm@7727: wenzelm@9416: fun init_browser_info remote_path curr_sess = make_browser_info wenzelm@11856: (Symtab.empty, [], Buffer.empty, Buffer.empty, init_graph remote_path curr_sess); wenzelm@7727: wenzelm@11856: fun map_browser_info f {theories, files, tex_index, html_index, graph} = wenzelm@11856: make_browser_info (f (theories, files, tex_index, html_index, graph)); wenzelm@7727: wenzelm@7727: wenzelm@7727: (* state *) wenzelm@7727: wenzelm@7727: val browser_info = ref empty_browser_info; wenzelm@11057: fun change_browser_info f = browser_info := map_browser_info f (! browser_info); wenzelm@7727: wenzelm@11057: val suppress_tex_source = ref false; wenzelm@11057: fun no_document f x = Library.setmp suppress_tex_source true f x; wenzelm@7727: wenzelm@7727: fun init_theory_info name info = wenzelm@11856: change_browser_info (fn (theories, files, tex_index, html_index, graph) => wenzelm@17412: (Symtab.update (name, info) theories, files, tex_index, html_index, graph)); wenzelm@7727: wenzelm@7727: fun change_theory_info name f = wenzelm@11856: change_browser_info (fn (info as (theories, files, tex_index, html_index, graph)) => wenzelm@17412: (case Symtab.lookup theories name of skalberg@15531: NONE => (warning ("Browser info: cannot access theory document " ^ quote name); info) wenzelm@17412: | SOME info => (Symtab.update (name, map_theory_info f info) theories, files, wenzelm@9416: tex_index, html_index, graph))); wenzelm@7727: wenzelm@7727: wenzelm@11856: fun add_file file = wenzelm@11856: change_browser_info (fn (theories, files, tex_index, html_index, graph) => wenzelm@11856: (theories, file :: files, tex_index, html_index, graph)); wenzelm@11856: wenzelm@7727: fun add_tex_index txt = wenzelm@11856: change_browser_info (fn (theories, files, tex_index, html_index, graph) => wenzelm@11856: (theories, files, Buffer.add txt tex_index, html_index, graph)); wenzelm@7727: wenzelm@7727: fun add_html_index txt = wenzelm@11856: change_browser_info (fn (theories, files, tex_index, html_index, graph) => wenzelm@11856: (theories, files, tex_index, Buffer.add txt html_index, graph)); wenzelm@7727: wenzelm@7727: fun add_graph_entry e = wenzelm@11856: change_browser_info (fn (theories, files, tex_index, html_index, graph) => wenzelm@11856: (theories, files, tex_index, html_index, ins_graph_entry e graph)); wenzelm@7727: wenzelm@11057: fun add_tex_source name txt = wenzelm@11057: if ! suppress_tex_source then () wenzelm@11057: else change_theory_info name (fn (tex_source, html_source, html) => wenzelm@11057: (Buffer.add txt tex_source, html_source, html)); wenzelm@7727: wenzelm@7727: fun add_html_source name txt = change_theory_info name (fn (tex_source, html_source, html) => wenzelm@7727: (tex_source, Buffer.add txt html_source, html)); wenzelm@7727: wenzelm@7727: fun add_html name txt = change_theory_info name (fn (tex_source, html_source, html) => wenzelm@7727: (tex_source, html_source, Buffer.add txt html)); wenzelm@7727: wenzelm@7727: wenzelm@7727: wenzelm@7727: (** global session state **) wenzelm@7727: wenzelm@7727: (* session_info *) wenzelm@7727: wenzelm@7727: type session_info = wenzelm@7727: {name: string, parent: string, session: string, path: string list, html_prefix: Path.T, wenzelm@17082: info: bool, doc_format: string, doc_graph: bool, documents: (string * string) list, wenzelm@17210: doc_prefix1: (Path.T * Path.T) option, doc_prefix2: (bool * Path.T) option, wenzelm@17210: remote_path: Url.T option, verbose: bool, readme: Path.T option}; wenzelm@7727: wenzelm@9416: fun make_session_info wenzelm@17082: (name, parent, session, path, html_prefix, info, doc_format, doc_graph, documents, wenzelm@17082: doc_prefix1, doc_prefix2, remote_path, verbose, readme) = wenzelm@7802: {name = name, parent = parent, session = session, path = path, html_prefix = html_prefix, wenzelm@17082: info = info, doc_format = doc_format, doc_graph = doc_graph, documents = documents, wenzelm@17082: doc_prefix1 = doc_prefix1, doc_prefix2 = doc_prefix2, remote_path = remote_path, wenzelm@17082: verbose = verbose, readme = readme}: session_info; wenzelm@7685: wenzelm@7685: wenzelm@7727: (* state *) wenzelm@7727: skalberg@15531: val session_info = ref (NONE: session_info option); wenzelm@7727: skalberg@15531: fun with_session x f = (case ! session_info of NONE => x | SOME info => f info); wenzelm@16426: fun with_context f = f (Context.theory_name (Context.the_context ())); wenzelm@7727: wenzelm@7727: wenzelm@7727: wenzelm@14922: (** document preparation **) wenzelm@7727: wenzelm@7727: (* maintain index *) wenzelm@7727: wenzelm@7727: val session_entries = wenzelm@7727: HTML.session_entries o wenzelm@14898: map (fn name => (Url.File (Path.append (Path.basic name) index_path), name)); wenzelm@7727: wenzelm@7727: fun get_entries dir = wenzelm@7727: split_lines (File.read (Path.append dir session_entries_path)); wenzelm@7727: wenzelm@7727: fun put_entries entries dir = wenzelm@7727: File.write (Path.append dir session_entries_path) (cat_lines entries); wenzelm@7727: wenzelm@7727: wenzelm@7727: fun create_index dir = wenzelm@7727: File.read (Path.append dir pre_index_path) ^ wenzelm@7727: session_entries (get_entries dir) ^ HTML.end_index wenzelm@7727: |> File.write (Path.append dir index_path); wenzelm@7727: wenzelm@7727: fun update_index dir name = wenzelm@7727: (case try get_entries dir of skalberg@15531: NONE => warning ("Browser info: cannot access session index of " ^ quote (Path.pack dir)) skalberg@15531: | SOME es => (put_entries ((es \ name) @ [name]) dir; create_index dir)); wenzelm@7727: wenzelm@7727: wenzelm@17082: (* document versions *) wenzelm@17082: wenzelm@17082: fun read_version str = wenzelm@17082: (case space_explode "=" str of wenzelm@17082: [name] => (name, "") wenzelm@17082: | [name, tags] => (name, tags) wenzelm@17082: | _ => error ("Malformed document version specification: " ^ quote str)); wenzelm@17082: wenzelm@17082: fun read_versions strs = wenzelm@19046: rev (distinct (eq_fst (op =)) (rev ((documentN, "") :: map read_version strs))) wenzelm@17082: |> filter_out (equal "-" o #2); wenzelm@17082: wenzelm@17082: wenzelm@7727: (* init session *) wenzelm@7727: wenzelm@12895: fun name_of_session elems = space_implode "/" ("Isabelle" :: elems); wenzelm@12895: wenzelm@17082: fun init build info doc doc_graph doc_versions path name doc_prefix2 wenzelm@17082: (remote_path, first_time) verbose = wenzelm@11911: if not build andalso not info andalso doc = "" andalso is_none doc_prefix2 then skalberg@15531: (browser_info := empty_browser_info; session_info := NONE) wenzelm@11856: else wenzelm@11856: let wenzelm@11856: val parent_name = name_of_session (Library.take (length path - 1, path)); wenzelm@11856: val session_name = name_of_session path; wenzelm@11856: val sess_prefix = Path.make path; wenzelm@11856: val html_prefix = Path.append (Path.expand output_path) sess_prefix; wenzelm@7727: wenzelm@17082: val (doc_prefix1, documents) = wenzelm@17082: if doc = "" then (NONE, []) wenzelm@17082: else if not (File.exists document_path) then (conditional verbose (fn () => wenzelm@17082: std_error "Warning: missing document directory\n"); (NONE, [])) wenzelm@17082: else (SOME (Path.append html_prefix document_path, html_prefix), wenzelm@17082: read_versions doc_versions); wenzelm@7727: wenzelm@11856: val parent_index_path = Path.append Path.parent index_path; wenzelm@11856: val index_up_lnk = if first_time then wenzelm@16263: Url.append (the remote_path) (Url.File (Path.append sess_prefix parent_index_path)) wenzelm@14898: else Url.File parent_index_path; wenzelm@11856: val readme = skalberg@15531: if File.exists readme_html_path then SOME readme_html_path skalberg@15531: else if File.exists readme_path then SOME readme_path skalberg@15531: else NONE; wenzelm@7727: wenzelm@17082: val docs = wenzelm@17082: (case readme of NONE => [] | SOME p => [(Url.File p, "README")]) @ wenzelm@17082: map (fn (name, _) => (Url.File (Path.ext doc (Path.basic name)), name)) documents; wenzelm@11856: val index_text = HTML.begin_index (index_up_lnk, parent_name) wenzelm@17082: (Url.File index_path, session_name) docs (Url.unpack "medium.html"); wenzelm@11856: in skalberg@15531: session_info := SOME (make_session_info (name, parent_name, session_name, path, html_prefix, wenzelm@17082: info, doc, doc_graph, documents, doc_prefix1, doc_prefix2, remote_path, verbose, readme)); wenzelm@11856: browser_info := init_browser_info remote_path path; wenzelm@11856: add_html_index index_text wenzelm@11856: end; wenzelm@7727: wenzelm@7727: wenzelm@17082: (* isatool wrappers *) wenzelm@17082: wenzelm@17082: fun isatool_document verbose format name tags path result_path = wenzelm@17082: let wenzelm@17082: val s = "\"$ISATOOL\" document -c -o '" ^ format ^ "' \ wenzelm@17082: \-n '" ^ name ^ "' -t '" ^ tags ^ "' " ^ File.shell_path path ^ wenzelm@17082: " 2>&1" ^ (if verbose then "" else " >/dev/null"); wenzelm@17082: val doc_path = Path.append result_path (Path.ext format (Path.basic name)); wenzelm@17082: in wenzelm@17082: conditional verbose (fn () => writeln s); wenzelm@17082: system s; wenzelm@17082: conditional (not (File.exists doc_path)) (fn () => wenzelm@17082: error ("No document: " ^ quote (show_path doc_path))); wenzelm@17082: doc_path wenzelm@17082: end; wenzelm@17082: wenzelm@17082: fun isatool_browser graph = wenzelm@17082: let wenzelm@17082: val pdf_path = File.tmp_path graph_pdf_path; wenzelm@17082: val eps_path = File.tmp_path graph_eps_path; wenzelm@17082: val gr_path = File.tmp_path graph_path; wenzelm@17082: val s = "browser -o " ^ File.shell_path pdf_path ^ " " ^ File.shell_path gr_path; wenzelm@17082: in wenzelm@17082: write_graph graph gr_path; wenzelm@17082: if File.isatool s <> 0 orelse not (File.exists eps_path) orelse not (File.exists pdf_path) wenzelm@17082: then error "Failed to prepare dependency graph" wenzelm@17082: else wenzelm@17082: let val pdf = File.read pdf_path and eps = File.read eps_path wenzelm@17082: in File.rm pdf_path; File.rm eps_path; File.rm gr_path; (pdf, eps) end wenzelm@17082: end; wenzelm@17082: wenzelm@17082: wenzelm@11856: (* finish session -- output all generated text *) wenzelm@11856: wenzelm@14922: fun write_tex src name path = wenzelm@14922: Buffer.write (Path.append path (tex_path name)) src; wenzelm@14922: wenzelm@14922: fun write_tex_index tex_index path = wenzelm@14922: write_tex (Buffer.add Latex.tex_trailer tex_index) doc_indexN path; wenzelm@14922: wenzelm@7685: wenzelm@11856: fun finish () = with_session () (fn {name, info, html_prefix, doc_format, doc_graph, wenzelm@17082: documents, doc_prefix1, doc_prefix2, path, verbose, readme, ...} => wenzelm@7727: let wenzelm@11856: val {theories, files, tex_index, html_index, graph} = ! browser_info; wenzelm@11856: val thys = Symtab.dest theories; wenzelm@9416: val parent_html_prefix = Path.append html_prefix Path.parent; wenzelm@7727: wenzelm@11856: fun finish_tex path (a, {tex_source, ...}: theory_info) = write_tex tex_source a path; wenzelm@11856: fun finish_html (a, {html, ...}: theory_info) = wenzelm@11856: Buffer.write (Path.append html_prefix (html_path a)) (Buffer.add HTML.end_theory html); wenzelm@11856: wenzelm@11856: val opt_graphs = wenzelm@16263: if doc_graph andalso (is_some doc_prefix1 orelse is_some doc_prefix2) then skalberg@15531: SOME (isatool_browser graph) skalberg@15531: else NONE; wenzelm@11856: wenzelm@17210: fun prepare_sources cp path = wenzelm@16263: (File.mkdir path; wenzelm@17210: if cp then File.copy_dir document_path path else (); wenzelm@17177: File.isatool ("latex -o sty " ^ File.shell_path (Path.append path (Path.basic "root.tex"))); skalberg@15531: (case opt_graphs of NONE => () | SOME (pdf, eps) => wenzelm@11856: (File.write (Path.append path graph_pdf_path) pdf; wenzelm@11856: File.write (Path.append path graph_eps_path) eps)); wenzelm@14922: write_tex_index tex_index path; skalberg@15570: List.app (finish_tex path) thys); wenzelm@7727: in wenzelm@11856: conditional info (fn () => wenzelm@11856: (File.mkdir (Path.append html_prefix session_path); wenzelm@11856: Buffer.write (Path.append html_prefix pre_index_path) html_index; wenzelm@11856: File.write (Path.append html_prefix session_entries_path) ""; wenzelm@11856: create_index html_prefix; wenzelm@11856: if length path > 1 then update_index parent_html_prefix name else (); wenzelm@16263: (case readme of NONE => () | SOME path => File.copy path html_prefix); wenzelm@11856: write_graph graph (Path.append html_prefix graph_path); wenzelm@16263: File.copy (Path.unpack "~~/lib/browser/GraphBrowser.jar") html_prefix; skalberg@15570: List.app (fn (a, txt) => File.write (Path.append html_prefix (Path.basic a)) txt) wenzelm@14898: (HTML.applet_pages name (Url.File index_path, name)); wenzelm@16263: File.copy (Path.unpack "~~/lib/html/isabelle.css") html_prefix; skalberg@15570: List.app finish_html thys; skalberg@15570: List.app (uncurry File.write) files; wenzelm@11856: conditional verbose (fn () => wenzelm@11856: std_error ("Browser info at " ^ show_path html_prefix ^ "\n")))); wenzelm@11856: wenzelm@17210: (case doc_prefix2 of NONE => () | SOME (cp, path) => wenzelm@17210: (prepare_sources cp path; wenzelm@12895: conditional verbose (fn () => std_error ("Document sources at " ^ show_path path ^ "\n")))); wenzelm@12895: wenzelm@17082: (case doc_prefix1 of NONE => () | SOME (path, result_path) => wenzelm@17082: documents |> List.app (fn (name, tags) => wenzelm@17082: let wenzelm@17210: val _ = prepare_sources true path; wenzelm@17082: val doc_path = isatool_document true doc_format name tags path result_path; wenzelm@17082: in wenzelm@17082: conditional verbose (fn () => std_error ("Document at " ^ show_path doc_path ^ "\n")) wenzelm@17082: end)); wenzelm@11856: wenzelm@7727: browser_info := empty_browser_info; skalberg@15531: session_info := NONE wenzelm@7727: end); wenzelm@7727: wenzelm@7727: wenzelm@7727: (* theory elements *) wenzelm@7727: wenzelm@7727: fun init_theory name = with_session () (fn _ => init_theory_info name empty_theory_info); wenzelm@7727: wenzelm@7727: fun verbatim_source name mk_text = wenzelm@7727: with_session () (fn _ => add_html_source name (HTML.verbatim_source (mk_text ()))); wenzelm@7727: wenzelm@9136: fun theory_output name s = wenzelm@9917: with_session () (fn _ => add_tex_source name (Latex.isabelle_file name s)); wenzelm@7727: wenzelm@7727: wenzelm@7727: fun parent_link remote_path curr_session name = wenzelm@9416: let val {name = _, session, is_local} = get_info (ThyInfo.theory name) skalberg@15531: in (if null session then NONE else wenzelm@16263: SOME (if is_some remote_path andalso not is_local then wenzelm@16263: Url.append (the remote_path) (Url.File berghofe@9044: (Path.append (Path.make session) (html_path name))) wenzelm@14898: else Url.File (Path.append (mk_rel_path curr_session session) berghofe@9044: (html_path name))), name) wenzelm@7727: end; wenzelm@7727: berghofe@15159: fun begin_theory optpath name raw_parents orig_files thy = wenzelm@9416: with_session thy (fn {name = sess_name, session, path, html_prefix, remote_path, ...} => wenzelm@7727: let wenzelm@7727: val parents = map (parent_link remote_path path) raw_parents; wenzelm@7727: val ml_path = ThyLoad.ml_path name; skalberg@15531: val files = map (apsnd SOME) orig_files @ wenzelm@16263: (if is_some (ThyLoad.check_file optpath ml_path) then [(ml_path, NONE)] else []); wenzelm@7727: wenzelm@7727: fun prep_file (raw_path, loadit) = berghofe@15159: (case ThyLoad.check_file optpath raw_path of skalberg@15531: SOME (path, _) => wenzelm@7727: let wenzelm@7727: val base = Path.base path; wenzelm@7727: val base_html = html_ext base; wenzelm@7727: in wenzelm@11856: add_file (Path.append html_prefix base_html, wenzelm@14898: HTML.ml_file (Url.File base) (File.read path)); skalberg@15531: (SOME (Url.File base_html), Url.File raw_path, loadit) wenzelm@7727: end skalberg@15531: | NONE => (warning ("Browser info: expected to find ML file" ^ quote (Path.pack raw_path)); skalberg@15531: (NONE, Url.File raw_path, loadit))); wenzelm@7727: wenzelm@7727: val files_html = map prep_file files; wenzelm@7727: wenzelm@7727: fun prep_html_source (tex_source, html_source, html) = wenzelm@7727: let wenzelm@14898: val txt = HTML.begin_theory (Url.File index_path, session) wenzelm@7727: name parents files_html (Buffer.content html_source) wenzelm@7727: in (tex_source, Buffer.empty, Buffer.add txt html) end; wenzelm@7727: wenzelm@9416: val entry = wenzelm@9416: {name = name, ID = ID_of path name, dir = sess_name, unfold = true, wenzelm@9416: path = Path.pack (html_path name), wenzelm@9416: parents = map (fn s => ID_of (#session (get_info (ThyInfo.theory s))) s) raw_parents}; wenzelm@7727: wenzelm@7727: in wenzelm@7727: change_theory_info name prep_html_source; wenzelm@9416: add_graph_entry entry; wenzelm@14898: add_html_index (HTML.theory_entry (Url.File (html_path name), name)); wenzelm@7727: add_tex_index (Latex.theory_entry name); wenzelm@16263: BrowserInfoData.put {name = sess_name, session = path, is_local = is_some remote_path} thy wenzelm@7727: end); wenzelm@7727: wenzelm@7727: wenzelm@13532: val hooks = ref ([]: (string -> (string * thm list) list -> unit) list); wenzelm@13532: fun add_hook f = hooks := (f :: ! hooks); wenzelm@13532: wenzelm@13532: fun results k xs = skalberg@15570: (List.app (fn f => (try (fn () => f k xs) (); ())) (! hooks); wenzelm@13532: with_session () (fn _ => with_context add_html (HTML.results k xs))); wenzelm@13532: wenzelm@13532: fun theorem a th = results "theorem" [(a, [th])]; wenzelm@13532: fun theorems a ths = results "theorems" [(a, ths)]; wenzelm@8088: fun chapter s = with_session () (fn _ => with_context add_html (HTML.chapter s)); wenzelm@7727: fun section s = with_session () (fn _ => with_context add_html (HTML.section s)); wenzelm@7727: fun subsection s = with_session () (fn _ => with_context add_html (HTML.subsection s)); wenzelm@7727: fun subsubsection s = with_session () (fn _ => with_context add_html (HTML.subsubsection s)); wenzelm@7727: wenzelm@7727: wenzelm@7727: wenzelm@14922: (** draft document output **) wenzelm@14922: wenzelm@14922: fun drafts doc_format src_paths = wenzelm@14922: let wenzelm@14935: fun prep_draft (tex_index, path) = wenzelm@14935: let wenzelm@14935: val base = Path.base path; wenzelm@14972: val name = wenzelm@14972: (case Path.pack (#1 (Path.split_ext base)) of "" => gensym "DUMMY" | s => s); wenzelm@14935: in wenzelm@14935: if File.exists path then wenzelm@14935: (Buffer.add (Latex.theory_entry name) tex_index, (name, base, File.read path)) wenzelm@14935: else error ("Bad file: " ^ quote (Path.pack path)) wenzelm@14935: end; wenzelm@14935: val (tex_index, srcs) = foldl_map prep_draft (Buffer.empty, src_paths); wenzelm@14935: wenzelm@17082: val doc_path = File.tmp_path document_path; wenzelm@17082: val result_path = Path.append doc_path Path.parent; wenzelm@14922: val _ = File.mkdir doc_path; wenzelm@14922: val root_path = Path.append doc_path (Path.basic "root.tex"); wenzelm@14922: val _ = File.copy (Path.unpack "~~/lib/texinputs/draft.tex") root_path; wenzelm@16263: val _ = File.isatool ("latex -o sty " ^ File.shell_path root_path); wenzelm@16263: val _ = File.isatool ("latex -o syms " ^ File.shell_path root_path); wenzelm@14922: wenzelm@14922: fun known name = wenzelm@14922: let val ss = split_lines (File.read (Path.append doc_path (Path.basic name))) wenzelm@14922: in fn s => s mem_string ss end; wenzelm@14922: val known_syms = known "syms.lst"; wenzelm@14922: val known_ctrls = known "ctrls.lst"; wenzelm@14922: skalberg@15570: val _ = srcs |> List.app (fn (name, base, txt) => wenzelm@14935: Symbol.explode txt wenzelm@14935: |> Latex.symbol_source (known_syms, known_ctrls) (Path.pack base) wenzelm@14935: |> File.write (Path.append doc_path (tex_path name))); wenzelm@14922: val _ = write_tex_index tex_index doc_path; wenzelm@17082: in isatool_document false doc_format documentN "" doc_path result_path end; wenzelm@14922: wenzelm@14922: wenzelm@7685: end; wenzelm@7685: wenzelm@6203: structure BasicPresent: BASIC_PRESENT = Present; wenzelm@6203: open BasicPresent;