src/Pure/System/build.scala
author wenzelm
Thu, 19 Jul 2012 14:24:40 +0200
changeset 48349 a78e5d399599
parent 48347 8bb27ab9e841
child 48350 09bf3b73e446
permissions -rw-r--r--
support Session.Queue with ordering and dependencies;

/*  Title:      Pure/System/build.scala
    Author:     Makarius

Build and manage Isabelle sessions.
*/

package isabelle


import java.io.File

import scala.collection.mutable
import scala.annotation.tailrec


object Build
{
  /** session information **/

  type Options = List[(String, Option[String])]

  object Session
  {
    object Key
    {
      object Ordering extends scala.math.Ordering[Key]
      {
        def compare(key1: Key, key2: Key): Int =
          key2.order compare key1.order match {
            case 0 => key1.name compare key2.name
            case ord => ord
          }
      }
    }

    sealed case class Key(name: String, order: Int)
    {
      override def toString: String = name
    }

    sealed case class Info(
      dir: Path,
      key: Key,
      parent: Option[String],
      description: String,
      options: Options,
      theories: List[(Options, String)],
      files: List[String])
    {
      def name: String = key.name
    }

    object Queue
    {
      val empty: Queue = new Queue()
    }

    final class Queue private(
      keys: Map[String, Key] = Map.empty,
      graph: Graph[Key, Info] = Graph.empty(Key.Ordering))
    {
      def lookup(name: String): Option[Info] =
        keys.get(name).map(graph.get_node(_))

      def + (info: Info): Queue =
      {
        val keys1 =
          if (keys.isDefinedAt(info.name)) error("Duplicate session: " + quote(info.name))
          else keys + (info.name -> info.key)

        val graph1 =
          try {
            graph.new_node(info.key, info).
              add_deps_acyclic(info.key, info.parent.toList.map(keys(_)))
          }
          catch {
            case exn: Graph.Cycles[_] =>
              error(cat_lines(exn.cycles.map(cycle =>
                "Cyclic session dependency of " +
                  cycle.map(key => quote(key.toString)).mkString(" via "))))
          }
        new Queue(keys1, graph1)
      }

      def topological_order: List[Info] =
        graph.topological_order.map(graph.get_node(_))
    }
  }


  /* parsing */

  val ROOT_NAME = "ROOT"

  private case class Session_Entry(
    name: String,
    reset: Boolean,
    order: Int,
    path: Option[String],
    parent: Option[String],
    description: String,
    options: Options,
    theories: List[(Options, List[String])],
    files: List[String])

  private object Parser extends Parse.Parser
  {
    val SESSION = "session"
    val IN = "in"
    val DESCRIPTION = "description"
    val OPTIONS = "options"
    val THEORIES = "theories"
    val FILES = "files"

    val syntax =
      Outer_Syntax.empty + "!" + "(" + ")" + "+" + "," + "=" + "[" + "]" +
        SESSION + IN + DESCRIPTION + OPTIONS + THEORIES + FILES

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

      val option =
        name ~ opt(keyword("=") ~! name ^^ { case _ ~ x => x }) ^^ { case x ~ y => (x, y) }
      val options = keyword("[") ~> repsep(option, keyword(",")) <~ keyword("]")

      val theories =
        keyword(THEORIES) ~! ((options | success(Nil)) ~ rep1(theory_name)) ^^
          { case _ ~ (x ~ y) => (x, y) }

      ((keyword(SESSION) ~! session_name) ^^ { case _ ~ x => x }) ~
        (keyword("!") ^^^ true | success(false)) ~
        (keyword("(") ~! (nat <~ keyword(")")) ^^ { case _ ~ x => x } | success(Integer.MAX_VALUE)) ~
        (opt(keyword(IN) ~! string ^^ { case _ ~ x => x })) ~
        (keyword("=") ~> opt(session_name <~ keyword("+"))) ~
        (keyword(DESCRIPTION) ~! text ^^ { case _ ~ x => x } | success("")) ~
        (keyword(OPTIONS) ~! options ^^ { case _ ~ x => x } | success(Nil)) ~
        rep(theories) ~
        (keyword(FILES) ~! rep1(path) ^^ { case _ ~ x => x } | success(Nil)) ^^
          { case a ~ b ~ c ~ d ~ e ~ f ~ g ~ h ~ i => Session_Entry(a, b, c, d, e, f, g, h, i) }
    }

    def parse_entries(root: File): List[Session_Entry] =
    {
      val toks = syntax.scan(Standard_System.read_file(root))
      parse_all(rep(session_entry), Token.reader(toks, root.toString)) match {
        case Success(result, _) => result
        case bad => error(bad.toString)
      }
    }
  }


  /* find sessions */

  def find_sessions(more_dirs: List[Path]): Session.Queue =
  {
    var sessions = Session.Queue.empty

    for {
      (dir, strict) <- Isabelle_System.components().map((_, false)) ++ more_dirs.map((_, true))
      root = Isabelle_System.platform_file(dir + Path.basic(ROOT_NAME))
      _ = (strict && !root.isFile && error("Bad session root file: " + quote(root.toString)))
      if root.isFile
      entry <- Parser.parse_entries(root)
    }
    {
      try {
        if (entry.name == "") error("Bad session name")

        val full_name =
          if (entry.name == "RAW" || entry.name == "Pure") {
            if (entry.parent.isDefined) error("Illegal parent session")
            else entry.name
          }
          else
            entry.parent match {
              case None => error("Missing parent session")
              case Some(parent) =>
                val parent_info =
                  sessions.lookup(parent) getOrElse
                    error("Undefined parent session: " + quote(parent))
                if (entry.reset) entry.name
                else parent_info.name + "-" + entry.name
            }

        val path =
          entry.path match {
            case Some(p) => Path.explode(p)
            case None => Path.basic(entry.name)
          }
        val thys = entry.theories.map({ case (x, ys) => ys.map(y => (x, y)) }).flatten
        val info =
          Session.Info(dir + path, Session.Key(full_name, entry.order),
            entry.parent, entry.description, entry.options, thys, entry.files)

        sessions += info
      }
      catch {
        case ERROR(msg) =>
          error(msg + "\nThe error(s) above occurred in session entry " +
            quote(entry.name) + " (file " + quote(root.toString) + ")")
      }
    }
    sessions
  }



  /** build **/

  def build(all_sessions: Boolean, build_images: Boolean, list_only: Boolean,
    more_dirs: List[Path], options: List[String], sessions: List[String]): Int =
  {
    println("more_dirs = " + more_dirs.toString)
    println("options = " + options.toString)
    println("sessions = " + sessions.toString)

    for (info <- find_sessions(more_dirs).topological_order)
      println(info.name + " in " + info.dir)

    0
  }


  /* command line entry point */

  def main(args: Array[String])
  {
    Command_Line.tool {
      args.toList match {
        case
          Properties.Value.Boolean(all_sessions) ::
          Properties.Value.Boolean(build_images) ::
          Properties.Value.Boolean(list_only) ::
          Command_Line.Chunks(more_dirs, options, sessions) =>
            build(all_sessions, build_images, list_only,
              more_dirs.map(Path.explode), options, sessions)
        case _ => error("Bad arguments:\n" + cat_lines(args))
      }
    }
  }
}