| author | wenzelm | 
| Fri, 24 Sep 2021 22:23:26 +0200 | |
| changeset 74362 | 0135a0c77b64 | 
| parent 73785 | b5fb99b985b4 | 
| child 74733 | 255e651a4c5f | 
| permissions | -rw-r--r-- | 
| 73718 | 1 | /* Title: Pure/Thy/document_build.scala | 
| 2 | Author: Makarius | |
| 3 | ||
| 4 | Build theory document (PDF) from session database. | |
| 5 | */ | |
| 6 | ||
| 7 | package isabelle | |
| 8 | ||
| 9 | ||
| 10 | object Document_Build | |
| 11 | {
 | |
| 12 | /* document variants */ | |
| 13 | ||
| 73778 | 14 | object Content | 
| 15 |   {
 | |
| 16 | def apply(path: Path, content: Bytes): Content = new Content_Bytes(path, content) | |
| 17 | def apply(path: Path, content: String): Content = new Content_String(path, content) | |
| 18 | } | |
| 19 | trait Content | |
| 73724 | 20 |   {
 | 
| 73778 | 21 | def write(dir: Path): Unit | 
| 22 | } | |
| 23 | final class Content_Bytes private[Document_Build](path: Path, content: Bytes) extends Content | |
| 24 |   {
 | |
| 25 | def write(dir: Path): Unit = Bytes.write(dir + path, content) | |
| 26 | } | |
| 27 | final class Content_String private[Document_Build](path: Path, content: String) extends Content | |
| 28 |   {
 | |
| 29 | def write(dir: Path): Unit = File.write(dir + path, content) | |
| 73724 | 30 | } | 
| 31 | ||
| 73718 | 32 | abstract class Document_Name | 
| 33 |   {
 | |
| 34 | def name: String | |
| 35 | def path: Path = Path.basic(name) | |
| 36 | ||
| 37 | override def toString: String = name | |
| 38 | } | |
| 39 | ||
| 40 | object Document_Variant | |
| 41 |   {
 | |
| 42 | def parse(name: String, tags: String): Document_Variant = | |
| 43 |       Document_Variant(name, Library.space_explode(',', tags))
 | |
| 44 | ||
| 45 | def parse(opt: String): Document_Variant = | |
| 46 |       Library.space_explode('=', opt) match {
 | |
| 47 | case List(name) => Document_Variant(name, Nil) | |
| 48 | case List(name, tags) => parse(name, tags) | |
| 49 |         case _ => error("Malformed document variant: " + quote(opt))
 | |
| 50 | } | |
| 51 | } | |
| 52 | ||
| 53 | sealed case class Document_Variant(name: String, tags: List[String]) extends Document_Name | |
| 54 |   {
 | |
| 55 |     def print_tags: String = tags.mkString(",")
 | |
| 56 | def print: String = if (tags.isEmpty) name else name + "=" + print_tags | |
| 57 | ||
| 73724 | 58 | def isabelletags: Content = | 
| 59 |     {
 | |
| 60 |       val path = Path.explode("isabelletags.sty")
 | |
| 73778 | 61 | val content = | 
| 73724 | 62 | Library.terminate_lines( | 
| 63 | tags.map(tag => | |
| 64 |             tag.toList match {
 | |
| 65 |               case '/' :: cs => "\\isafoldtag{" + cs.mkString + "}"
 | |
| 66 |               case '-' :: cs => "\\isadroptag{" + cs.mkString + "}"
 | |
| 67 |               case '+' :: cs => "\\isakeeptag{" + cs.mkString + "}"
 | |
| 68 |               case cs => "\\isakeeptag{" + cs.mkString + "}"
 | |
| 69 | })) | |
| 73778 | 70 | Content(path, content) | 
| 73724 | 71 | } | 
| 73718 | 72 | } | 
| 73 | ||
| 74 | sealed case class Document_Input(name: String, sources: SHA1.Digest) | |
| 75 | extends Document_Name | |
| 76 | ||
| 77 | sealed case class Document_Output(name: String, sources: SHA1.Digest, log_xz: Bytes, pdf: Bytes) | |
| 78 | extends Document_Name | |
| 79 |   {
 | |
| 80 | def log: String = log_xz.uncompress().text | |
| 81 | def log_lines: List[String] = split_lines(log) | |
| 82 | ||
| 83 | def write(db: SQL.Database, session_name: String): Unit = | |
| 84 | write_document(db, session_name, this) | |
| 73719 | 85 | |
| 86 | def write(dir: Path): Path = | |
| 87 |     {
 | |
| 88 | val path = dir + Path.basic(name).pdf | |
| 89 | Isabelle_System.make_directory(path.expand.dir) | |
| 90 | Bytes.write(path, pdf) | |
| 91 | path | |
| 92 | } | |
| 73718 | 93 | } | 
| 94 | ||
| 95 | ||
| 96 | /* SQL data model */ | |
| 97 | ||
| 98 | object Data | |
| 99 |   {
 | |
| 100 |     val session_name = SQL.Column.string("session_name").make_primary_key
 | |
| 101 |     val name = SQL.Column.string("name").make_primary_key
 | |
| 102 |     val sources = SQL.Column.string("sources")
 | |
| 103 |     val log_xz = SQL.Column.bytes("log_xz")
 | |
| 104 |     val pdf = SQL.Column.bytes("pdf")
 | |
| 105 | ||
| 106 |     val table = SQL.Table("isabelle_documents", List(session_name, name, sources, log_xz, pdf))
 | |
| 107 | ||
| 108 | def where_equal(session_name: String, name: String = ""): SQL.Source = | |
| 109 | "WHERE " + Data.session_name.equal(session_name) + | |
| 110 | (if (name == "") "" else " AND " + Data.name.equal(name)) | |
| 111 | } | |
| 112 | ||
| 113 | def read_documents(db: SQL.Database, session_name: String): List[Document_Input] = | |
| 114 |   {
 | |
| 115 | val select = Data.table.select(List(Data.name, Data.sources), Data.where_equal(session_name)) | |
| 116 | db.using_statement(select)(stmt => | |
| 117 | stmt.execute_query().iterator(res => | |
| 118 |       {
 | |
| 119 | val name = res.string(Data.name) | |
| 120 | val sources = res.string(Data.sources) | |
| 121 | Document_Input(name, SHA1.fake(sources)) | |
| 122 | }).toList) | |
| 123 | } | |
| 124 | ||
| 125 | def read_document(db: SQL.Database, session_name: String, name: String): Option[Document_Output] = | |
| 126 |   {
 | |
| 127 | val select = Data.table.select(sql = Data.where_equal(session_name, name)) | |
| 128 | db.using_statement(select)(stmt => | |
| 129 |     {
 | |
| 130 | val res = stmt.execute_query() | |
| 131 |       if (res.next()) {
 | |
| 132 | val name = res.string(Data.name) | |
| 133 | val sources = res.string(Data.sources) | |
| 134 | val log_xz = res.bytes(Data.log_xz) | |
| 135 | val pdf = res.bytes(Data.pdf) | |
| 136 | Some(Document_Output(name, SHA1.fake(sources), log_xz, pdf)) | |
| 137 | } | |
| 138 | else None | |
| 139 | }) | |
| 140 | } | |
| 141 | ||
| 142 | def write_document(db: SQL.Database, session_name: String, doc: Document_Output): Unit = | |
| 143 |   {
 | |
| 144 | db.using_statement(Data.table.insert())(stmt => | |
| 145 |     {
 | |
| 146 | stmt.string(1) = session_name | |
| 147 | stmt.string(2) = doc.name | |
| 148 | stmt.string(3) = doc.sources.toString | |
| 149 | stmt.bytes(4) = doc.log_xz | |
| 150 | stmt.bytes(5) = doc.pdf | |
| 151 | stmt.execute() | |
| 152 | }) | |
| 153 | } | |
| 154 | ||
| 155 | ||
| 73719 | 156 | /* context */ | 
| 157 | ||
| 73724 | 158 | val isabelle_styles: List[Path] = | 
| 159 |     List("comment.sty", "isabelle.sty", "isabellesym.sty", "pdfsetup.sty", "railsetup.sty").
 | |
| 160 |       map(name => Path.explode("~~/lib/texinputs") + Path.basic(name))
 | |
| 161 | ||
| 73719 | 162 | def context( | 
| 163 | session: String, | |
| 164 | deps: Sessions.Deps, | |
| 165 | db_context: Sessions.Database_Context, | |
| 166 | progress: Progress = new Progress): Context = | |
| 167 |   {
 | |
| 168 | val info = deps.sessions_structure(session) | |
| 169 | val base = deps(session) | |
| 170 | val hierarchy = deps.sessions_structure.hierarchy(session) | |
| 171 | new Context(info, base, hierarchy, db_context, progress) | |
| 172 | } | |
| 173 | ||
| 174 | final class Context private[Document_Build]( | |
| 175 | info: Sessions.Info, | |
| 176 | base: Sessions.Base, | |
| 177 | hierarchy: List[String], | |
| 178 | db_context: Sessions.Database_Context, | |
| 179 | val progress: Progress = new Progress) | |
| 180 |   {
 | |
| 181 | /* session info */ | |
| 182 | ||
| 183 | def session: String = info.name | |
| 184 | def options: Options = info.options | |
| 185 | ||
| 73743 | 186 |     def document_bibliography: Boolean = options.bool("document_bibliography")
 | 
| 187 | ||
| 73735 | 188 | def document_preprocessor: Option[String] = | 
| 189 |       proper_string(options.string("document_preprocessor"))
 | |
| 190 | ||
| 73723 | 191 | def document_logo: Option[String] = | 
| 192 |       options.string("document_logo") match {
 | |
| 193 | case "" => None | |
| 194 |         case "_" => Some("")
 | |
| 195 | case name => Some(name) | |
| 196 | } | |
| 197 | ||
| 73721 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 198 |     def document_build: String = options.string("document_build")
 | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 199 | |
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 200 | def get_engine(): Engine = | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 201 |     {
 | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 202 | val name = document_build | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 203 |       engines.find(_.name == name).getOrElse(error("Bad document_build engine " + quote(name)))
 | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 204 | } | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 205 | |
| 73719 | 206 | def get_export(theory: String, name: String): Export.Entry = | 
| 207 | db_context.get_export(hierarchy, theory, name) | |
| 208 | ||
| 209 | ||
| 210 | /* document content */ | |
| 211 | ||
| 212 | def documents: List[Document_Variant] = info.documents | |
| 213 | ||
| 214 | def session_theories: List[Document.Node.Name] = base.session_theories | |
| 215 | def document_theories: List[Document.Node.Name] = session_theories ::: base.document_theories | |
| 216 | ||
| 217 | lazy val tex_files: List[Content] = | |
| 218 | for (name <- document_theories) | |
| 219 |       yield {
 | |
| 220 | val path = Path.basic(tex_name(name)) | |
| 73785 | 221 | val xml = YXML.parse_body(get_export(name.theory, Export.DOCUMENT_LATEX).text) | 
| 73780 
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
 wenzelm parents: 
73778diff
changeset | 222 | val content = Latex.output(xml, file_pos = name.path.implode_symbolic) | 
| 73778 | 223 | Content(path, content) | 
| 73719 | 224 | } | 
| 225 | ||
| 226 | lazy val session_graph: Content = | |
| 227 |     {
 | |
| 228 | val path = Presentation.session_graph_path | |
| 73778 | 229 | val content = graphview.Graph_File.make_pdf(options, base.session_graph_display) | 
| 230 | Content(path, content) | |
| 73719 | 231 | } | 
| 232 | ||
| 233 | lazy val session_tex: Content = | |
| 234 |     {
 | |
| 235 |       val path = Path.basic("session.tex")
 | |
| 73778 | 236 | val content = | 
| 73719 | 237 | Library.terminate_lines( | 
| 238 |           base.session_theories.map(name => "\\input{" + tex_name(name) + "}"))
 | |
| 73778 | 239 | Content(path, content) | 
| 73719 | 240 | } | 
| 241 | ||
| 73723 | 242 | lazy val isabelle_logo: Option[Content] = | 
| 243 |     {
 | |
| 244 | document_logo.map(logo_name => | |
| 245 |         Isabelle_System.with_tmp_file("logo", ext = "pdf")(tmp_path =>
 | |
| 246 |         {
 | |
| 247 | Logo.create_logo(logo_name, output_file = tmp_path, quiet = true) | |
| 248 |           val path = Path.basic("isabelle_logo.pdf")
 | |
| 73778 | 249 | val content = Bytes.read(tmp_path) | 
| 250 | Content(path, content) | |
| 73723 | 251 | })) | 
| 252 | } | |
| 253 | ||
| 73719 | 254 | |
| 255 | /* document directory */ | |
| 256 | ||
| 257 | def prepare_directory(dir: Path, doc: Document_Variant): Directory = | |
| 258 |     {
 | |
| 259 | val doc_dir = dir + Path.basic(doc.name) | |
| 260 | Isabelle_System.make_directory(doc_dir) | |
| 261 | ||
| 262 | ||
| 263 | /* sources */ | |
| 264 | ||
| 73724 | 265 | isabelle_styles.foreach(Isabelle_System.copy_file(_, doc_dir)) | 
| 266 | doc.isabelletags.write(doc_dir) | |
| 267 | ||
| 73719 | 268 |       for ((base_dir, src) <- info.document_files) {
 | 
| 269 | Isabelle_System.copy_file_base(info.dir + base_dir, src, doc_dir) | |
| 270 | } | |
| 271 | ||
| 272 | session_tex.write(doc_dir) | |
| 273 | tex_files.foreach(_.write(doc_dir)) | |
| 274 | ||
| 275 | val root_name1 = "root_" + doc.name | |
| 276 | val root_name = if ((doc_dir + Path.explode(root_name1).tex).is_file) root_name1 else "root" | |
| 277 | ||
| 73735 | 278 |       for (name <- document_preprocessor) {
 | 
| 279 | def message(s: String): String = s + " for document_preprocessor=" + quote(name) | |
| 280 | val path = doc_dir + Path.explode(name) | |
| 281 |         if (path.is_file) {
 | |
| 282 |           try { Isabelle_System.bash(File.bash_path(path), cwd = doc_dir.file).check }
 | |
| 283 |           catch { case ERROR(msg) => cat_error(msg, message("The error(s) above occurred")) }
 | |
| 284 | } | |
| 285 |         else error(message("Missing executable"))
 | |
| 286 | } | |
| 287 | ||
| 73723 | 288 | val digests1 = List(doc.print, document_logo.toString, document_build).map(SHA1.digest) | 
| 289 | val digests2 = File.find_files(doc_dir.file, follow_links = true).map(SHA1.digest) | |
| 290 | val sources = SHA1.digest_set(digests1 ::: digests2) | |
| 73719 | 291 | |
| 292 | ||
| 73723 | 293 | /* derived material (without SHA1 digest) */ | 
| 294 | ||
| 295 | isabelle_logo.foreach(_.write(doc_dir)) | |
| 73719 | 296 | |
| 297 | session_graph.write(doc_dir) | |
| 298 | ||
| 299 | Directory(doc_dir, doc, root_name, sources) | |
| 300 | } | |
| 301 | ||
| 302 | def old_document(directory: Directory): Option[Document_Output] = | |
| 303 |       for {
 | |
| 304 | old_doc <- db_context.input_database(session)(read_document(_, _, directory.doc.name)) | |
| 305 | if old_doc.sources == directory.sources | |
| 306 | } | |
| 73720 | 307 | yield old_doc | 
| 73719 | 308 | } | 
| 309 | ||
| 310 | sealed case class Directory( | |
| 311 | doc_dir: Path, | |
| 312 | doc: Document_Variant, | |
| 313 | root_name: String, | |
| 314 | sources: SHA1.Digest) | |
| 315 |   {
 | |
| 73737 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 316 | def root_name_script(ext: String = ""): String = | 
| 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 317 | Bash.string(if (ext.isEmpty) root_name else root_name + "." + ext) | 
| 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 318 | |
| 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 319 | def conditional_script(ext: String, exe: String, after: String = ""): String = | 
| 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 320 | "if [ -f " + root_name_script(ext) + " ]\n" + | 
| 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 321 | "then\n" + | 
| 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 322 | " " + exe + " " + root_name_script() + "\n" + | 
| 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 323 | (if (after.isEmpty) "" else " " + after) + | 
| 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 324 | "fi\n" | 
| 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 325 | |
| 73719 | 326 | def log_errors(): List[String] = | 
| 327 | Latex.latex_errors(doc_dir, root_name) ::: | |
| 328 | Bibtex.bibtex_errors(doc_dir, root_name) | |
| 329 | ||
| 330 | def make_document(log: List[String], errors: List[String]): Document_Output = | |
| 331 |     {
 | |
| 332 | val root_pdf = Path.basic(root_name).pdf | |
| 333 | val result_pdf = doc_dir + root_pdf | |
| 334 | ||
| 335 |       if (errors.nonEmpty) {
 | |
| 336 |         val errors1 = errors ::: List("Failed to build document " + quote(doc.name))
 | |
| 337 | throw new Build_Error(log, Exn.cat_message(errors1: _*)) | |
| 338 | } | |
| 339 |       else if (!result_pdf.is_file) {
 | |
| 340 | val message = "Bad document result: expected to find " + root_pdf | |
| 341 | throw new Build_Error(log, message) | |
| 342 | } | |
| 343 |       else {
 | |
| 344 | val log_xz = Bytes(cat_lines(log)).compress() | |
| 345 | val pdf = Bytes.read(result_pdf) | |
| 346 | Document_Output(doc.name, sources, log_xz, pdf) | |
| 347 | } | |
| 348 | } | |
| 349 | } | |
| 350 | ||
| 73718 | 351 | |
| 73721 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 352 | /* build engines */ | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 353 | |
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 354 | lazy val engines: List[Engine] = Isabelle_System.make_services(classOf[Engine]) | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 355 | |
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 356 | abstract class Engine(val name: String) extends Isabelle_System.Service | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 357 |   {
 | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 358 | override def toString: String = name | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 359 | |
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 360 | def prepare_directory(context: Context, dir: Path, doc: Document_Variant): Directory | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 361 | def build_document(context: Context, directory: Directory, verbose: Boolean): Document_Output | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 362 | } | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 363 | |
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 364 | abstract class Bash_Engine(name: String) extends Engine(name) | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 365 |   {
 | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 366 | def prepare_directory(context: Context, dir: Path, doc: Document_Variant): Directory = | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 367 | context.prepare_directory(dir, doc) | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 368 | |
| 73759 | 369 | def use_pdflatex: Boolean = false | 
| 73737 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 370 | def latex_script(context: Context, directory: Directory): String = | 
| 73759 | 371 | (if (use_pdflatex) "$ISABELLE_PDFLATEX" else "$ISABELLE_LUALATEX") + | 
| 73737 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 372 | " " + directory.root_name_script() + "\n" | 
| 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 373 | |
| 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 374 | def bibtex_script(context: Context, directory: Directory, latex: Boolean = false): String = | 
| 73743 | 375 |     {
 | 
| 376 | val ext = if (context.document_bibliography) "aux" else "bib" | |
| 377 | directory.conditional_script(ext, "$ISABELLE_BIBTEX", | |
| 73737 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 378 | after = if (latex) latex_script(context, directory) else "") | 
| 73743 | 379 | } | 
| 73737 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 380 | |
| 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 381 | def makeindex_script(context: Context, directory: Directory, latex: Boolean = false): String = | 
| 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 382 |       directory.conditional_script("idx", "$ISABELLE_MAKEINDEX",
 | 
| 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 383 | after = if (latex) latex_script(context, directory) else "") | 
| 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 384 | |
| 73759 | 385 | def use_build_script: Boolean = false | 
| 73737 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 386 | def build_script(context: Context, directory: Directory): String = | 
| 73721 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 387 |     {
 | 
| 73759 | 388 |       val has_build_script = (directory.doc_dir + Path.explode("build")).is_file
 | 
| 73721 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 389 | |
| 73759 | 390 |       if (!use_build_script && has_build_script) {
 | 
| 73721 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 391 |         error("Unexpected document build script for option document_build=" +
 | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 392 | quote(context.document_build)) | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 393 | } | 
| 73759 | 394 |       else if (use_build_script && !has_build_script) error("Missing document build script")
 | 
| 395 | else if (has_build_script) "./build pdf " + Bash.string(directory.doc.name) | |
| 73721 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 396 |       else {
 | 
| 73737 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 397 | "set -e\n" + | 
| 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 398 | latex_script(context, directory) + | 
| 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 399 | bibtex_script(context, directory, latex = true) + | 
| 73738 
d701bd96e323
more robust: allow \printindex within the document;
 wenzelm parents: 
73737diff
changeset | 400 | makeindex_script(context, directory) + | 
| 73737 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 401 | latex_script(context, directory) + | 
| 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 402 | makeindex_script(context, directory, latex = true) | 
| 73721 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 403 | } | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 404 | } | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 405 | |
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 406 | def build_document(context: Context, directory: Directory, verbose: Boolean): Document_Output = | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 407 |     {
 | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 408 | val result = | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 409 | context.progress.bash( | 
| 73737 
6638323d2774
clarified bash scripts, with public interfaces for user-defined Document_Build.Engine;
 wenzelm parents: 
73735diff
changeset | 410 | build_script(context, directory), | 
| 73721 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 411 | cwd = directory.doc_dir.file, | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 412 | echo = verbose, | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 413 | watchdog = Time.seconds(0.5)) | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 414 | |
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 415 | val log = result.out_lines ::: result.err_lines | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 416 | val errors = (if (result.ok) Nil else List(result.err)) ::: directory.log_errors() | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 417 | directory.make_document(log, errors) | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 418 | } | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 419 | } | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 420 | |
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 421 |   class LuaLaTeX_Engine extends Bash_Engine("lualatex")
 | 
| 73759 | 422 |   class PDFLaTeX_Engine extends Bash_Engine("pdflatex") { override def use_pdflatex: Boolean = true }
 | 
| 423 |   class Build_Engine extends Bash_Engine("build") { override def use_build_script: Boolean = true }
 | |
| 73721 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 424 | |
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 425 | |
| 73718 | 426 | /* build documents */ | 
| 427 | ||
| 428 | def tex_name(name: Document.Node.Name): String = name.theory_base_name + ".tex" | |
| 429 | ||
| 430 | class Build_Error(val log_lines: List[String], val message: String) | |
| 431 | extends Exn.User_Error(message) | |
| 432 | ||
| 433 | def build_documents( | |
| 73719 | 434 | context: Context, | 
| 73718 | 435 | output_sources: Option[Path] = None, | 
| 436 | output_pdf: Option[Path] = None, | |
| 73721 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 437 | verbose: Boolean = false): List[Document_Output] = | 
| 73718 | 438 |   {
 | 
| 73719 | 439 | val progress = context.progress | 
| 73721 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 440 | val engine = context.get_engine() | 
| 73718 | 441 | |
| 442 | val documents = | |
| 73719 | 443 | for (doc <- context.documents) | 
| 73718 | 444 |       yield {
 | 
| 445 |         Isabelle_System.with_tmp_dir("document")(tmp_dir =>
 | |
| 446 |         {
 | |
| 73719 | 447 |           progress.echo("Preparing " + context.session + "/" + doc.name + " ...")
 | 
| 73718 | 448 | val start = Time.now() | 
| 449 | ||
| 73721 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 450 | output_sources.foreach(engine.prepare_directory(context, _, doc)) | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 451 | val directory = engine.prepare_directory(context, tmp_dir, doc) | 
| 73718 | 452 | |
| 73721 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 453 | val document = | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 454 | context.old_document(directory) getOrElse | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 455 | engine.build_document(context, directory, verbose) | 
| 73718 | 456 | |
| 73721 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 457 | val stop = Time.now() | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 458 | val timing = stop - start | 
| 73718 | 459 | |
| 73721 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 460 |           progress.echo("Finished " + context.session + "/" + doc.name +
 | 
| 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 461 |             " (" + timing.message_hms + " elapsed time)")
 | 
| 73718 | 462 | |
| 73721 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 463 | document | 
| 73718 | 464 | }) | 
| 465 | } | |
| 466 | ||
| 467 |     for (dir <- output_pdf; doc <- documents) {
 | |
| 73719 | 468 | val path = doc.write(dir) | 
| 73718 | 469 |       progress.echo("Document at " + path.absolute)
 | 
| 470 | } | |
| 471 | ||
| 472 | documents | |
| 473 | } | |
| 474 | ||
| 475 | ||
| 476 | /* Isabelle tool wrapper */ | |
| 477 | ||
| 478 | val isabelle_tool = | |
| 479 |     Isabelle_Tool("document", "prepare session theory document", Scala_Project.here, args =>
 | |
| 480 |     {
 | |
| 481 | var output_sources: Option[Path] = None | |
| 482 | var output_pdf: Option[Path] = None | |
| 483 | var verbose_latex = false | |
| 484 | var dirs: List[Path] = Nil | |
| 485 | var options = Options.init() | |
| 486 | var verbose_build = false | |
| 487 | ||
| 488 | val getopts = Getopts( | |
| 489 | """ | |
| 490 | Usage: isabelle document [OPTIONS] SESSION | |
| 491 | ||
| 492 | Options are: | |
| 493 | -O DIR output directory for LaTeX sources and resulting PDF | |
| 494 | -P DIR output directory for resulting PDF | |
| 495 | -S DIR output directory for LaTeX sources | |
| 496 | -V verbose latex | |
| 497 | -d DIR include session directory | |
| 498 | -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) | |
| 499 | -v verbose build | |
| 500 | ||
| 501 | Prepare the theory document of a session. | |
| 502 | """, | |
| 503 | "O:" -> (arg => | |
| 504 |           {
 | |
| 505 | val dir = Path.explode(arg) | |
| 506 | output_sources = Some(dir) | |
| 507 | output_pdf = Some(dir) | |
| 508 | }), | |
| 509 |         "P:" -> (arg => { output_pdf = Some(Path.explode(arg)) }),
 | |
| 510 |         "S:" -> (arg => { output_sources = Some(Path.explode(arg)) }),
 | |
| 511 | "V" -> (_ => verbose_latex = true), | |
| 512 | "d:" -> (arg => dirs = dirs ::: List(Path.explode(arg))), | |
| 513 | "o:" -> (arg => options = options + arg), | |
| 514 | "v" -> (_ => verbose_build = true)) | |
| 515 | ||
| 516 | val more_args = getopts(args) | |
| 517 | val session = | |
| 518 |         more_args match {
 | |
| 519 | case List(a) => a | |
| 520 | case _ => getopts.usage() | |
| 521 | } | |
| 522 | ||
| 523 | val progress = new Console_Progress(verbose = verbose_build) | |
| 524 | val store = Sessions.store(options) | |
| 525 | ||
| 526 |       progress.interrupt_handler {
 | |
| 527 | val res = | |
| 528 | Build.build(options, selection = Sessions.Selection.session(session), | |
| 529 | dirs = dirs, progress = progress, verbose = verbose_build) | |
| 530 |         if (!res.ok) error("Failed to build session " + quote(session))
 | |
| 531 | ||
| 532 | val deps = | |
| 533 | Sessions.load_structure(options + "document=pdf", dirs = dirs). | |
| 534 | selection_deps(Sessions.Selection.session(session)) | |
| 535 | ||
| 536 |         if (output_sources.isEmpty && output_pdf.isEmpty) {
 | |
| 537 |           progress.echo_warning("No output directory")
 | |
| 538 | } | |
| 539 | ||
| 540 | using(store.open_database_context())(db_context => | |
| 73719 | 541 |         {
 | 
| 542 | build_documents(context(session, deps, db_context, progress = progress), | |
| 73718 | 543 | output_sources = output_sources, output_pdf = output_pdf, | 
| 73721 
52030acb19ac
option document_build refers to build engine in Isabelle/Scala;
 wenzelm parents: 
73720diff
changeset | 544 | verbose = verbose_latex) | 
| 73719 | 545 | }) | 
| 73718 | 546 | } | 
| 547 | }) | |
| 548 | } |