diff --git a/src/Pure/Thy/document_build.scala b/src/Pure/Thy/document_build.scala --- a/src/Pure/Thy/document_build.scala +++ b/src/Pure/Thy/document_build.scala @@ -1,565 +1,581 @@ /* Title: Pure/Thy/document_build.scala Author: Makarius Build theory document (PDF) from session database. */ package isabelle object Document_Build { /* document variants */ abstract class Document_Name { def name: String def path: Path = Path.basic(name) override def toString: String = name } object Document_Variant { def parse(opt: String): Document_Variant = Library.space_explode('=', opt) match { case List(name) => Document_Variant(name, Latex.Tags.empty) case List(name, tags) => Document_Variant(name, Latex.Tags(tags)) case _ => error("Malformed document variant: " + quote(opt)) } } sealed case class Document_Variant(name: String, tags: Latex.Tags) extends Document_Name { def print: String = if (tags.toString.isEmpty) name else name + "=" + tags.toString } sealed case class Document_Input(name: String, sources: SHA1.Digest) extends Document_Name { override def toString: String = name } sealed case class Document_Output(name: String, sources: SHA1.Digest, log_xz: Bytes, pdf: Bytes) extends Document_Name { override def toString: String = name def log: String = log_xz.uncompress().text def log_lines: List[String] = split_lines(log) def write(db: SQL.Database, session_name: String): Unit = write_document(db, session_name, this) def write(dir: Path): Path = { val path = dir + Path.basic(name).pdf Isabelle_System.make_directory(path.expand.dir) Bytes.write(path, pdf) path } } /* SQL data model */ object Data { val session_name = SQL.Column.string("session_name").make_primary_key val name = SQL.Column.string("name").make_primary_key val sources = SQL.Column.string("sources") val log_xz = SQL.Column.bytes("log_xz") val pdf = SQL.Column.bytes("pdf") val table = SQL.Table("isabelle_documents", List(session_name, name, sources, log_xz, pdf)) def where_equal(session_name: String, name: String = ""): SQL.Source = "WHERE " + Data.session_name.equal(session_name) + (if (name == "") "" else " AND " + Data.name.equal(name)) } def read_documents(db: SQL.Database, session_name: String): List[Document_Input] = { val select = Data.table.select(List(Data.name, Data.sources), Data.where_equal(session_name)) db.using_statement(select)(stmt => stmt.execute_query().iterator({ res => val name = res.string(Data.name) val sources = res.string(Data.sources) Document_Input(name, SHA1.fake_digest(sources)) }).toList) } def read_document( db: SQL.Database, session_name: String, name: String ): Option[Document_Output] = { val select = Data.table.select(sql = Data.where_equal(session_name, name)) db.using_statement(select)({ stmt => val res = stmt.execute_query() if (res.next()) { val name = res.string(Data.name) val sources = res.string(Data.sources) val log_xz = res.bytes(Data.log_xz) val pdf = res.bytes(Data.pdf) Some(Document_Output(name, SHA1.fake_digest(sources), log_xz, pdf)) } else None }) } def write_document(db: SQL.Database, session_name: String, doc: Document_Output): Unit = { db.using_statement(Data.table.insert()){ stmt => stmt.string(1) = session_name stmt.string(2) = doc.name stmt.string(3) = doc.sources.toString stmt.bytes(4) = doc.log_xz stmt.bytes(5) = doc.pdf stmt.execute() } } /* background context */ def session_background( options: Options, session: String, dirs: List[Path] = Nil, progress: Progress = new Progress, verbose: Boolean = false ): Sessions.Background = { Sessions.load_structure(options + "document", dirs = dirs). selection_deps(Sessions.Selection.session(session), progress = progress, verbose = verbose). background(session) } /* document context */ val texinputs: Path = Path.explode("~~/lib/texinputs") val isabelle_styles: List[Path] = List("isabelle.sty", "isabellesym.sty", "pdfsetup.sty", "railsetup.sty"). map(name => texinputs + Path.basic(name)) def running_script(title: String): String = "echo " + Bash.string("Running program \"" + title + "\" ...") + ";" def detect_running_script(s: String): Option[String] = for { s1 <- Library.try_unprefix("Running program \"", s) s2 <- Library.try_unsuffix("\" ...", s1) } yield s2 - sealed case class Document_Latex(name: Document.Node.Name, body: XML.Body) { + sealed case class Document_Latex( + name: Document.Node.Name, + body: XML.Body, + line_pos: Properties.T => Option[Int] + ) { def content: File.Content_XML = File.content(Path.basic(tex_name(name)), body) def file_pos: String = File.symbolic_path(name.path) def write(latex_output: Latex.Output, dir: Path): Unit = - content.output(latex_output.make(_, file_pos = file_pos)).write(dir) + content.output(latex_output.make(_, file_pos = file_pos, line_pos = line_pos)) + .write(dir) } def context( session_context: Export.Session_Context, document_session: Option[Sessions.Base] = None, document_selection: Document.Node.Name => Boolean = _ => true, progress: Progress = new Progress ): Context = new Context(session_context, document_session, document_selection, progress) final class Context private[Document_Build]( val session_context: Export.Session_Context, document_session: Option[Sessions.Base], document_selection: Document.Node.Name => Boolean, val progress: Progress ) { context => /* session info */ private val base = document_session getOrElse session_context.session_base private val info = session_context.sessions_structure(base.session_name) def session: String = info.name def options: Options = info.options override def toString: String = session val classpath: List[File.Content] = session_context.classpath() def document_bibliography: Boolean = options.bool("document_bibliography") def document_logo: Option[String] = options.string("document_logo") match { case "" => None case "_" => Some("") case name => Some(name) } def document_build: String = options.string("document_build") def get_engine(): Engine = { val name = document_build Classpath(jar_contents = classpath).make_services(classOf[Engine]) .find(_.name == name).getOrElse(error("Bad document_build engine " + quote(name))) } /* document content */ def documents: List[Document_Variant] = info.documents def session_document_theories: List[Document.Node.Name] = base.proper_session_theories def all_document_theories: List[Document.Node.Name] = base.all_document_theories lazy val isabelle_logo: Option[File.Content] = { document_logo.map(logo_name => Isabelle_System.with_tmp_file("logo", ext = "pdf") { tmp_path => Logo.create_logo(logo_name, output_file = tmp_path, quiet = true) val path = Path.basic("isabelle_logo.pdf") val content = Bytes.read(tmp_path) File.content(path, content) }) } lazy val session_graph: File.Content = { val path = Browser_Info.session_graph_path val content = graphview.Graph_File.make_pdf(options, base.session_graph_display) File.content(path, content) } lazy val session_tex: File.Content = { val path = Path.basic("session.tex") val content = Library.terminate_lines( session_document_theories.map(name => "\\input{" + tex_name(name) + "}")) File.content(path, content) } lazy val document_latex: List[Document_Latex] = for (name <- all_document_theories) yield { val body = if (document_selection(name)) { val entry = session_context(name.theory, Export.DOCUMENT_LATEX, permissive = true) YXML.parse_body(entry.text) } else Nil - Document_Latex(name, body) + + def line_pos(props: Properties.T): Option[Int] = + Position.Line.unapply(props) orElse { + for { + snapshot <- session_context.document_snapshot + id <- Position.Id.unapply(props) + offset <- Position.Offset.unapply(props) + pos <- snapshot.find_command_position(id, offset) + } yield pos.line1 + } + + Document_Latex(name, body, line_pos) } /* build document */ def build_document(doc: Document_Variant, verbose: Boolean = false): Document_Output = { Isabelle_System.with_tmp_dir("document") { tmp_dir => val engine = get_engine() val directory = engine.prepare_directory(context, tmp_dir, doc) engine.build_document(context, directory, verbose) } } /* document directory */ def make_directory(dir: Path, doc: Document_Variant): Path = Isabelle_System.make_directory(dir + Path.basic(doc.name)) def prepare_directory( dir: Path, doc: Document_Variant, latex_output: Latex.Output ): Directory = { val doc_dir = make_directory(dir, doc) /* actual sources: with SHA1 digest */ isabelle_styles.foreach(Latex.copy_file(_, doc_dir)) val comment_latex = latex_output.options.bool("document_comment_latex") if (!comment_latex) Latex.copy_file(texinputs + Path.basic("comment.sty"), doc_dir) doc.tags.sty(comment_latex).write(doc_dir) for ((base_dir, src) <- info.document_files) { Latex.copy_file_base(info.dir + base_dir, src, doc_dir) } session_tex.write(doc_dir) document_latex.foreach(_.write(latex_output, doc_dir)) val root_name1 = "root_" + doc.name val root_name = if ((doc_dir + Path.explode(root_name1).tex).is_file) root_name1 else "root" val digests1 = List(doc.print, document_logo.toString, document_build).map(SHA1.digest) val digests2 = File.find_files(doc_dir.file, follow_links = true).map(SHA1.digest) val sources = SHA1.digest_set(digests1 ::: digests2) /* derived material: without SHA1 digest */ isabelle_logo.foreach(_.write(doc_dir)) session_graph.write(doc_dir) Directory(doc_dir, doc, root_name, sources) } def old_document(directory: Directory): Option[Document_Output] = for { db <- session_context.session_db() old_doc <- read_document(db, session, directory.doc.name) if old_doc.sources == directory.sources } yield old_doc } sealed case class Directory( doc_dir: Path, doc: Document_Variant, root_name: String, sources: SHA1.Digest ) { def root_name_script(ext: String = ""): String = Bash.string(if (ext.isEmpty) root_name else root_name + "." + ext) def conditional_script( ext: String, exe: String, title: String = "", after: String = "" ): String = { "if [ -f " + root_name_script(ext) + " ]\n" + "then\n" + " " + (if (title.nonEmpty) running_script(title) else "") + exe + " " + root_name_script() + "\n" + (if (after.isEmpty) "" else " " + after) + "fi\n" } def log_errors(): List[String] = Latex.latex_errors(doc_dir, root_name) ::: Bibtex.bibtex_errors(doc_dir, root_name) def make_document(log: List[String], errors: List[String]): Document_Output = { val root_pdf = Path.basic(root_name).pdf val result_pdf = doc_dir + root_pdf if (errors.nonEmpty) { val message = "Failed to build document " + quote(doc.name) throw new Build_Error(log, errors ::: List(message)) } else if (!result_pdf.is_file) { val message = "Bad document result: expected to find " + root_pdf throw new Build_Error(log, List(message)) } else { val log_xz = Bytes(cat_lines(log)).compress() val pdf = Bytes.read(result_pdf) Document_Output(doc.name, sources, log_xz, pdf) } } } /* build engines */ abstract class Engine(val name: String) extends Isabelle_System.Service { override def toString: String = name def prepare_directory(context: Context, dir: Path, doc: Document_Variant): Directory def build_document(context: Context, directory: Directory, verbose: Boolean): Document_Output } abstract class Bash_Engine(name: String) extends Engine(name) { def prepare_directory(context: Context, dir: Path, doc: Document_Variant): Directory = context.prepare_directory(dir, doc, new Latex.Output(context.options)) def use_pdflatex: Boolean = false def running_latex: String = running_script(if (use_pdflatex) "pdflatex" else "lualatex") def latex_script(context: Context, directory: Directory): String = running_latex + (if (use_pdflatex) "$ISABELLE_PDFLATEX" else "$ISABELLE_LUALATEX") + " " + directory.root_name_script() + "\n" def bibtex_script(context: Context, directory: Directory, latex: Boolean = false): String = { val ext = if (context.document_bibliography) "aux" else "bib" directory.conditional_script(ext, "$ISABELLE_BIBTEX", title = "bibtex", after = if (latex) latex_script(context, directory) else "") } def makeindex_script(context: Context, directory: Directory, latex: Boolean = false): String = directory.conditional_script("idx", "$ISABELLE_MAKEINDEX", title = "makeindex", after = if (latex) latex_script(context, directory) else "") def use_build_script: Boolean = false def build_script(context: Context, directory: Directory): String = { val has_build_script = (directory.doc_dir + Path.explode("build")).is_file if (!use_build_script && has_build_script) { error("Unexpected document build script for option document_build=" + quote(context.document_build)) } else if (use_build_script && !has_build_script) error("Missing document build script") else if (has_build_script) "./build pdf " + Bash.string(directory.doc.name) else { "set -e\n" + latex_script(context, directory) + bibtex_script(context, directory, latex = true) + makeindex_script(context, directory) + latex_script(context, directory) + makeindex_script(context, directory, latex = true) } } def build_document( context: Context, directory: Directory, verbose: Boolean ): Document_Output = { val result = context.progress.bash( build_script(context, directory), cwd = directory.doc_dir.file, echo = verbose, watchdog = Time.seconds(0.5)) val log = result.out_lines ::: result.err_lines val err = result.err val errors1 = directory.log_errors() val errors2 = if (result.ok) errors1 else if (err.nonEmpty) err :: errors1 else if (errors1.nonEmpty) errors1 else List("Error") directory.make_document(log, errors2) } } class LuaLaTeX_Engine extends Bash_Engine("lualatex") class PDFLaTeX_Engine extends Bash_Engine("pdflatex") { override def use_pdflatex: Boolean = true } class Build_Engine extends Bash_Engine("build") { override def use_build_script: Boolean = true } class LIPIcs_Engine(name: String) extends Bash_Engine(name) { def lipics_options(options: Options): Options = options + "document_heading_prefix=" + "document_comment_latex" override def use_pdflatex: Boolean = true override def prepare_directory(context: Context, dir: Path, doc: Document_Variant): Directory = { val doc_dir = context.make_directory(dir, doc) Build_LIPIcs.document_files.foreach(Latex.copy_file(_, doc_dir)) val latex_output = new Latex.Output(lipics_options(context.options)) context.prepare_directory(dir, doc, latex_output) } } class LIPIcs_LuaLaTeX_Engine extends LIPIcs_Engine("lipics") class LIPIcs_PDFLaTeX_Engine extends LIPIcs_Engine("lipics_pdflatex") { override def use_pdflatex: Boolean = true } /* build documents */ def tex_name(name: Document.Node.Name): String = name.theory_base_name + ".tex" class Build_Error(val log_lines: List[String], val log_errors: List[String]) extends Exn.User_Error(Exn.cat_message(log_errors: _*)) def build_documents( context: Context, output_sources: Option[Path] = None, output_pdf: Option[Path] = None, verbose: Boolean = false ): List[Document_Output] = { val progress = context.progress val engine = context.get_engine() val documents = for (doc <- context.documents) yield { Isabelle_System.with_tmp_dir("document") { tmp_dir => progress.echo("Preparing " + context.session + "/" + doc.name + " ...") val start = Time.now() output_sources.foreach(engine.prepare_directory(context, _, doc)) val directory = engine.prepare_directory(context, tmp_dir, doc) val document = context.old_document(directory) getOrElse engine.build_document(context, directory, verbose) val stop = Time.now() val timing = stop - start progress.echo("Finished " + context.session + "/" + doc.name + " (" + timing.message_hms + " elapsed time)") document } } for (dir <- output_pdf; doc <- documents) { val path = doc.write(dir) progress.echo("Document at " + path.absolute) } documents } /* Isabelle tool wrapper */ val isabelle_tool = Isabelle_Tool("document", "prepare session theory document", Scala_Project.here, { args => var output_sources: Option[Path] = None var output_pdf: Option[Path] = None var verbose_latex = false var dirs: List[Path] = Nil var options = Options.init() var verbose_build = false val getopts = Getopts(""" Usage: isabelle document [OPTIONS] SESSION Options are: -O DIR output directory for LaTeX sources and resulting PDF -P DIR output directory for resulting PDF -S DIR output directory for LaTeX sources -V verbose latex -d DIR include session directory -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) -v verbose build Prepare the theory document of a session. """, "O:" -> (arg => { val dir = Path.explode(arg) output_sources = Some(dir) output_pdf = Some(dir) }), "P:" -> (arg => { output_pdf = Some(Path.explode(arg)) }), "S:" -> (arg => { output_sources = Some(Path.explode(arg)) }), "V" -> (_ => verbose_latex = true), "d:" -> (arg => dirs = dirs ::: List(Path.explode(arg))), "o:" -> (arg => options = options + arg), "v" -> (_ => verbose_build = true)) val more_args = getopts(args) val session = more_args match { case List(a) => a case _ => getopts.usage() } val progress = new Console_Progress(verbose = verbose_build) progress.interrupt_handler { val build_results = Build.build(options, selection = Sessions.Selection.session(session), dirs = dirs, progress = progress, verbose = verbose_build) if (!build_results.ok) error("Failed to build session " + quote(session)) if (output_sources.isEmpty && output_pdf.isEmpty) { progress.echo_warning("No output directory") } val session_background = Document_Build.session_background(options, session, dirs = dirs) using(Export.open_session_context(build_results.store, session_background)) { session_context => build_documents( context(session_context, progress = progress), output_sources = output_sources, output_pdf = output_pdf, verbose = verbose_latex) } } }) } diff --git a/src/Pure/Thy/export.scala b/src/Pure/Thy/export.scala --- a/src/Pure/Thy/export.scala +++ b/src/Pure/Thy/export.scala @@ -1,598 +1,598 @@ /* Title: Pure/Thy/export.scala Author: Makarius Manage theory exports: compressed blobs. */ package isabelle import scala.annotation.tailrec import scala.util.matching.Regex object Export { /* artefact names */ val DOCUMENT_ID = "PIDE/document_id" val FILES = "PIDE/files" val MARKUP = "PIDE/markup" val MESSAGES = "PIDE/messages" val DOCUMENT_PREFIX = "document/" val DOCUMENT_LATEX = DOCUMENT_PREFIX + "latex" val DOCUMENT_CITATIONS = DOCUMENT_PREFIX + "citations" val THEORY_PREFIX: String = "theory/" val PROOFS_PREFIX: String = "proofs/" def explode_name(s: String): List[String] = space_explode('/', s) def implode_name(elems: Iterable[String]): String = elems.mkString("/") /* SQL data model */ object Data { val session_name = SQL.Column.string("session_name").make_primary_key val theory_name = SQL.Column.string("theory_name").make_primary_key val name = SQL.Column.string("name").make_primary_key val executable = SQL.Column.bool("executable") val compressed = SQL.Column.bool("compressed") val body = SQL.Column.bytes("body") val table = SQL.Table("isabelle_exports", List(session_name, theory_name, name, executable, compressed, body)) def where_equal(session_name: String, theory_name: String = "", name: String = ""): SQL.Source = "WHERE " + Data.session_name.equal(session_name) + (if (theory_name == "") "" else " AND " + Data.theory_name.equal(theory_name)) + (if (name == "") "" else " AND " + Data.name.equal(name)) } def compound_name(a: String, b: String): String = if (a.isEmpty) b else a + ":" + b sealed case class Entry_Name(session: String = "", theory: String = "", name: String = "") { val compound_name: String = Export.compound_name(theory, name) def make_path(prune: Int = 0): Path = { val elems = theory :: space_explode('/', name) if (elems.length < prune + 1) { error("Cannot prune path by " + prune + " element(s): " + Path.make(elems)) } else Path.make(elems.drop(prune)) } def readable(db: SQL.Database): Boolean = { val select = Data.table.select(List(Data.name), Data.where_equal(session, theory, name)) db.using_statement(select)(stmt => stmt.execute_query().next()) } def read(db: SQL.Database, cache: XML.Cache): Option[Entry] = { val select = Data.table.select(List(Data.executable, Data.compressed, Data.body), Data.where_equal(session, theory, name)) db.using_statement(select) { stmt => val res = stmt.execute_query() if (res.next()) { val executable = res.bool(Data.executable) val compressed = res.bool(Data.compressed) val bytes = res.bytes(Data.body) val body = Future.value(compressed, bytes) Some(Entry(this, executable, body, cache)) } else None } } } def read_theory_names(db: SQL.Database, session_name: String): List[String] = { val select = Data.table.select(List(Data.theory_name), Data.where_equal(session_name), distinct = true) + SQL.order_by(List(Data.theory_name)) db.using_statement(select)(stmt => stmt.execute_query().iterator(_.string(Data.theory_name)).toList) } def read_entry_names(db: SQL.Database, session_name: String): List[Entry_Name] = { val select = Data.table.select(List(Data.theory_name, Data.name), Data.where_equal(session_name)) + SQL.order_by(List(Data.theory_name, Data.name)) db.using_statement(select)(stmt => stmt.execute_query().iterator(res => Entry_Name(session = session_name, theory = res.string(Data.theory_name), name = res.string(Data.name))).toList) } def message(msg: String, theory_name: String, name: String): String = msg + " " + quote(name) + " for theory " + quote(theory_name) object Entry { def apply( entry_name: Entry_Name, executable: Boolean, body: Future[(Boolean, Bytes)], cache: XML.Cache ): Entry = new Entry(entry_name, executable, body, cache) def empty(theory_name: String, name: String): Entry = Entry(Entry_Name(theory = theory_name, name = name), false, Future.value(false, Bytes.empty), XML.Cache.none) def make( session_name: String, args: Protocol.Export.Args, bytes: Bytes, cache: XML.Cache ): Entry = { val body = if (args.compress) Future.fork(bytes.maybe_compress(cache = cache.compress)) else Future.value((false, bytes)) val entry_name = Entry_Name(session = session_name, theory = args.theory_name, name = args.name) Entry(entry_name, args.executable, body, cache) } } final class Entry private( val entry_name: Entry_Name, val executable: Boolean, body: Future[(Boolean, Bytes)], cache: XML.Cache ) { def session_name: String = entry_name.session def theory_name: String = entry_name.theory def name: String = entry_name.name override def toString: String = name def is_finished: Boolean = body.is_finished def cancel(): Unit = body.cancel() def compound_name: String = entry_name.compound_name def name_has_prefix(s: String): Boolean = name.startsWith(s) val name_elems: List[String] = explode_name(name) def name_extends(elems: List[String]): Boolean = name_elems.startsWith(elems) && name_elems != elems def bytes: Bytes = { val (compressed, bs) = body.join if (compressed) bs.uncompress(cache = cache.compress) else bs } def text: String = bytes.text def yxml: XML.Body = YXML.parse_body(UTF8.decode_permissive(bytes), cache = cache) def write(db: SQL.Database): Unit = { val (compressed, bs) = body.join db.using_statement(Data.table.insert()) { stmt => stmt.string(1) = session_name stmt.string(2) = theory_name stmt.string(3) = name stmt.bool(4) = executable stmt.bool(5) = compressed stmt.bytes(6) = bs stmt.execute() } } } def make_regex(pattern: String): Regex = { @tailrec def make(result: List[String], depth: Int, chs: List[Char]): Regex = chs match { case '*' :: '*' :: rest => make("[^:]*" :: result, depth, rest) case '*' :: rest => make("[^:/]*" :: result, depth, rest) case '?' :: rest => make("[^:/]" :: result, depth, rest) case '\\' :: c :: rest => make(("\\" + c) :: result, depth, rest) case '{' :: rest => make("(" :: result, depth + 1, rest) case ',' :: rest if depth > 0 => make("|" :: result, depth, rest) case '}' :: rest if depth > 0 => make(")" :: result, depth - 1, rest) case c :: rest if ".+()".contains(c) => make(("\\" + c) :: result, depth, rest) case c :: rest => make(c.toString :: result, depth, rest) case Nil => result.reverse.mkString.r } make(Nil, 0, pattern.toList) } def make_matcher(pats: List[String]): Entry_Name => Boolean = { val regs = pats.map(make_regex) (entry_name: Entry_Name) => regs.exists(_.matches(entry_name.compound_name)) } /* database consumer thread */ def consumer(db: SQL.Database, cache: XML.Cache, progress: Progress = new Progress): Consumer = new Consumer(db, cache, progress) class Consumer private[Export](db: SQL.Database, cache: XML.Cache, progress: Progress) { private val errors = Synchronized[List[String]](Nil) private val consumer = Consumer_Thread.fork_bulk[(Entry, Boolean)](name = "export")( bulk = { case (entry, _) => entry.is_finished }, consume = { (args: List[(Entry, Boolean)]) => val results = db.transaction { for ((entry, strict) <- args) yield { if (progress.stopped) { entry.cancel() Exn.Res(()) } else if (entry.entry_name.readable(db)) { if (strict) { val msg = message("Duplicate export", entry.theory_name, entry.name) errors.change(msg :: _) } Exn.Res(()) } else Exn.capture { entry.write(db) } } } (results, true) }) def make_entry(session_name: String, args0: Protocol.Export.Args, body: Bytes): Unit = { if (!progress.stopped && !body.is_empty) { val args = if (db.is_server) args0.copy(compress = false) else args0 consumer.send(Entry.make(session_name, args, body, cache) -> args.strict) } } def shutdown(close: Boolean = false): List[String] = { consumer.shutdown() if (close) db.close() errors.value.reverse ::: (if (progress.stopped) List("Export stopped") else Nil) } } /* context for database access */ def open_database_context(store: Sessions.Store): Database_Context = { val database_server = if (store.database_server) Some(store.open_database_server()) else None new Database_Context(store, database_server) } def open_session_context0(store: Sessions.Store, session: String): Session_Context = open_database_context(store).open_session0(session, close_database_context = true) def open_session_context( store: Sessions.Store, session_background: Sessions.Background, document_snapshot: Option[Document.Snapshot] = None ): Session_Context = { open_database_context(store).open_session( session_background, document_snapshot = document_snapshot, close_database_context = true) } class Database_Context private[Export]( val store: Sessions.Store, val database_server: Option[SQL.Database] ) extends AutoCloseable { database_context => override def toString: String = { val s = database_server match { case Some(db) => db.toString case None => "input_dirs = " + store.input_dirs.map(_.absolute).mkString(", ") } "Database_Context(" + s + ")" } def cache: Term.Cache = store.cache def close(): Unit = database_server.foreach(_.close()) def open_database(session: String, output: Boolean = false): Session_Database = database_server match { case Some(db) => new Session_Database(session, db) case None => new Session_Database(session, store.open_database(session, output = output)) { override def close(): Unit = db.close() } } def open_session0(session: String, close_database_context: Boolean = false): Session_Context = open_session(Sessions.background0(session), close_database_context = close_database_context) def open_session( session_background: Sessions.Background, document_snapshot: Option[Document.Snapshot] = None, close_database_context: Boolean = false ): Session_Context = { val session_name = session_background.check_errors.session_name val session_hierarchy = session_background.sessions_structure.build_hierarchy(session_name) val session_databases = database_server match { case Some(db) => session_hierarchy.map(name => new Session_Database(name, db)) case None => val attempts = session_hierarchy.map(name => name -> store.try_open_database(name, server = false)) attempts.collectFirst({ case (name, None) => name }) match { case Some(bad) => for ((_, Some(db)) <- attempts) db.close() store.error_database(bad) case None => for ((name, Some(db)) <- attempts) yield { new Session_Database(name, db) { override def close(): Unit = this.db.close() } } } } new Session_Context( database_context, session_background, session_databases, document_snapshot) { override def close(): Unit = { session_databases.foreach(_.close()) if (close_database_context) database_context.close() } } } } class Session_Database private[Export](val session: String, val db: SQL.Database) extends AutoCloseable { def close(): Unit = () lazy private [Export] val theory_names: List[String] = read_theory_names(db, session) lazy private [Export] val entry_names: List[Entry_Name] = read_entry_names(db, session) } class Session_Context private[Export]( val database_context: Database_Context, session_background: Sessions.Background, db_hierarchy: List[Session_Database], - document_snapshot: Option[Document.Snapshot] + val document_snapshot: Option[Document.Snapshot] ) extends AutoCloseable { session_context => def close(): Unit = () def cache: Term.Cache = database_context.cache def sessions_structure: Sessions.Structure = session_background.sessions_structure def session_base: Sessions.Base = session_background.base def session_name: String = if (document_snapshot.isDefined) Sessions.DRAFT else session_base.session_name def session_database(session: String = session_name): Option[Session_Database] = db_hierarchy.find(_.session == session) def session_db(session: String = session_name): Option[SQL.Database] = session_database(session = session).map(_.db) def session_stack: List[String] = ((if (document_snapshot.isDefined) List(session_name) else Nil) ::: db_hierarchy.map(_.session)).reverse private def select[A]( session: String, select: Session_Database => List[A], project: Entry_Name => A, sort_key: A => String ): List[A] = { def result(name: String): List[A] = if (name == Sessions.DRAFT) { (for { snapshot <- document_snapshot.iterator entry_name <- snapshot.all_exports.keysIterator } yield project(entry_name)).toSet.toList.sortBy(sort_key) } else session_database(name).map(select).getOrElse(Nil) if (session.nonEmpty) result(session) else session_stack.flatMap(result) } def entry_names(session: String = session_name): List[Entry_Name] = select(session, _.entry_names, identity, _.compound_name) def theory_names(session: String = session_name): List[String] = select(session, _.theory_names, _.theory, identity) def get(theory: String, name: String): Option[Entry] = { def snapshot_entry: Option[Entry] = for { snapshot <- document_snapshot entry_name = Entry_Name(session = Sessions.DRAFT, theory = theory, name = name) entry <- snapshot.all_exports.get(entry_name) } yield entry def db_entry: Option[Entry] = db_hierarchy.view.map(database => Export.Entry_Name(session = database.session, theory = theory, name = name) .read(database.db, cache)) .collectFirst({ case Some(entry) => entry }) snapshot_entry orElse db_entry } def apply(theory: String, name: String, permissive: Boolean = false): Entry = get(theory, name) match { case None if permissive => Entry.empty(theory, name) case None => error("Missing export entry " + quote(compound_name(theory, name))) case Some(entry) => entry } def theory(theory: String, other_cache: Option[Term.Cache] = None): Theory_Context = new Theory_Context(session_context, theory, other_cache) def get_source_file(name: String): Option[Sessions.Source_File] = { val store = database_context.store (for { database <- db_hierarchy.iterator file <- store.read_sources(database.db, database.session, name = name).iterator } yield file).nextOption() } def source_file(name: String): Sessions.Source_File = get_source_file(name).getOrElse(error("Missing session source file " + quote(name))) def classpath(): List[File.Content] = { (for { session <- session_stack.iterator info <- sessions_structure.get(session).iterator if info.export_classpath.nonEmpty matcher = make_matcher(info.export_classpath) entry_name <- entry_names(session = session).iterator if matcher(entry_name) entry <- get(entry_name.theory, entry_name.name).iterator } yield File.content(entry.entry_name.make_path(), entry.bytes)).toList } override def toString: String = "Export.Session_Context(" + commas_quote(session_stack) + ")" } class Theory_Context private[Export]( val session_context: Session_Context, val theory: String, other_cache: Option[Term.Cache] ) { def cache: Term.Cache = other_cache getOrElse session_context.cache def get(name: String): Option[Entry] = session_context.get(theory, name) def apply(name: String, permissive: Boolean = false): Entry = session_context.apply(theory, name, permissive = permissive) def yxml(name: String): XML.Body = get(name) match { case Some(entry) => entry.yxml case None => Nil } def document_id(): Option[Long] = apply(DOCUMENT_ID, permissive = true).text match { case Value.Long(id) => Some(id) case _ => None } def files0(permissive: Boolean = false): List[String] = split_lines(apply(FILES, permissive = permissive).text) def files(permissive: Boolean = false): Option[(String, List[String])] = files0(permissive = permissive) match { case Nil => None case a :: bs => Some((a, bs)) } override def toString: String = "Export.Theory_Context(" + quote(theory) + ")" } /* export to file-system */ def export_files( store: Sessions.Store, session_name: String, export_dir: Path, progress: Progress = new Progress, export_prune: Int = 0, export_list: Boolean = false, export_patterns: List[String] = Nil ): Unit = { using(store.open_database(session_name)) { db => val entry_names = read_entry_names(db, session_name) // list if (export_list) { for (entry_name <- entry_names) progress.echo(entry_name.compound_name) } // export if (export_patterns.nonEmpty) { val matcher = make_matcher(export_patterns) for { entry_name <- entry_names if matcher(entry_name) entry <- entry_name.read(db, store.cache) } { val path = export_dir + entry_name.make_path(prune = export_prune) progress.echo("export " + path + (if (entry.executable) " (executable)" else "")) Isabelle_System.make_directory(path.dir) val bytes = entry.bytes if (!path.is_file || Bytes.read(path) != bytes) Bytes.write(path, bytes) File.set_executable(path, entry.executable) } } } } /* Isabelle tool wrapper */ val default_export_dir: Path = Path.explode("export") val isabelle_tool = Isabelle_Tool("export", "retrieve theory exports", Scala_Project.here, { args => /* arguments */ var export_dir = default_export_dir var dirs: List[Path] = Nil var export_list = false var no_build = false var options = Options.init() var export_prune = 0 var export_patterns: List[String] = Nil val getopts = Getopts(""" Usage: isabelle export [OPTIONS] SESSION Options are: -O DIR output directory for exported files (default: """ + default_export_dir + """) -d DIR include session directory -l list exports -n no build of session -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) -p NUM prune path of exported files by NUM elements -x PATTERN extract files matching pattern (e.g. "*:**" for all) List or export theory exports for SESSION: named blobs produced by isabelle build. Option -l or -x is required; option -x may be repeated. The PATTERN language resembles glob patterns in the shell, with ? and * (both excluding ":" and "/"), ** (excluding ":"), and [abc] or [^abc], and variants {pattern1,pattern2,pattern3}. """, "O:" -> (arg => export_dir = Path.explode(arg)), "d:" -> (arg => dirs = dirs ::: List(Path.explode(arg))), "l" -> (_ => export_list = true), "n" -> (_ => no_build = true), "o:" -> (arg => options = options + arg), "p:" -> (arg => export_prune = Value.Int.parse(arg)), "x:" -> (arg => export_patterns ::= arg)) val more_args = getopts(args) val session_name = more_args match { case List(session_name) if export_list || export_patterns.nonEmpty => session_name case _ => getopts.usage() } val progress = new Console_Progress() /* build */ if (!no_build) { val rc = progress.interrupt_handler { Build.build_logic(options, session_name, progress = progress, dirs = dirs) } if (rc != Process_Result.RC.ok) sys.exit(rc) } /* export files */ val store = Sessions.store(options) export_files(store, session_name, export_dir, progress = progress, export_prune = export_prune, export_list = export_list, export_patterns = export_patterns) }) } diff --git a/src/Pure/Thy/latex.scala b/src/Pure/Thy/latex.scala --- a/src/Pure/Thy/latex.scala +++ b/src/Pure/Thy/latex.scala @@ -1,477 +1,483 @@ /* Title: Pure/Thy/latex.scala Author: Makarius Support for LaTeX. */ package isabelle import java.io.{File => JFile} import scala.annotation.tailrec import scala.collection.mutable import scala.collection.immutable.TreeMap import scala.util.matching.Regex object Latex { /* output name for LaTeX macros */ private val output_name_map: Map[Char, String] = Map('_' -> "UNDERSCORE", '\'' -> "PRIME", '0' -> "ZERO", '1' -> "ONE", '2' -> "TWO", '3' -> "THREE", '4' -> "FOUR", '5' -> "FIVE", '6' -> "SIX", '7' -> "SEVEN", '8' -> "EIGHT", '9' -> "NINE") def output_name(name: String): String = if (name.exists(output_name_map.keySet)) { val res = new StringBuilder for (c <- name) { output_name_map.get(c) match { case None => res += c case Some(s) => res ++= s } } res.toString } else name /* cite: references to bibliography */ object Cite { sealed case class Value(kind: String, citation: String, location: XML.Body) def unapply(tree: XML.Tree): Option[Value] = tree match { case XML.Elem(Markup(Markup.Latex_Cite.name, props), body) => val kind = Markup.Kind.unapply(props).getOrElse("cite") val citations = Markup.Citations.get(props) Some(Value(kind, citations, body)) case _ => None } } /* index entries */ def index_escape(str: String): String = { val special1 = "!\"@|" val special2 = "\\{}#" if (str.exists(c => special1.contains(c) || special2.contains(c))) { val res = new StringBuilder for (c <- str) { if (special1.contains(c)) { res ++= "\\char" res ++= Value.Int(c) } else { if (special2.contains(c)) { res += '"'} res += c } } res.toString } else str } object Index_Item { sealed case class Value(text: Text, like: String) def unapply(tree: XML.Tree): Option[Value] = tree match { case XML.Wrapped_Elem(Markup.Latex_Index_Item(_), text, like) => Some(Value(text, XML.content(like))) case _ => None } } object Index_Entry { sealed case class Value(items: List[Index_Item.Value], kind: String) def unapply(tree: XML.Tree): Option[Value] = tree match { case XML.Elem(Markup.Latex_Index_Entry(kind), body) => val items = body.map(Index_Item.unapply) if (items.forall(_.isDefined)) Some(Value(items.map(_.get), kind)) else None case _ => None } } /* tags */ object Tags { object Op extends Enumeration { val fold, drop, keep = Value } val standard = "document,theory,proof,ML,visible,-invisible,important,unimportant" private def explode(spec: String): List[String] = Library.space_explode(',', spec) def apply(spec: String): Tags = new Tags(spec, (explode(standard) ::: explode(spec)).foldLeft(TreeMap.empty[String, Op.Value]) { case (m, tag) => tag.toList match { case '/' :: cs => m + (cs.mkString -> Op.fold) case '-' :: cs => m + (cs.mkString -> Op.drop) case '+' :: cs => m + (cs.mkString -> Op.keep) case cs => m + (cs.mkString -> Op.keep) } }) val empty: Tags = apply("") } class Tags private(spec: String, map: TreeMap[String, Tags.Op.Value]) { override def toString: String = spec def get(name: String): Option[Tags.Op.Value] = map.get(name) def sty(comment_latex: Boolean): File.Content = { val path = Path.explode("isabelletags.sty") val comment = if (comment_latex) """\usepackage{comment}""" else """%plain TeX version of comment package -- much faster! \let\isafmtname\fmtname\def\fmtname{plain} \usepackage{comment} \let\fmtname\isafmtname""" val tags = (for ((name, op) <- map.iterator) yield "\\isa" + op + "tag{" + name + "}").toList File.content(path, comment + """ \newcommand{\isakeeptag}[1]% {\includecomment{isadelim#1}\includecomment{isatag#1}\csarg\def{isafold#1}{}} \newcommand{\isadroptag}[1]% {\excludecomment{isadelim#1}\excludecomment{isatag#1}\csarg\def{isafold#1}{}} \newcommand{\isafoldtag}[1]% {\includecomment{isadelim#1}\excludecomment{isatag#1}\csarg\def{isafold#1}{\isafold{#1}}} """ + Library.terminate_lines(tags)) } } /* output text and positions */ type Text = XML.Body def position(a: String, b: String): String = "%:%" + a + "=" + b + "%:%\n" def init_position(file_pos: String): List[String] = if (file_pos.isEmpty) Nil else List("\\endinput\n", position(Markup.FILE, file_pos)) def append_position(path: Path, file_pos: String): Unit = { val pos = init_position(file_pos).mkString if (pos.nonEmpty) { val sep = if (File.read(path).endsWith("\n")) "" else "\n" File.append(path, sep + pos) } } def copy_file(src: Path, dst: Path): Unit = { Isabelle_System.copy_file(src, dst) if (src.is_latex) { val target = if (dst.is_dir) dst + src.base else dst val file_pos = File.symbolic_path(src) append_position(target, file_pos) } } def copy_file_base(base_dir: Path, src: Path, target_dir: Path): Unit = { Isabelle_System.copy_file_base(base_dir, src, target_dir) if (src.is_latex) { val file_pos = File.symbolic_path(base_dir + src) append_position(target_dir + src, file_pos) } } class Output(val options: Options) { def latex_output(latex_text: Text): String = make(latex_text) def latex_macro0(name: String, optional_argument: String = ""): Text = XML.string("\\" + name + optional_argument) def latex_macro(name: String, body: Text, optional_argument: String = ""): Text = XML.enclose("\\" + name + optional_argument + "{", "}", body) def latex_environment(name: String, body: Text, optional_argument: String = ""): Text = XML.enclose( "%\n\\begin{" + name + "}" + optional_argument + "%\n", "%\n\\end{" + name + "}", body) def latex_heading(kind: String, body: Text, optional_argument: String = ""): Text = XML.enclose( "%\n\\" + options.string("document_heading_prefix") + kind + optional_argument + "{", "%\n}\n", body) def latex_body(kind: String, body: Text, optional_argument: String = ""): Text = latex_environment("isamarkup" + kind, body, optional_argument) def latex_tag(name: String, body: Text, delim: Boolean = false): Text = { val s = output_name(name) val kind = if (delim) "delim" else "tag" val end = if (delim) "" else "{\\isafold" + s + "}%\n" if (options.bool("document_comment_latex")) { XML.enclose( "%\n\\begin{isa" + kind + s + "}\n", "%\n\\end{isa" + kind + s + "}\n" + end, body) } else { XML.enclose( "%\n\\isa" + kind + s + "\n", "%\n\\endisa" + kind + s + "\n" + end, body) } } def cite(value: Cite.Value): Text = { latex_macro0(value.kind) ::: (if (value.location.isEmpty) Nil else XML.string("[") ::: value.location ::: XML.string("]")) ::: XML.string("{" + value.citation + "}") } def index_item(item: Index_Item.Value): String = { val like = if (item.like.isEmpty) "" else index_escape(item.like) + "@" val text = index_escape(latex_output(item.text)) like + text } def index_entry(entry: Index_Entry.Value): Text = { val items = entry.items.map(index_item).mkString("!") val kind = if (entry.kind.isEmpty) "" else "|" + index_escape(entry.kind) latex_macro("index", XML.string(items + kind)) } /* standard output of text with per-line positions */ def unknown_elem(elem: XML.Elem, pos: Position.T): XML.Body = error("Unknown latex markup element " + quote(elem.name) + Position.here(pos) + ":\n" + XML.string_of_tree(elem)) - def make(latex_text: Text, file_pos: String = ""): String = { + def make( + latex_text: Text, + file_pos: String = "", + line_pos: Properties.T => Option[Int] = Position.Line.unapply + ): String = { var line = 1 val result = new mutable.ListBuffer[String] val positions = new mutable.ListBuffer[String] ++= init_position(file_pos) val file_position = if (file_pos.isEmpty) Position.none else Position.File(file_pos) def traverse(xml: XML.Body): Unit = { xml.foreach { case XML.Text(s) => line += s.count(_ == '\n') result += s case elem @ XML.Elem(markup, body) => val a = Markup.Optional_Argument.get(markup.properties) traverse { markup match { case Markup.Document_Latex(props) => - for (l <- Position.Line.unapply(props) if positions.nonEmpty) { - val s = position(Value.Int(line), Value.Int(l)) - if (positions.last != s) positions += s + if (positions.nonEmpty) { + for (l <- line_pos(props)) { + val s = position(Value.Int(line), Value.Int(l)) + if (positions.last != s) positions += s + } } body case Markup.Latex_Output(_) => XML.string(latex_output(body)) case Markup.Latex_Macro0(name) if body.isEmpty => latex_macro0(name, a) case Markup.Latex_Macro(name) => latex_macro(name, body, a) case Markup.Latex_Environment(name) => latex_environment(name, body, a) case Markup.Latex_Heading(kind) => latex_heading(kind, body, a) case Markup.Latex_Body(kind) => latex_body(kind, body, a) case Markup.Latex_Delim(name) => latex_tag(name, body, delim = true) case Markup.Latex_Tag(name) => latex_tag(name, body) case Markup.Latex_Cite(_) => elem match { case Cite(value) => cite(value) case _ => unknown_elem(elem, file_position) } case Markup.Latex_Index_Entry(_) => elem match { case Index_Entry(entry) => index_entry(entry) case _ => unknown_elem(elem, file_position) } case _ => unknown_elem(elem, file_position) } } } } traverse(latex_text) result ++= positions result.mkString } } /* generated .tex file */ private val File_Pattern = """^%:%file=(.+)%:%$""".r private val Line_Pattern = """^*%:%(\d+)=(\d+)%:%$""".r def read_tex_file(tex_file: Path): Tex_File = { val positions = Line.logical_lines(File.read(tex_file)).reverse. takeWhile(_ != "\\endinput").reverse val source_file = positions match { case File_Pattern(file) :: _ => Some(file) case _ => None } val source_lines = if (source_file.isEmpty) Nil else positions.flatMap(line => line match { case Line_Pattern(Value.Int(line), Value.Int(source_line)) => Some(line -> source_line) case _ => None }) new Tex_File(tex_file, source_file, source_lines) } final class Tex_File private[Latex]( tex_file: Path, source_file: Option[String], source_lines: List[(Int, Int)] ) { override def toString: String = tex_file.toString def source_position(l: Int): Option[Position.T] = source_file match { case None => None case Some(file) => val source_line = source_lines.iterator.takeWhile({ case (m, _) => m <= l }). foldLeft(0) { case (_, (_, n)) => n } if (source_line == 0) None else Some(Position.Line_File(source_line, file)) } def position(line: Int): Position.T = source_position(line) getOrElse Position.Line_File(line, source_file.getOrElse(tex_file.implode)) } /* latex log */ def latex_errors(dir: Path, root_name: String): List[String] = { val root_log_path = dir + Path.explode(root_name).ext("log") if (root_log_path.is_file) { for { (msg, pos) <- filter_errors(dir, File.read(root_log_path)) } yield "Latex error" + Position.here(pos) + ":\n" + Library.indent_lines(2, msg) } else Nil } def filter_errors(dir: Path, root_log: String): List[(String, Position.T)] = { val seen_files = Synchronized(Map.empty[JFile, Tex_File]) def check_tex_file(path: Path): Option[Tex_File] = seen_files.change_result(seen => seen.get(path.file) match { case None => if (path.is_file) { val tex_file = read_tex_file(path) (Some(tex_file), seen + (path.file -> tex_file)) } else (None, seen) case some => (some, seen) }) def tex_file_position(path: Path, line: Int): Position.T = check_tex_file(path) match { case Some(tex_file) => tex_file.position(line) case None => Position.Line_File(line, path.implode) } object File_Line_Error { val Pattern: Regex = """^(.*?\.\w\w\w):(\d+): (.*)$""".r def unapply(line: String): Option[(Path, Int, String)] = line match { case Pattern(file, Value.Int(line), message) => val path = File.standard_path(file) if (Path.is_wellformed(path)) { val file = (dir + Path.explode(path)).canonical val msg = Library.perhaps_unprefix("LaTeX Error: ", message) if (file.is_file) Some((file, line, msg)) else None } else None case _ => None } } val Line_Error = """^l\.\d+ (.*)$""".r val More_Error = List( "", "