| author | wenzelm | 
| Tue, 09 Dec 2014 18:29:45 +0100 | |
| changeset 59117 | caddfa6ca534 | 
| parent 55041 | 368ee97e03ce | 
| child 59448 | 149d2bc5ddb6 | 
| permissions | -rw-r--r-- | 
| 24584 | 1 | (* Title: Pure/Thy/html.ML | 
| 9415 | 2 | Author: Markus Wenzel and Stefan Berghofer, TU Muenchen | 
| 6324 | 3 | |
| 9415 | 4 | HTML presentation elements. | 
| 6324 | 5 | *) | 
| 6 | ||
| 7 | signature HTML = | |
| 8 | sig | |
| 23724 | 9 |   val html_mode: ('a -> 'b) -> 'a -> 'b
 | 
| 23622 | 10 | type text = string | 
| 6324 | 11 | val plain: string -> text | 
| 12 | val name: string -> text | |
| 13 | val keyword: string -> text | |
| 23622 | 14 | val command: string -> text | 
| 6324 | 15 | val href_name: string -> text -> text | 
| 6649 | 16 | val href_path: Url.T -> text -> text | 
| 17 | val href_opt_path: Url.T option -> text -> text | |
| 6376 | 18 | val para: text -> text | 
| 6324 | 19 | val preform: text -> text | 
| 20 | val verbatim: string -> text | |
| 23724 | 21 | val begin_document: string -> text | 
| 22 | val end_document: text | |
| 51399 
6ac3c29a300e
discontinued "isabelle usedir" option -r (reset session path);
 wenzelm parents: 
50233diff
changeset | 23 | val begin_session_index: string -> (Url.T * string) list -> Url.T -> text | 
| 9415 | 24 | val applet_pages: string -> Url.T * string -> (string * string) list | 
| 6649 | 25 | val theory_entry: Url.T * string -> text | 
| 54456 | 26 | val theory: string -> (Url.T option * string) list -> text -> text | 
| 6324 | 27 | end; | 
| 28 | ||
| 29 | structure HTML: HTML = | |
| 30 | struct | |
| 31 | ||
| 32 | ||
| 33 | (** HTML print modes **) | |
| 34 | ||
| 35 | (* mode *) | |
| 36 | ||
| 37 | val htmlN = "HTML"; | |
| 37146 
f652333bbf8e
renamed structure PrintMode to Print_Mode, keeping the old name as legacy alias for some time;
 wenzelm parents: 
34003diff
changeset | 38 | fun html_mode f x = Print_Mode.with_modes [htmlN] f x; | 
| 6324 | 39 | |
| 40 | ||
| 33985 | 41 | (* common markup *) | 
| 42 | ||
| 43 | fun span class = ("<span class=" ^ quote (XML.text class) ^ ">", "</span>");
 | |
| 44 | ||
| 46866 
b190913c3c41
simplified span class in conformance to Scala version;
 wenzelm parents: 
45666diff
changeset | 45 | val _ = Markup.add_mode htmlN (span o fst); | 
| 33985 | 46 | |
| 47 | ||
| 6324 | 48 | (* symbol output *) | 
| 49 | ||
| 6338 | 50 | local | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
48670diff
changeset | 51 | val hidden = span Markup.hiddenN |-> enclose; | 
| 33985 | 52 | |
| 43592 
e67d104c0c50
HTML: render control symbols more like Isabelle/Scala/jEdit;
 wenzelm parents: 
43548diff
changeset | 53 | (* FIXME proper unicode -- produced on Scala side *) | 
| 16196 | 54 | val html_syms = Symtab.make | 
| 34003 
610e41138486
output_syms: permissive treatment of control symbols, cf. Scala version;
 wenzelm parents: 
33986diff
changeset | 55 |    [("", (0, "")),
 | 
| 33986 
041dc6d8d344
output linefeed as </br> -- workaround problem with <pre> in Lobo Browser 0.98.4;
 wenzelm parents: 
33985diff
changeset | 56 |     ("'", (1, "'")),
 | 
| 23622 | 57 |     ("\\<exclamdown>", (1, "¡")),
 | 
| 58 |     ("\\<cent>", (1, "¢")),
 | |
| 59 |     ("\\<pounds>", (1, "£")),
 | |
| 60 |     ("\\<currency>", (1, "¤")),
 | |
| 61 |     ("\\<yen>", (1, "¥")),
 | |
| 62 |     ("\\<bar>", (1, "¦")),
 | |
| 63 |     ("\\<section>", (1, "§")),
 | |
| 64 |     ("\\<dieresis>", (1, "¨")),
 | |
| 65 |     ("\\<copyright>", (1, "©")),
 | |
| 66 |     ("\\<ordfeminine>", (1, "ª")),
 | |
| 67 |     ("\\<guillemotleft>", (1, "«")),
 | |
| 68 |     ("\\<not>", (1, "¬")),
 | |
| 69 |     ("\\<hyphen>", (1, "­")),
 | |
| 70 |     ("\\<registered>", (1, "®")),
 | |
| 71 |     ("\\<inverse>", (1, "¯")),
 | |
| 72 |     ("\\<degree>", (1, "°")),
 | |
| 73 |     ("\\<plusminus>", (1, "±")),
 | |
| 74 |     ("\\<acute>", (1, "´")),
 | |
| 75 |     ("\\<paragraph>", (1, "¶")),
 | |
| 76 |     ("\\<cdot>", (1, "·")),
 | |
| 77 |     ("\\<cedilla>", (1, "¸")),
 | |
| 78 |     ("\\<ordmasculine>", (1, "º")),
 | |
| 79 |     ("\\<guillemotright>", (1, "»")),
 | |
| 80 |     ("\\<onequarter>", (1, "¼")),
 | |
| 81 |     ("\\<onehalf>", (1, "½")),
 | |
| 82 |     ("\\<threequarters>", (1, "¾")),
 | |
| 83 |     ("\\<questiondown>", (1, "¿")),
 | |
| 84 |     ("\\<times>", (1, "×")),
 | |
| 85 |     ("\\<div>", (1, "÷")),
 | |
| 86 |     ("\\<circ>", (1, "o")),
 | |
| 87 |     ("\\<Alpha>", (1, "Α")),
 | |
| 88 |     ("\\<Beta>", (1, "Β")),
 | |
| 89 |     ("\\<Gamma>", (1, "Γ")),
 | |
| 90 |     ("\\<Delta>", (1, "Δ")),
 | |
| 91 |     ("\\<Epsilon>", (1, "Ε")),
 | |
| 92 |     ("\\<Zeta>", (1, "Ζ")),
 | |
| 93 |     ("\\<Eta>", (1, "Η")),
 | |
| 94 |     ("\\<Theta>", (1, "Θ")),
 | |
| 95 |     ("\\<Iota>", (1, "Ι")),
 | |
| 96 |     ("\\<Kappa>", (1, "Κ")),
 | |
| 97 |     ("\\<Lambda>", (1, "Λ")),
 | |
| 98 |     ("\\<Mu>", (1, "Μ")),
 | |
| 99 |     ("\\<Nu>", (1, "Ν")),
 | |
| 100 |     ("\\<Xi>", (1, "Ξ")),
 | |
| 101 |     ("\\<Omicron>", (1, "Ο")),
 | |
| 102 |     ("\\<Pi>", (1, "Π")),
 | |
| 103 |     ("\\<Rho>", (1, "Ρ")),
 | |
| 104 |     ("\\<Sigma>", (1, "Σ")),
 | |
| 105 |     ("\\<Tau>", (1, "Τ")),
 | |
| 106 |     ("\\<Upsilon>", (1, "Υ")),
 | |
| 107 |     ("\\<Phi>", (1, "Φ")),
 | |
| 108 |     ("\\<Chi>", (1, "Χ")),
 | |
| 109 |     ("\\<Psi>", (1, "Ψ")),
 | |
| 110 |     ("\\<Omega>", (1, "Ω")),
 | |
| 111 |     ("\\<alpha>", (1, "α")),
 | |
| 112 |     ("\\<beta>", (1, "β")),
 | |
| 113 |     ("\\<gamma>", (1, "γ")),
 | |
| 114 |     ("\\<delta>", (1, "δ")),
 | |
| 115 |     ("\\<epsilon>", (1, "ε")),
 | |
| 116 |     ("\\<zeta>", (1, "ζ")),
 | |
| 117 |     ("\\<eta>", (1, "η")),
 | |
| 118 |     ("\\<theta>", (1, "ϑ")),
 | |
| 119 |     ("\\<iota>", (1, "ι")),
 | |
| 120 |     ("\\<kappa>", (1, "κ")),
 | |
| 121 |     ("\\<lambda>", (1, "λ")),
 | |
| 122 |     ("\\<mu>", (1, "μ")),
 | |
| 123 |     ("\\<nu>", (1, "ν")),
 | |
| 124 |     ("\\<xi>", (1, "ξ")),
 | |
| 125 |     ("\\<omicron>", (1, "ο")),
 | |
| 126 |     ("\\<pi>", (1, "π")),
 | |
| 127 |     ("\\<rho>", (1, "ρ")),
 | |
| 128 |     ("\\<sigma>", (1, "σ")),
 | |
| 129 |     ("\\<tau>", (1, "τ")),
 | |
| 130 |     ("\\<upsilon>", (1, "υ")),
 | |
| 131 |     ("\\<phi>", (1, "φ")),
 | |
| 132 |     ("\\<chi>", (1, "χ")),
 | |
| 133 |     ("\\<psi>", (1, "ψ")),
 | |
| 134 |     ("\\<omega>", (1, "ω")),
 | |
| 135 |     ("\\<bullet>", (1, "•")),
 | |
| 136 |     ("\\<dots>", (1, "…")),
 | |
| 137 |     ("\\<wp>", (1, "℘")),
 | |
| 138 |     ("\\<forall>", (1, "∀")),
 | |
| 139 |     ("\\<partial>", (1, "∂")),
 | |
| 140 |     ("\\<exists>", (1, "∃")),
 | |
| 141 |     ("\\<emptyset>", (1, "∅")),
 | |
| 142 |     ("\\<nabla>", (1, "∇")),
 | |
| 143 |     ("\\<in>", (1, "∈")),
 | |
| 144 |     ("\\<notin>", (1, "∉")),
 | |
| 145 |     ("\\<Prod>", (1, "∏")),
 | |
| 146 |     ("\\<Sum>", (1, "∑")),
 | |
| 147 |     ("\\<star>", (1, "∗")),
 | |
| 148 |     ("\\<propto>", (1, "∝")),
 | |
| 149 |     ("\\<infinity>", (1, "∞")),
 | |
| 150 |     ("\\<angle>", (1, "∠")),
 | |
| 151 |     ("\\<and>", (1, "∧")),
 | |
| 152 |     ("\\<or>", (1, "∨")),
 | |
| 153 |     ("\\<inter>", (1, "∩")),
 | |
| 154 |     ("\\<union>", (1, "∪")),
 | |
| 155 |     ("\\<sim>", (1, "∼")),
 | |
| 156 |     ("\\<cong>", (1, "≅")),
 | |
| 157 |     ("\\<approx>", (1, "≈")),
 | |
| 158 |     ("\\<noteq>", (1, "≠")),
 | |
| 159 |     ("\\<equiv>", (1, "≡")),
 | |
| 160 |     ("\\<le>", (1, "≤")),
 | |
| 161 |     ("\\<ge>", (1, "≥")),
 | |
| 162 |     ("\\<subset>", (1, "⊂")),
 | |
| 163 |     ("\\<supset>", (1, "⊃")),
 | |
| 164 |     ("\\<subseteq>", (1, "⊆")),
 | |
| 165 |     ("\\<supseteq>", (1, "⊇")),
 | |
| 166 |     ("\\<oplus>", (1, "⊕")),
 | |
| 167 |     ("\\<otimes>", (1, "⊗")),
 | |
| 168 |     ("\\<bottom>", (1, "⊥")),
 | |
| 169 |     ("\\<lceil>", (1, "⌈")),
 | |
| 170 |     ("\\<rceil>", (1, "⌉")),
 | |
| 171 |     ("\\<lfloor>", (1, "⌊")),
 | |
| 172 |     ("\\<rfloor>", (1, "⌋")),
 | |
| 173 |     ("\\<langle>", (1, "⟨")),
 | |
| 174 |     ("\\<rangle>", (1, "⟩")),
 | |
| 175 |     ("\\<lozenge>", (1, "◊")),
 | |
| 176 |     ("\\<spadesuit>", (1, "♠")),
 | |
| 177 |     ("\\<clubsuit>", (1, "♣")),
 | |
| 178 |     ("\\<heartsuit>", (1, "♥")),
 | |
| 179 |     ("\\<diamondsuit>", (1, "♦")),
 | |
| 180 |     ("\\<lbrakk>", (2, "[|")),
 | |
| 181 |     ("\\<rbrakk>", (2, "|]")),
 | |
| 182 |     ("\\<Longrightarrow>", (3, "==>")),
 | |
| 183 |     ("\\<Rightarrow>", (2, "=>")),
 | |
| 184 |     ("\\<And>", (2, "!!")),
 | |
| 185 |     ("\\<Colon>", (2, "::")),
 | |
| 186 |     ("\\<lparr>", (2, "(|")),
 | |
| 187 |     ("\\<rparr>", (2, "|)),")),
 | |
| 188 |     ("\\<longleftrightarrow>", (3, "<->")),
 | |
| 189 |     ("\\<longrightarrow>", (3, "-->")),
 | |
| 190 |     ("\\<rightarrow>", (2, "->")),
 | |
| 55033 | 191 |     ("\\<open>", (1, "‹")),
 | 
| 192 |     ("\\<close>", (1, "›")),
 | |
| 55041 | 193 |     ("\\<newline>", (1, "⏎")),
 | 
| 43592 
e67d104c0c50
HTML: render control symbols more like Isabelle/Scala/jEdit;
 wenzelm parents: 
43548diff
changeset | 194 |     ("\\<^bsub>", (0, hidden "⇘" ^ "<sub>")),
 | 
| 
e67d104c0c50
HTML: render control symbols more like Isabelle/Scala/jEdit;
 wenzelm parents: 
43548diff
changeset | 195 |     ("\\<^esub>", (0, hidden "⇙" ^ "</sub>")),
 | 
| 
e67d104c0c50
HTML: render control symbols more like Isabelle/Scala/jEdit;
 wenzelm parents: 
43548diff
changeset | 196 |     ("\\<^bsup>", (0, hidden "⇗" ^ "<sup>")),
 | 
| 
e67d104c0c50
HTML: render control symbols more like Isabelle/Scala/jEdit;
 wenzelm parents: 
43548diff
changeset | 197 |     ("\\<^esup>", (0, hidden "⇖" ^ "</sup>"))];
 | 
| 17178 | 198 | |
| 16196 | 199 | fun output_sym s = | 
| 23622 | 200 | if Symbol.is_raw s then (1, Symbol.decode_raw s) | 
| 20742 | 201 | else | 
| 33983 
cfbf1ff6170d
output "'" as "'" which is a bit more portable ("'" is defined in XML/XHTML, but not in old-style HTML4);
 wenzelm parents: 
33222diff
changeset | 202 | (case Symtab.lookup html_syms s of | 
| 
cfbf1ff6170d
output "'" as "'" which is a bit more portable ("'" is defined in XML/XHTML, but not in old-style HTML4);
 wenzelm parents: 
33222diff
changeset | 203 | SOME x => x | 
| 26709 | 204 | | NONE => (size s, XML.text s)); | 
| 6324 | 205 | |
| 33985 | 206 | fun output_markup (bg, en) s1 s2 = | 
| 207 | let val (n, txt) = output_sym s2 | |
| 208 | in (n, hidden s1 ^ enclose bg en txt) end; | |
| 14534 | 209 | |
| 33985 | 210 |   val output_sub = output_markup ("<sub>", "</sub>");
 | 
| 211 |   val output_sup = output_markup ("<sup>", "</sup>");
 | |
| 212 | val output_bold = output_markup (span "bold"); | |
| 213 | ||
| 41482 | 214 | fun output_syms [] (result, width) = (implode (rev result), width) | 
| 215 | | output_syms (s1 :: rest) (result, width) = | |
| 216 | let | |
| 217 |           val (s2, ss) = (case rest of [] => ("", []) | s2 :: ss => (s2, ss));
 | |
| 218 | val ((w, s), r) = | |
| 43592 
e67d104c0c50
HTML: render control symbols more like Isabelle/Scala/jEdit;
 wenzelm parents: 
43548diff
changeset | 219 | if s1 = "\\<^sub>" then (output_sub "⇩" s2, ss) | 
| 
e67d104c0c50
HTML: render control symbols more like Isabelle/Scala/jEdit;
 wenzelm parents: 
43548diff
changeset | 220 | else if s1 = "\\<^sup>" then (output_sup "⇧" s2, ss) | 
| 
e67d104c0c50
HTML: render control symbols more like Isabelle/Scala/jEdit;
 wenzelm parents: 
43548diff
changeset | 221 | else if s1 = "\\<^bold>" then (output_bold "❙" s2, ss) | 
| 41482 | 222 | else (output_sym s1, rest); | 
| 223 | in output_syms r (s :: result, width + w) end; | |
| 6338 | 224 | in | 
| 6324 | 225 | |
| 41482 | 226 | fun output_width str = output_syms (Symbol.explode str) ([], 0); | 
| 6338 | 227 | val output = #1 o output_width; | 
| 228 | ||
| 26709 | 229 | val _ = Output.add_mode htmlN output_width Symbol.encode_raw; | 
| 230 | ||
| 6338 | 231 | end; | 
| 232 | ||
| 6324 | 233 | |
| 234 | ||
| 235 | (** HTML markup **) | |
| 236 | ||
| 237 | type text = string; | |
| 238 | ||
| 239 | ||
| 240 | (* atoms *) | |
| 241 | ||
| 242 | val plain = output; | |
| 19265 | 243 | val name = enclose "<span class=\"name\">" "</span>" o output; | 
| 244 | val keyword = enclose "<span class=\"keyword\">" "</span>" o output; | |
| 23622 | 245 | val command = enclose "<span class=\"command\">" "</span>" o output; | 
| 6324 | 246 | |
| 247 | ||
| 248 | (* misc *) | |
| 249 | ||
| 28840 | 250 | fun href_name s txt = "<a href=" ^ quote s ^ ">" ^ txt ^ "</a>"; | 
| 21858 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
20742diff
changeset | 251 | fun href_path path txt = href_name (Url.implode path) txt; | 
| 6376 | 252 | |
| 15531 | 253 | fun href_opt_path NONE txt = txt | 
| 254 | | href_opt_path (SOME p) txt = href_path p txt; | |
| 6376 | 255 | |
| 12413 | 256 | fun para txt = "\n<p>" ^ txt ^ "</p>\n"; | 
| 27830 | 257 | fun preform txt = "<pre>" ^ txt ^ "</pre>"; | 
| 6324 | 258 | val verbatim = preform o output; | 
| 259 | ||
| 260 | ||
| 261 | (* document *) | |
| 262 | ||
| 263 | fun begin_document title = | |
| 40946 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 264 | "<?xml version=\"1.0\" encoding=\"utf-8\" ?>\n\ | 
| 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 265 | \<!DOCTYPE html PUBLIC \"-//W3C//DTD XHTML 1.0 Transitional//EN\" \ | 
| 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 266 | \\"http://www.w3.org/TR/xhtml1/DTD/xhtml1-transitional.dtd\">\n\ | 
| 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 267 | \<html xmlns=\"http://www.w3.org/1999/xhtml\">\n\ | 
| 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 268 | \<head>\n\ | 
| 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 269 | \<meta http-equiv=\"Content-Type\" content=\"text/html; charset=utf-8\"/>\n\ | 
| 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 270 |   \<title>" ^ plain (title ^ " (" ^ Distribution.version ^ ")") ^ "</title>\n\
 | 
| 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 271 | \<link media=\"all\" rel=\"stylesheet\" type=\"text/css\" href=\"isabelle.css\"/>\n\ | 
| 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 272 | \</head>\n\ | 
| 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 273 | \\n\ | 
| 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 274 | \<body>\n\ | 
| 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 275 | \<div class=\"head\">\ | 
| 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 276 | \<h1>" ^ plain title ^ "</h1>\n"; | 
| 6324 | 277 | |
| 14541 | 278 | val end_document = "\n</div>\n</body>\n</html>\n"; | 
| 6324 | 279 | |
| 6338 | 280 | |
| 281 | (* session index *) | |
| 282 | ||
| 51399 
6ac3c29a300e
discontinued "isabelle usedir" option -r (reset session path);
 wenzelm parents: 
50233diff
changeset | 283 | fun begin_session_index session docs graph = | 
| 
6ac3c29a300e
discontinued "isabelle usedir" option -r (reset session path);
 wenzelm parents: 
50233diff
changeset | 284 |   begin_document ("Session " ^ plain session) ^
 | 
| 6754 | 285 |   para ("View " ^ href_path graph "theory dependencies" ^
 | 
| 27829 | 286 | implode (map (fn (p, name) => "<br/>\nView " ^ href_path p name) docs)) ^ | 
| 37941 
1d812ff95a14
refrain from generating <hr/> and from "hiding" it in isabelle.css -- the latter might be used in other situations as well;
 wenzelm parents: 
37940diff
changeset | 287 | "\n</div>\n<div class=\"theories\">\n<h2>Theories</h2>\n<ul>\n"; | 
| 6338 | 288 | |
| 6649 | 289 | fun choice chs s = space_implode " " (map (fn (s', lnk) => | 
| 290 | enclose "[" "]" (if s = s' then keyword s' else href_name lnk s')) chs); | |
| 291 | ||
| 51399 
6ac3c29a300e
discontinued "isabelle usedir" option -r (reset session path);
 wenzelm parents: 
50233diff
changeset | 292 | fun back_link (path, name) = para (href_path path "Back" ^ " to index of " ^ plain name); | 
| 
6ac3c29a300e
discontinued "isabelle usedir" option -r (reset session path);
 wenzelm parents: 
50233diff
changeset | 293 | |
| 9415 | 294 | fun applet_pages session back = | 
| 6649 | 295 | let | 
| 9415 | 296 | val sizes = | 
| 297 |      [("small", "small.html", ("500", "400")),
 | |
| 298 |       ("medium", "medium.html", ("650", "520")),
 | |
| 299 |       ("large", "large.html", ("800", "640"))];
 | |
| 6649 | 300 | |
| 9415 | 301 | fun applet_page (size, name, (width, height)) = | 
| 6754 | 302 | let | 
| 303 | val browser_size = "Set browser size: " ^ | |
| 9415 | 304 | choice (map (fn (y, z, _) => (y, z)) sizes) size; | 
| 6754 | 305 | in | 
| 306 |         (name, begin_document ("Theory dependencies of " ^ session) ^
 | |
| 307 | back_link back ^ | |
| 9415 | 308 | para browser_size ^ | 
| 37941 
1d812ff95a14
refrain from generating <hr/> and from "hiding" it in isabelle.css -- the latter might be used in other situations as well;
 wenzelm parents: 
37940diff
changeset | 309 | "\n</div>\n<div class=\"graphbrowser\">\n\ | 
| 17080 | 310 | \<applet code=\"GraphBrowser/GraphBrowser.class\" \ | 
| 14541 | 311 | \archive=\"GraphBrowser.jar\" \ | 
| 27832 
992c6d984925
<applet>: more XHTML 1.0 Transitional conformance;
 wenzelm parents: 
27830diff
changeset | 312 | \width=" ^ quote width ^ " height=" ^ quote height ^ ">\n\ | 
| 
992c6d984925
<applet>: more XHTML 1.0 Transitional conformance;
 wenzelm parents: 
27830diff
changeset | 313 | \<param name=\"graphfile\" value=\"" ^ "session.graph" ^ "\"/>\n\ | 
| 37941 
1d812ff95a14
refrain from generating <hr/> and from "hiding" it in isabelle.css -- the latter might be used in other situations as well;
 wenzelm parents: 
37940diff
changeset | 314 | \</applet>" ^ end_document) | 
| 6649 | 315 | end; | 
| 9415 | 316 | in map applet_page sizes end; | 
| 6649 | 317 | |
| 318 | ||
| 51399 
6ac3c29a300e
discontinued "isabelle usedir" option -r (reset session path);
 wenzelm parents: 
50233diff
changeset | 319 | fun theory_entry (p, s) = "<li>" ^ href_path p (plain s) ^ "</li>\n"; | 
| 6324 | 320 | |
| 321 | ||
| 322 | (* theory *) | |
| 323 | ||
| 54456 | 324 | fun theory A Bs txt = | 
| 325 |   begin_document ("Theory " ^ A) ^ "\n" ^
 | |
| 326 | command "theory" ^ " " ^ name A ^ "<br/>\n" ^ | |
| 54455 
1d977436c1bf
removed remains of HTML presentation of auxiliary files -- inactive since Isabelle2013;
 wenzelm parents: 
53194diff
changeset | 327 | keyword "imports" ^ " " ^ space_implode " " (map (uncurry href_opt_path o apsnd name) Bs) ^ | 
| 54456 | 328 | "<br/>\n" ^ | 
| 329 | enclose "\n</div>\n<div class=\"source\">\n<pre>" "</pre>\n" txt ^ | |
| 330 | end_document; | |
| 16267 
0773b17922d8
present new-style theory header, with 'imports' and 'uses';
 wenzelm parents: 
16196diff
changeset | 331 | |
| 6324 | 332 | end; |