src/Pure/Thy/sessions.scala
author wenzelm
Fri, 07 Apr 2017 11:50:49 +0200
changeset 65420 695d4e22345a
parent 65419 457e4fbed731
child 65422 b606c98e6d10
permissions -rw-r--r--
support for static session imports, without affect build hierarchy;

/*  Title:      Pure/Thy/sessions.scala
    Author:     Makarius

Isabelle session information.
*/

package isabelle

import java.nio.ByteBuffer
import java.nio.channels.FileChannel
import java.nio.file.StandardOpenOption
import java.sql.PreparedStatement

import scala.collection.SortedSet
import scala.collection.mutable


object Sessions
{
  /* base info and source dependencies */

  def is_pure(name: String): Boolean = name == Thy_Header.PURE

  object Base
  {
    def pure(options: Options): Base = session_base(options, Thy_Header.PURE)

    lazy val bootstrap: Base =
      Base(keywords = Thy_Header.bootstrap_header, syntax = Thy_Header.bootstrap_syntax)
  }

  sealed case class Base(
    global_theories: Set[String] = Set.empty,
    loaded_theories: Set[String] = Set.empty,
    known_theories: Map[String, Document.Node.Name] = Map.empty,
    keywords: Thy_Header.Keywords = Nil,
    syntax: Outer_Syntax = Outer_Syntax.empty,
    sources: List[(Path, SHA1.Digest)] = Nil,
    session_graph: Graph_Display.Graph = Graph_Display.empty_graph)
  {
    def loaded_theory(name: Document.Node.Name): Boolean =
      loaded_theories.contains(name.theory)
  }

  sealed case class Deps(sessions: Map[String, Base])
  {
    def is_empty: Boolean = sessions.isEmpty
    def apply(name: String): Base = sessions(name)
    def sources(name: String): List[SHA1.Digest] = sessions(name).sources.map(_._2)
  }

  def deps(sessions: T,
      progress: Progress = No_Progress,
      inlined_files: Boolean = false,
      verbose: Boolean = false,
      list_files: Boolean = false,
      check_keywords: Set[String] = Set.empty,
      global_theories: Set[String] = Set.empty): Deps =
  {
    Deps((Map.empty[String, Base] /: sessions.imports_topological_order)({
      case (sessions, (name, info)) =>
        if (progress.stopped) throw Exn.Interrupt()

        try {
          val parent_base =
            info.parent match {
              case None => Base.bootstrap
              case Some(parent) => sessions(parent)
            }
          val resources = new Resources(name, parent_base)

          if (verbose || list_files) {
            val groups =
              if (info.groups.isEmpty) ""
              else info.groups.mkString(" (", " ", ")")
            progress.echo("Session " + info.chapter + "/" + name + groups)
          }

          val thy_deps =
          {
            val root_theories =
              info.theories.flatMap({ case (_, thys) =>
                thys.map(thy => (resources.import_name(info.dir.implode, thy), info.pos))
              })
            val thy_deps = resources.thy_info.dependencies(root_theories)

            thy_deps.errors match {
              case Nil => thy_deps
              case errs => error(cat_lines(errs))
            }
          }

          val known_theories =
            (parent_base.known_theories /: thy_deps.deps)({ case (known, dep) =>
              val name = dep.name
              known.get(name.theory) match {
                case Some(name1) if name != name1 =>
                  error("Duplicate theory " + quote(name.node) + " vs. " + quote(name1.node))
                case _ =>
                  known + (name.theory -> name) +
                    (Long_Name.base_name(name.theory) -> name)  // legacy
              }
            })

          val loaded_theories = thy_deps.loaded_theories
          val keywords = thy_deps.keywords
          val syntax = thy_deps.syntax

          val theory_files = thy_deps.deps.map(dep => Path.explode(dep.name.node))
          val loaded_files =
            if (inlined_files) {
              val pure_files =
                if (is_pure(name)) {
                  val roots = Thy_Header.ml_roots.map(p => info.dir + Path.explode(p._1))
                  val files =
                    roots.flatMap(root => resources.loaded_files(syntax, File.read(root))).
                      map(file => info.dir + Path.explode(file))
                  roots ::: files
                }
                else Nil
              pure_files ::: thy_deps.loaded_files
            }
            else Nil

          val all_files =
            (theory_files ::: loaded_files :::
              info.files.map(file => info.dir + file) :::
              info.document_files.map(file => info.dir + file._1 + file._2)).map(_.expand)

          if (list_files)
            progress.echo(cat_lines(all_files.map(_.implode).sorted.map("  " + _)))

          if (check_keywords.nonEmpty)
            Check_Keywords.check_keywords(progress, syntax.keywords, check_keywords, theory_files)

          val base =
            Base(global_theories = global_theories,
              loaded_theories = loaded_theories,
              known_theories = known_theories,
              keywords = keywords,
              syntax = syntax,
              sources = all_files.map(p => (p, SHA1.digest(p.file))),
              session_graph = thy_deps.session_graph(info.parent getOrElse "", parent_base))

          sessions + (name -> base)
        }
        catch {
          case ERROR(msg) =>
            cat_error(msg, "The error(s) above occurred in session " +
              quote(name) + Position.here(info.pos))
        }
    }))
  }

  def session_base(options: Options, session: String, dirs: List[Path] = Nil): Base =
  {
    val full_sessions = load(options, dirs = dirs)
    val (_, selected_sessions) =
      full_sessions.selection(Selection(sessions = List(session)))

    deps(selected_sessions, global_theories = full_sessions.global_theories)(session)
  }


  /* cumulative session info */

  sealed case class Info(
    chapter: String,
    select: Boolean,
    pos: Position.T,
    groups: List[String],
    dir: Path,
    parent: Option[String],
    description: String,
    options: Options,
    imports: List[String],
    theories: List[(Options, List[String])],
    global_theories: List[String],
    files: List[Path],
    document_files: List[(Path, Path)],
    meta_digest: SHA1.Digest)
  {
    def timeout: Time = Time.seconds(options.real("timeout") * options.real("timeout_scale"))
  }

  object Selection
  {
    val all: Selection = Selection(all_sessions = true)
  }

  sealed case class Selection(
    requirements: Boolean = false,
    all_sessions: Boolean = false,
    exclude_session_groups: List[String] = Nil,
    exclude_sessions: List[String] = Nil,
    session_groups: List[String] = Nil,
    sessions: List[String] = Nil)
  {
    def apply(graph: Graph[String, Info]): (List[String], Graph[String, Info]) =
    {
      val bad_sessions =
        SortedSet((exclude_sessions ::: sessions).filterNot(graph.defined(_)): _*).toList
      if (bad_sessions.nonEmpty) error("Undefined session(s): " + commas_quote(bad_sessions))

      val excluded =
      {
        val exclude_group = exclude_session_groups.toSet
        val exclude_group_sessions =
          (for {
            (name, (info, _)) <- graph.iterator
            if graph.get_node(name).groups.exists(exclude_group)
          } yield name).toList
        graph.all_succs(exclude_group_sessions ::: exclude_sessions).toSet
      }

      val pre_selected =
      {
        if (all_sessions) graph.keys
        else {
          val select_group = session_groups.toSet
          val select = sessions.toSet
          (for {
            (name, (info, _)) <- graph.iterator
            if info.select || select(name) || graph.get_node(name).groups.exists(select_group)
          } yield name).toList
        }
      }.filterNot(excluded)

      val selected =
        if (requirements) (graph.all_preds(pre_selected).toSet -- pre_selected).toList
        else pre_selected

      (selected, graph.restrict(graph.all_preds(selected).toSet))
    }
  }

  def make(infos: Traversable[(String, Info)]): T =
  {
    def add_edges(graph: Graph[String, Info], kind: String, edges: Info => Traversable[String])
      : Graph[String, Info] =
    {
      def add_edge(pos: Position.T, name: String, g: Graph[String, Info], parent: String) =
      {
        if (!g.defined(parent))
          error("Bad " + kind + " session " + quote(parent) + " for " +
            quote(name) + Position.here(pos))

        try { g.add_edge_acyclic(parent, name) }
        catch {
          case exn: Graph.Cycles[_] =>
            error(cat_lines(exn.cycles.map(cycle =>
              "Cyclic session dependency of " +
                cycle.map(c => quote(c.toString)).mkString(" via "))) + Position.here(pos))
        }
      }
      (graph /: graph.iterator) {
        case (g, (name, (info, _))) => (g /: edges(info))(add_edge(info.pos, name, _, _))
      }
    }

    val graph0 =
      (Graph.string[Info] /: infos) {
        case (graph, (name, info)) =>
          if (graph.defined(name))
            error("Duplicate session " + quote(name) + Position.here(info.pos) +
              Position.here(graph.get_node(name).pos))
          else graph.new_node(name, info)
      }
    val graph1 = add_edges(graph0, "parent", _.parent)
    val graph2 = add_edges(graph1, "imports", _.imports)

    new T(graph1, graph2)
  }

  final class T private[Sessions](
      val build_graph: Graph[String, Info],
      val imports_graph: Graph[String, Info])
  {
    def apply(name: String): Info = imports_graph.get_node(name)
    def get(name: String): Option[Info] =
      if (imports_graph.defined(name)) Some(imports_graph.get_node(name)) else None

    def global_theories: Set[String] =
      (for {
        (_, (info, _)) <- imports_graph.iterator
        name <- info.global_theories.iterator }
       yield name).toSet

    def selection(select: Selection): (List[String], T) =
    {
      val (_, build_graph1) = select(build_graph)
      val (selected, imports_graph1) = select(imports_graph)
      (selected, new T(build_graph1, imports_graph1))
    }

    def build_ancestors(name: String): List[String] =
      build_graph.all_preds(List(name)).tail.reverse

    def build_descendants(names: List[String]): List[String] =
      build_graph.all_succs(names)

    def build_topological_order: List[(String, Info)] =
      build_graph.topological_order.map(name => (name, apply(name)))

    def imports_topological_order: List[(String, Info)] =
      imports_graph.topological_order.map(name => (name, apply(name)))

    override def toString: String =
      imports_graph.keys_iterator.mkString("Sessions.T(", ", ", ")")
  }


  /* parser */

  val ROOT = Path.explode("ROOT")
  val ROOTS = Path.explode("ROOTS")

  private val CHAPTER = "chapter"
  private val SESSION = "session"
  private val IN = "in"
  private val DESCRIPTION = "description"
  private val OPTIONS = "options"
  private val SESSIONS = "sessions"
  private val THEORIES = "theories"
  private val GLOBAL = "global"
  private val FILES = "files"
  private val DOCUMENT_FILES = "document_files"

  lazy val root_syntax =
    Outer_Syntax.init() + "(" + ")" + "+" + "," + "=" + "[" + "]" + GLOBAL + IN +
      (CHAPTER, Keyword.THY_DECL) +
      (SESSION, Keyword.THY_DECL) +
      (DESCRIPTION, Keyword.QUASI_COMMAND) +
      (OPTIONS, Keyword.QUASI_COMMAND) +
      (SESSIONS, Keyword.QUASI_COMMAND) +
      (THEORIES, Keyword.QUASI_COMMAND) +
      (FILES, Keyword.QUASI_COMMAND) +
      (DOCUMENT_FILES, Keyword.QUASI_COMMAND)

  private object Parser extends Parse.Parser with Options.Parser
  {
    private abstract class Entry
    private sealed case class Chapter(name: String) extends Entry
    private sealed case class Session_Entry(
      pos: Position.T,
      name: String,
      groups: List[String],
      path: String,
      parent: Option[String],
      description: String,
      options: List[Options.Spec],
      imports: List[String],
      theories: List[(List[Options.Spec], List[(String, Boolean)])],
      files: List[String],
      document_files: List[(String, String)]) extends Entry

    private val chapter: Parser[Chapter] =
    {
      val chapter_name = atom("chapter name", _.is_name)

      command(CHAPTER) ~! chapter_name ^^ { case _ ~ a => Chapter(a) }
    }

    private val session_entry: Parser[Session_Entry] =
    {
      val session_name = atom("session name", _.is_name)

      val option =
        option_name ~ opt($$$("=") ~! option_value ^^
          { case _ ~ x => x }) ^^ { case x ~ y => (x, y) }
      val options = $$$("[") ~> rep1sep(option, $$$(",")) <~ $$$("]")

      val global =
        ($$$("(") ~! $$$(GLOBAL) ~ $$$(")")) ^^ { case _ => true } | success(false)

      val theory_entry =
        theory_name ~ global ^^ { case x ~ y => (x, y) }

      val theories =
        $$$(THEORIES) ~!
          ((options | success(Nil)) ~ rep(theory_entry)) ^^
          { case _ ~ (x ~ y) => (x, y) }

      val document_files =
        $$$(DOCUMENT_FILES) ~!
          (($$$("(") ~! ($$$(IN) ~! (path ~ $$$(")"))) ^^
              { case _ ~ (_ ~ (x ~ _)) => x } | success("document")) ~
            rep1(path)) ^^ { case _ ~ (x ~ y) => y.map((x, _)) }

      command(SESSION) ~!
        (position(session_name) ~
          (($$$("(") ~! (rep1(name) <~ $$$(")")) ^^ { case _ ~ x => x }) | success(Nil)) ~
          (($$$(IN) ~! path ^^ { case _ ~ x => x }) | success(".")) ~
          ($$$("=") ~!
            (opt(session_name ~! $$$("+") ^^ { case x ~ _ => x }) ~
              (($$$(DESCRIPTION) ~! text ^^ { case _ ~ x => x }) | success("")) ~
              (($$$(OPTIONS) ~! options ^^ { case _ ~ x => x }) | success(Nil)) ~
              (($$$(SESSIONS) ~! rep(session_name)  ^^ { case _ ~ x => x }) | success(Nil)) ~
              rep1(theories) ~
              (($$$(FILES) ~! rep1(path) ^^ { case _ ~ x => x }) | success(Nil)) ~
              (rep(document_files) ^^ (x => x.flatten))))) ^^
        { case _ ~ ((a, pos) ~ b ~ c ~ (_ ~ (d ~ e ~ f ~ g ~ h ~ i ~ j))) =>
            Session_Entry(pos, a, b, c, d, e, f, g, h, i, j) }
    }

    def parse(options: Options, select: Boolean, dir: Path): List[(String, Info)] =
    {
      def make_info(entry_chapter: String, entry: Session_Entry): (String, Info) =
      {
        try {
          val name = entry.name

          if (name == "") error("Bad session name")
          if (is_pure(name) && entry.parent.isDefined) error("Illegal parent session")
          if (!is_pure(name) && !entry.parent.isDefined) error("Missing parent session")

          val session_options = options ++ entry.options

          val theories =
            entry.theories.map({ case (opts, thys) => (session_options ++ opts, thys.map(_._1)) })

          val global_theories =
            for { (_, thys) <- entry.theories; (thy, global) <- thys if global }
            yield {
              val thy_name = Path.explode(thy).expand.base.implode
              if (Long_Name.is_qualified(thy_name))
                error("Bad qualified name for global theory " + quote(thy_name))
              else thy_name
            }

          val files = entry.files.map(Path.explode(_))
          val document_files =
            entry.document_files.map({ case (s1, s2) => (Path.explode(s1), Path.explode(s2)) })

          val meta_digest =
            SHA1.digest((entry_chapter, name, entry.parent, entry.options,
              entry.imports, entry.theories, entry.files, entry.document_files).toString)

          val info =
            Info(entry_chapter, select, entry.pos, entry.groups, dir + Path.explode(entry.path),
              entry.parent, entry.description, session_options, entry.imports, theories,
              global_theories, files, document_files, meta_digest)

          (name, info)
        }
        catch {
          case ERROR(msg) =>
            error(msg + "\nThe error(s) above occurred in session entry " +
              quote(entry.name) + Position.here(entry.pos))
        }
      }

      val root = dir + ROOT
      if (root.is_file) {
        val toks = Token.explode(root_syntax.keywords, File.read(root))
        val start = Token.Pos.file(root.implode)

        parse_all(rep(chapter | session_entry), Token.reader(toks, start)) match {
          case Success(result, _) =>
            var entry_chapter = "Unsorted"
            val infos = new mutable.ListBuffer[(String, Info)]
            result.foreach {
              case Chapter(name) => entry_chapter = name
              case entry: Session_Entry => infos += make_info(entry_chapter, entry)
            }
            infos.toList
          case bad => error(bad.toString)
        }
      }
      else Nil
    }
  }


  /* load sessions from certain directories */

  private def is_session_dir(dir: Path): Boolean =
    (dir + ROOT).is_file || (dir + ROOTS).is_file

  private def check_session_dir(dir: Path): Path =
    if (is_session_dir(dir)) dir
    else error("Bad session root directory: " + dir.toString)

  def load(options: Options, dirs: List[Path] = Nil, select_dirs: List[Path] = Nil): T =
  {
    def load_dir(select: Boolean, dir: Path): List[(String, Info)] =
      load_root(select, dir) ::: load_roots(select, dir)

    def load_root(select: Boolean, dir: Path): List[(String, Info)] =
      Parser.parse(options, select, dir)

    def load_roots(select: Boolean, dir: Path): List[(String, Info)] =
    {
      val roots = dir + ROOTS
      if (roots.is_file) {
        for {
          line <- split_lines(File.read(roots))
          if !(line == "" || line.startsWith("#"))
          dir1 =
            try { check_session_dir(dir + Path.explode(line)) }
            catch {
              case ERROR(msg) =>
                error(msg + "\nThe error(s) above occurred in session catalog " + roots.toString)
            }
          info <- load_dir(select, dir1)
        } yield info
      }
      else Nil
    }

    val default_dirs = Isabelle_System.components().filter(is_session_dir(_))
    dirs.foreach(check_session_dir(_))
    select_dirs.foreach(check_session_dir(_))

    make(
      for {
        (select, dir) <- (default_dirs ::: dirs).map((false, _)) ::: select_dirs.map((true, _))
        info <- load_dir(select, dir)
      } yield info)
  }



  /** heap file with SHA1 digest **/

  private val sha1_prefix = "SHA1:"

  def read_heap_digest(heap: Path): Option[String] =
  {
    if (heap.is_file) {
      val file = FileChannel.open(heap.file.toPath, StandardOpenOption.READ)
      try {
        val len = file.size
        val n = sha1_prefix.length + SHA1.digest_length
        if (len >= n) {
          file.position(len - n)

          val buf = ByteBuffer.allocate(n)
          var i = 0
          var m = 0
          do {
            m = file.read(buf)
            if (m != -1) i += m
          }
          while (m != -1 && n > i)

          if (i == n) {
            val prefix = new String(buf.array(), 0, sha1_prefix.length, UTF8.charset)
            val s = new String(buf.array(), sha1_prefix.length, SHA1.digest_length, UTF8.charset)
            if (prefix == sha1_prefix) Some(s) else None
          }
          else None
        }
        else None
      }
      finally { file.close }
    }
    else None
  }

  def write_heap_digest(heap: Path): String =
    read_heap_digest(heap) match {
      case None =>
        val s = SHA1.digest(heap).rep
        File.append(heap, sha1_prefix + s)
        s
      case Some(s) => s
    }



  /** persistent store **/

  object Session_Info
  {
    val session_name = SQL.Column.string("session_name", primary_key = true)

    // Build_Log.Session_Info
    val session_timing = SQL.Column.bytes("session_timing")
    val command_timings = SQL.Column.bytes("command_timings")
    val ml_statistics = SQL.Column.bytes("ml_statistics")
    val task_statistics = SQL.Column.bytes("task_statistics")
    val build_log_columns =
      List(session_name, session_timing, command_timings, ml_statistics, task_statistics)

    // Build.Session_Info
    val sources = SQL.Column.string("sources")
    val input_heaps = SQL.Column.string("input_heaps")
    val output_heap = SQL.Column.string("output_heap")
    val return_code = SQL.Column.int("return_code")
    val build_columns = List(sources, input_heaps, output_heap, return_code)

    val table = SQL.Table("isabelle_session_info", build_log_columns ::: build_columns)

    def where_session_name(name: String): String =
      "WHERE " + session_name.sql_name + " = " + SQL.quote_string(name)

    def select_statement(db: SQL.Database, name: String, columns: List[SQL.Column])
        : PreparedStatement =
      db.select_statement(table, columns, where_session_name(name))

    def delete_statement(db: SQL.Database, name: String): PreparedStatement =
      db.delete_statement(table, where_session_name(name))
  }

  def store(system_mode: Boolean = false): Store = new Store(system_mode)

  class Store private[Sessions](system_mode: Boolean)
  {
    /* file names */

    def database(name: String): Path = Path.basic("log") + Path.basic(name).ext("db")
    def log(name: String): Path = Path.basic("log") + Path.basic(name)
    def log_gz(name: String): Path = log(name).ext("gz")


    /* SQL database content */

    val xml_cache: XML.Cache = new XML.Cache()

    def encode_properties(ps: Properties.T): Bytes =
      Bytes(YXML.string_of_body(XML.Encode.properties(ps)))

    def decode_properties(bs: Bytes): Properties.T =
      xml_cache.props(XML.Decode.properties(YXML.parse_body(bs.text)))

    def compress_properties(ps: List[Properties.T], options: XZ.Options = XZ.options()): Bytes =
    {
      if (ps.isEmpty) Bytes.empty
      else Bytes(YXML.string_of_body(XML.Encode.list(XML.Encode.properties)(ps))).compress(options)
    }

    def uncompress_properties(bs: Bytes): List[Properties.T] =
    {
      if (bs.isEmpty) Nil
      else
        XML.Decode.list(XML.Decode.properties)(YXML.parse_body(bs.uncompress().text)).
          map(xml_cache.props(_))
    }

    def read_bytes(db: SQL.Database, name: String, column: SQL.Column): Bytes =
      using(Session_Info.select_statement(db, name, List(column)))(stmt =>
      {
        val rs = stmt.executeQuery
        if (!rs.next) Bytes.empty else db.bytes(rs, column)
      })

    def read_properties(db: SQL.Database, name: String, column: SQL.Column): List[Properties.T] =
      uncompress_properties(read_bytes(db, name, column))


    /* output */

    val browser_info: Path =
      if (system_mode) Path.explode("~~/browser_info")
      else Path.explode("$ISABELLE_BROWSER_INFO")

    val output_dir: Path =
      if (system_mode) Path.explode("~~/heaps/$ML_IDENTIFIER")
      else Path.explode("$ISABELLE_OUTPUT")

    override def toString: String = "Store(output_dir = " + output_dir.expand + ")"

    def prepare_output() { Isabelle_System.mkdirs(output_dir + Path.basic("log")) }


    /* input */

    private val input_dirs =
      if (system_mode) List(output_dir)
      else {
        val ml_ident = Path.explode("$ML_IDENTIFIER").expand
        output_dir :: Path.split(Isabelle_System.getenv_strict("ISABELLE_PATH")).map(_ + ml_ident)
      }

    def find_database_heap(name: String): Option[(Path, Option[String])] =
      input_dirs.find(dir => (dir + database(name)).is_file).map(dir =>
        (dir + database(name), read_heap_digest(dir + Path.basic(name))))

    def find_database(name: String): Option[Path] =
      input_dirs.map(_ + database(name)).find(_.is_file)

    def heap(name: String): Path =
      input_dirs.map(_ + Path.basic(name)).find(_.is_file) getOrElse
        error("Unknown logic " + quote(name) + " -- no heap file found in:\n" +
          cat_lines(input_dirs.map(dir => "  " + dir.expand.implode)))


    /* session info */

    def write_session_info(
      db: SQL.Database,
      name: String,
      build_log: Build_Log.Session_Info,
      build: Build.Session_Info)
    {
      db.transaction {
        db.create_table(Session_Info.table)
        using(Session_Info.delete_statement(db, name))(_.execute)
        using(db.insert_statement(Session_Info.table))(stmt =>
        {
          db.set_string(stmt, 1, name)
          db.set_bytes(stmt, 2, encode_properties(build_log.session_timing))
          db.set_bytes(stmt, 3, compress_properties(build_log.command_timings))
          db.set_bytes(stmt, 4, compress_properties(build_log.ml_statistics))
          db.set_bytes(stmt, 5, compress_properties(build_log.task_statistics))
          db.set_string(stmt, 6, cat_lines(build.sources))
          db.set_string(stmt, 7, cat_lines(build.input_heaps))
          db.set_string(stmt, 8, build.output_heap getOrElse "")
          db.set_int(stmt, 9, build.return_code)
          stmt.execute()
        })
      }
    }

    def read_session_timing(db: SQL.Database, name: String): Properties.T =
      decode_properties(read_bytes(db, name, Session_Info.session_timing))

    def read_command_timings(db: SQL.Database, name: String): List[Properties.T] =
      read_properties(db, name, Session_Info.command_timings)

    def read_ml_statistics(db: SQL.Database, name: String): List[Properties.T] =
      read_properties(db, name, Session_Info.ml_statistics)

    def read_task_statistics(db: SQL.Database, name: String): List[Properties.T] =
      read_properties(db, name, Session_Info.task_statistics)

    def read_build_log(db: SQL.Database, name: String,
      command_timings: Boolean = false,
      ml_statistics: Boolean = false,
      task_statistics: Boolean = false): Build_Log.Session_Info =
    {
      Build_Log.Session_Info(
        session_timing = read_session_timing(db, name),
        command_timings = if (command_timings) read_command_timings(db, name) else Nil,
        ml_statistics = if (ml_statistics) read_ml_statistics(db, name) else Nil,
        task_statistics = if (task_statistics) read_task_statistics(db, name) else Nil)
    }

    def read_build(db: SQL.Database, name: String): Option[Build.Session_Info] =
      using(Session_Info.select_statement(db, name, Session_Info.build_columns))(stmt =>
      {
        val rs = stmt.executeQuery
        if (!rs.next) None
        else {
          Some(
            Build.Session_Info(
              split_lines(db.string(rs, Session_Info.sources)),
              split_lines(db.string(rs, Session_Info.input_heaps)),
              db.string(rs, Session_Info.output_heap) match { case "" => None case s => Some(s) },
              db.int(rs, Session_Info.return_code)))
        }
      })
  }
}