src/Pure/Tools/build.scala
changeset 50686 d703e3aafa8c
parent 50566 b43c4f660320
child 50707 5b2bf7611662
equal deleted inserted replaced
50685:293e8ec4dfc8 50686:d703e3aafa8c
       
     1 /*  Title:      Pure/Tools/build.scala
       
     2     Author:     Makarius
       
     3     Options:    :folding=explicit:collapseFolds=1:
       
     4 
       
     5 Build and manage Isabelle sessions.
       
     6 */
       
     7 
       
     8 package isabelle
       
     9 
       
    10 
       
    11 import java.util.{Timer, TimerTask}
       
    12 import java.io.{BufferedInputStream, FileInputStream,
       
    13   BufferedReader, InputStreamReader, IOException}
       
    14 import java.util.zip.GZIPInputStream
       
    15 
       
    16 import scala.collection.SortedSet
       
    17 import scala.annotation.tailrec
       
    18 
       
    19 
       
    20 object Build
       
    21 {
       
    22   /** progress context **/
       
    23 
       
    24   class Progress {
       
    25     def echo(msg: String) {}
       
    26     def stopped: Boolean = false
       
    27   }
       
    28 
       
    29   object Ignore_Progress extends Progress
       
    30 
       
    31   object Console_Progress extends Progress {
       
    32     override def echo(msg: String) { java.lang.System.out.println(msg) }
       
    33   }
       
    34 
       
    35 
       
    36 
       
    37   /** session information **/
       
    38 
       
    39   // external version
       
    40   sealed case class Session_Entry(
       
    41     pos: Position.T,
       
    42     name: String,
       
    43     groups: List[String],
       
    44     path: String,
       
    45     parent: Option[String],
       
    46     description: String,
       
    47     options: List[Options.Spec],
       
    48     theories: List[(List[Options.Spec], List[String])],
       
    49     files: List[String])
       
    50 
       
    51   // internal version
       
    52   sealed case class Session_Info(
       
    53     select: Boolean,
       
    54     pos: Position.T,
       
    55     groups: List[String],
       
    56     dir: Path,
       
    57     parent: Option[String],
       
    58     description: String,
       
    59     options: Options,
       
    60     theories: List[(Options, List[Path])],
       
    61     files: List[Path],
       
    62     entry_digest: SHA1.Digest)
       
    63 
       
    64   def is_pure(name: String): Boolean = name == "RAW" || name == "Pure"
       
    65 
       
    66   def session_info(options: Options, select: Boolean, dir: Path, entry: Session_Entry)
       
    67       : (String, Session_Info) =
       
    68     try {
       
    69       val name = entry.name
       
    70 
       
    71       if (name == "") error("Bad session name")
       
    72       if (is_pure(name) && entry.parent.isDefined) error("Illegal parent session")
       
    73       if (!is_pure(name) && !entry.parent.isDefined) error("Missing parent session")
       
    74 
       
    75       val session_options = options ++ entry.options
       
    76 
       
    77       val theories =
       
    78         entry.theories.map({ case (opts, thys) =>
       
    79           (session_options ++ opts, thys.map(Path.explode(_))) })
       
    80       val files = entry.files.map(Path.explode(_))
       
    81       val entry_digest = SHA1.digest((name, entry.parent, entry.options, entry.theories).toString)
       
    82 
       
    83       val info =
       
    84         Session_Info(select, entry.pos, entry.groups, dir + Path.explode(entry.path),
       
    85           entry.parent, entry.description, session_options, theories, files, entry_digest)
       
    86 
       
    87       (name, info)
       
    88     }
       
    89     catch {
       
    90       case ERROR(msg) =>
       
    91         error(msg + "\nThe error(s) above occurred in session entry " +
       
    92           quote(entry.name) + Position.here(entry.pos))
       
    93     }
       
    94 
       
    95 
       
    96   /* session tree */
       
    97 
       
    98   object Session_Tree
       
    99   {
       
   100     def apply(infos: Seq[(String, Session_Info)]): Session_Tree =
       
   101     {
       
   102       val graph1 =
       
   103         (Graph.string[Session_Info] /: infos) {
       
   104           case (graph, (name, info)) =>
       
   105             if (graph.defined(name))
       
   106               error("Duplicate session " + quote(name) + Position.here(info.pos))
       
   107             else graph.new_node(name, info)
       
   108         }
       
   109       val graph2 =
       
   110         (graph1 /: graph1.entries) {
       
   111           case (graph, (name, (info, _))) =>
       
   112             info.parent match {
       
   113               case None => graph
       
   114               case Some(parent) =>
       
   115                 if (!graph.defined(parent))
       
   116                   error("Bad parent session " + quote(parent) + " for " +
       
   117                     quote(name) + Position.here(info.pos))
       
   118 
       
   119                 try { graph.add_edge_acyclic(parent, name) }
       
   120                 catch {
       
   121                   case exn: Graph.Cycles[_] =>
       
   122                     error(cat_lines(exn.cycles.map(cycle =>
       
   123                       "Cyclic session dependency of " +
       
   124                         cycle.map(c => quote(c.toString)).mkString(" via "))) +
       
   125                           Position.here(info.pos))
       
   126                 }
       
   127             }
       
   128         }
       
   129       new Session_Tree(graph2)
       
   130     }
       
   131   }
       
   132 
       
   133   final class Session_Tree private(val graph: Graph[String, Session_Info])
       
   134     extends PartialFunction[String, Session_Info]
       
   135   {
       
   136     def apply(name: String): Session_Info = graph.get_node(name)
       
   137     def isDefinedAt(name: String): Boolean = graph.defined(name)
       
   138 
       
   139     def selection(requirements: Boolean, all_sessions: Boolean,
       
   140       session_groups: List[String], sessions: List[String]): (List[String], Session_Tree) =
       
   141     {
       
   142       val bad_sessions = sessions.filterNot(isDefinedAt(_))
       
   143       if (!bad_sessions.isEmpty) error("Undefined session(s): " + commas_quote(bad_sessions))
       
   144 
       
   145       val pre_selected =
       
   146       {
       
   147         if (all_sessions) graph.keys.toList
       
   148         else {
       
   149           val select_group = session_groups.toSet
       
   150           val select = sessions.toSet
       
   151           (for {
       
   152             (name, (info, _)) <- graph.entries
       
   153             if info.select || select(name) || apply(name).groups.exists(select_group)
       
   154           } yield name).toList
       
   155         }
       
   156       }
       
   157       val selected =
       
   158         if (requirements) (graph.all_preds(pre_selected).toSet -- pre_selected).toList
       
   159         else pre_selected
       
   160 
       
   161       val tree1 = new Session_Tree(graph.restrict(graph.all_preds(selected).toSet))
       
   162       (selected, tree1)
       
   163     }
       
   164 
       
   165     def topological_order: List[(String, Session_Info)] =
       
   166       graph.topological_order.map(name => (name, apply(name)))
       
   167 
       
   168     override def toString: String = graph.entries.map(_._1).toList.sorted.mkString(",")
       
   169   }
       
   170 
       
   171 
       
   172   /* parser */
       
   173 
       
   174   private val SESSION = "session"
       
   175   private val IN = "in"
       
   176   private val DESCRIPTION = "description"
       
   177   private val OPTIONS = "options"
       
   178   private val THEORIES = "theories"
       
   179   private val FILES = "files"
       
   180 
       
   181   lazy val root_syntax =
       
   182     Outer_Syntax.init() + "(" + ")" + "+" + "," + "=" + "[" + "]" +
       
   183       (SESSION, Keyword.THY_DECL) + IN + DESCRIPTION + OPTIONS + THEORIES + FILES
       
   184 
       
   185   private object Parser extends Parse.Parser
       
   186   {
       
   187     def session_entry(pos: Position.T): Parser[Session_Entry] =
       
   188     {
       
   189       val session_name = atom("session name", _.is_name)
       
   190 
       
   191       val option =
       
   192         name ~ opt(keyword("=") ~! name ^^ { case _ ~ x => x }) ^^ { case x ~ y => (x, y) }
       
   193       val options = keyword("[") ~> rep1sep(option, keyword(",")) <~ keyword("]")
       
   194 
       
   195       val theories =
       
   196         keyword(THEORIES) ~! ((options | success(Nil)) ~ rep(theory_name)) ^^
       
   197           { case _ ~ (x ~ y) => (x, y) }
       
   198 
       
   199       command(SESSION) ~!
       
   200         (session_name ~
       
   201           ((keyword("(") ~! (rep1(name) <~ keyword(")")) ^^ { case _ ~ x => x }) | success(Nil)) ~
       
   202           ((keyword(IN) ~! path ^^ { case _ ~ x => x }) | success(".")) ~
       
   203           (keyword("=") ~!
       
   204             (opt(session_name ~! keyword("+") ^^ { case x ~ _ => x }) ~
       
   205               ((keyword(DESCRIPTION) ~! text ^^ { case _ ~ x => x }) | success("")) ~
       
   206               ((keyword(OPTIONS) ~! options ^^ { case _ ~ x => x }) | success(Nil)) ~
       
   207               rep1(theories) ~
       
   208               ((keyword(FILES) ~! rep1(path) ^^ { case _ ~ x => x }) | success(Nil))))) ^^
       
   209         { case _ ~ (a ~ b ~ c ~ (_ ~ (d ~ e ~ f ~ g ~ h))) =>
       
   210             Session_Entry(pos, a, b, c, d, e, f, g, h) }
       
   211     }
       
   212 
       
   213     def parse_entries(root: Path): List[Session_Entry] =
       
   214     {
       
   215       val toks = root_syntax.scan(File.read(root))
       
   216       parse_all(rep(session_entry(root.position)), Token.reader(toks, root.implode)) match {
       
   217         case Success(result, _) => result
       
   218         case bad => error(bad.toString)
       
   219       }
       
   220     }
       
   221   }
       
   222 
       
   223 
       
   224   /* find sessions within certain directories */
       
   225 
       
   226   private val ROOT = Path.explode("ROOT")
       
   227   private val ROOTS = Path.explode("ROOTS")
       
   228 
       
   229   private def is_session_dir(dir: Path): Boolean =
       
   230     (dir + ROOT).is_file || (dir + ROOTS).is_file
       
   231 
       
   232   private def check_session_dir(dir: Path): Path =
       
   233     if (is_session_dir(dir)) dir
       
   234     else error("Bad session root directory: " + dir.toString)
       
   235 
       
   236   def find_sessions(options: Options, more_dirs: List[(Boolean, Path)]): Session_Tree =
       
   237   {
       
   238     def find_dir(select: Boolean, dir: Path): List[(String, Session_Info)] =
       
   239       find_root(select, dir) ::: find_roots(select, dir)
       
   240 
       
   241     def find_root(select: Boolean, dir: Path): List[(String, Session_Info)] =
       
   242     {
       
   243       val root = dir + ROOT
       
   244       if (root.is_file) Parser.parse_entries(root).map(session_info(options, select, dir, _))
       
   245       else Nil
       
   246     }
       
   247 
       
   248     def find_roots(select: Boolean, dir: Path): List[(String, Session_Info)] =
       
   249     {
       
   250       val roots = dir + ROOTS
       
   251       if (roots.is_file) {
       
   252         for {
       
   253           line <- split_lines(File.read(roots))
       
   254           if !(line == "" || line.startsWith("#"))
       
   255           dir1 =
       
   256             try { check_session_dir(dir + Path.explode(line)) }
       
   257             catch {
       
   258               case ERROR(msg) =>
       
   259                 error(msg + "\nThe error(s) above occurred in session catalog " + roots.toString)
       
   260             }
       
   261           info <- find_dir(select, dir1)
       
   262         } yield info
       
   263       }
       
   264       else Nil
       
   265     }
       
   266 
       
   267     val default_dirs = Isabelle_System.components().filter(is_session_dir(_)).map((false, _))
       
   268 
       
   269     more_dirs foreach { case (_, dir) => check_session_dir(dir) }
       
   270 
       
   271     Session_Tree(
       
   272       for {
       
   273         (select, dir) <- default_dirs ::: more_dirs
       
   274         info <- find_dir(select, dir)
       
   275       } yield info)
       
   276   }
       
   277 
       
   278 
       
   279 
       
   280   /** build **/
       
   281 
       
   282   /* queue */
       
   283 
       
   284   object Queue
       
   285   {
       
   286     def apply(tree: Session_Tree): Queue =
       
   287     {
       
   288       val graph = tree.graph
       
   289 
       
   290       def outdegree(name: String): Int = graph.imm_succs(name).size
       
   291       def timeout(name: String): Double = tree(name).options.real("timeout")
       
   292 
       
   293       object Ordering extends scala.math.Ordering[String]
       
   294       {
       
   295         def compare(name1: String, name2: String): Int =
       
   296           outdegree(name2) compare outdegree(name1) match {
       
   297             case 0 =>
       
   298               timeout(name2) compare timeout(name1) match {
       
   299                 case 0 => name1 compare name2
       
   300                 case ord => ord
       
   301               }
       
   302             case ord => ord
       
   303           }
       
   304       }
       
   305 
       
   306       new Queue(graph, SortedSet(graph.keys.toList: _*)(Ordering))
       
   307     }
       
   308   }
       
   309 
       
   310   final class Queue private(graph: Graph[String, Session_Info], order: SortedSet[String])
       
   311   {
       
   312     def is_inner(name: String): Boolean = !graph.is_maximal(name)
       
   313 
       
   314     def is_empty: Boolean = graph.is_empty
       
   315 
       
   316     def - (name: String): Queue = new Queue(graph.del_node(name), order - name)
       
   317 
       
   318     def dequeue(skip: String => Boolean): Option[(String, Session_Info)] =
       
   319     {
       
   320       val it = order.iterator.dropWhile(name => skip(name) || !graph.is_minimal(name))
       
   321       if (it.hasNext) { val name = it.next; Some((name, graph.get_node(name))) }
       
   322       else None
       
   323     }
       
   324   }
       
   325 
       
   326 
       
   327   /* source dependencies and static content */
       
   328 
       
   329   sealed case class Session_Content(
       
   330     loaded_theories: Set[String],
       
   331     syntax: Outer_Syntax,
       
   332     sources: List[(Path, SHA1.Digest)],
       
   333     errors: List[String])
       
   334   {
       
   335     def check_errors: Session_Content =
       
   336     {
       
   337       if (errors.isEmpty) this
       
   338       else error(cat_lines(errors))
       
   339     }
       
   340   }
       
   341 
       
   342   sealed case class Deps(deps: Map[String, Session_Content])
       
   343   {
       
   344     def is_empty: Boolean = deps.isEmpty
       
   345     def apply(name: String): Session_Content = deps(name)
       
   346     def sources(name: String): List[SHA1.Digest] = deps(name).sources.map(_._2)
       
   347   }
       
   348 
       
   349   def dependencies(progress: Build.Progress, inlined_files: Boolean,
       
   350       verbose: Boolean, list_files: Boolean, tree: Session_Tree): Deps =
       
   351     Deps((Map.empty[String, Session_Content] /: tree.topological_order)(
       
   352       { case (deps, (name, info)) =>
       
   353           val (preloaded, parent_syntax, parent_errors) =
       
   354             info.parent match {
       
   355               case None =>
       
   356                 (Set.empty[String], Outer_Syntax.init_pure(), Nil)
       
   357               case Some(parent_name) =>
       
   358                 val parent = deps(parent_name)
       
   359                 (parent.loaded_theories, parent.syntax, parent.errors)
       
   360             }
       
   361           val thy_info = new Thy_Info(new Thy_Load(preloaded, parent_syntax))
       
   362 
       
   363           if (verbose || list_files) {
       
   364             val groups =
       
   365               if (info.groups.isEmpty) ""
       
   366               else info.groups.mkString(" (", " ", ")")
       
   367             progress.echo("Session " + name + groups)
       
   368           }
       
   369 
       
   370           val thy_deps =
       
   371             thy_info.dependencies(inlined_files,
       
   372               info.theories.map(_._2).flatten.
       
   373                 map(thy => Thy_Load.path_node_name(info.dir + Thy_Load.thy_path(thy))))
       
   374 
       
   375           val loaded_theories = thy_deps.loaded_theories
       
   376           val syntax = thy_deps.make_syntax
       
   377 
       
   378           val all_files =
       
   379             (thy_deps.deps.map({ case dep =>
       
   380               val thy = Path.explode(dep.name.node)
       
   381               val uses = dep.join_header.uses.map(p =>
       
   382                 Path.explode(dep.name.dir) + Path.explode(p._1))
       
   383               thy :: uses
       
   384             }).flatten ::: info.files.map(file => info.dir + file)).map(_.expand)
       
   385 
       
   386           if (list_files)
       
   387             progress.echo(cat_lines(all_files.map(_.implode).sorted.map("  " + _)))
       
   388 
       
   389           val sources =
       
   390             try { all_files.map(p => (p, SHA1.digest(p.file))) }
       
   391             catch {
       
   392               case ERROR(msg) =>
       
   393                 error(msg + "\nThe error(s) above occurred in session " +
       
   394                   quote(name) + Position.here(info.pos))
       
   395             }
       
   396           val errors = parent_errors ::: thy_deps.deps.map(dep => dep.join_header.errors).flatten
       
   397 
       
   398           deps + (name -> Session_Content(loaded_theories, syntax, sources, errors))
       
   399       }))
       
   400 
       
   401   def session_content(inlined_files: Boolean, dirs: List[Path], session: String): Session_Content =
       
   402   {
       
   403     val options = Options.init()
       
   404     val (_, tree) =
       
   405       find_sessions(options, dirs.map((false, _))).selection(false, false, Nil, List(session))
       
   406     dependencies(Build.Ignore_Progress, inlined_files, false, false, tree)(session)
       
   407   }
       
   408 
       
   409   def outer_syntax(session: String): Outer_Syntax =
       
   410     session_content(false, Nil, session).check_errors.syntax
       
   411 
       
   412 
       
   413   /* jobs */
       
   414 
       
   415   private class Job(name: String, info: Session_Info, output: Path, do_output: Boolean,
       
   416     verbose: Boolean, browser_info: Path)
       
   417   {
       
   418     // global browser info dir
       
   419     if (info.options.bool("browser_info") && !(browser_info + Path.explode("index.html")).is_file)
       
   420     {
       
   421       browser_info.file.mkdirs()
       
   422       File.copy(Path.explode("~~/lib/logo/isabelle.gif"),
       
   423         browser_info + Path.explode("isabelle.gif"))
       
   424       File.write(browser_info + Path.explode("index.html"),
       
   425         File.read(Path.explode("~~/lib/html/library_index_header.template")) +
       
   426         File.read(Path.explode("~~/lib/html/library_index_content.template")) +
       
   427         File.read(Path.explode("~~/lib/html/library_index_footer.template")))
       
   428     }
       
   429 
       
   430     def output_path: Option[Path] = if (do_output) Some(output) else None
       
   431 
       
   432     private val parent = info.parent.getOrElse("")
       
   433 
       
   434     private val args_file = File.tmp_file("args")
       
   435     File.write(args_file, YXML.string_of_body(
       
   436       if (is_pure(name)) Options.encode(info.options)
       
   437       else
       
   438         {
       
   439           import XML.Encode._
       
   440               pair(bool, pair(Options.encode, pair(bool, pair(Path.encode, pair(string,
       
   441                 pair(string, list(pair(Options.encode, list(Path.encode)))))))))(
       
   442               (do_output, (info.options, (verbose, (browser_info, (parent,
       
   443                 (name, info.theories)))))))
       
   444         }))
       
   445 
       
   446     private val env =
       
   447       Map("INPUT" -> parent, "TARGET" -> name, "OUTPUT" -> Isabelle_System.standard_path(output),
       
   448         (if (is_pure(name)) "ISABELLE_PROCESS_OPTIONS" else "ARGS_FILE") ->
       
   449           Isabelle_System.posix_path(args_file.getPath))
       
   450 
       
   451     private val script =
       
   452       if (is_pure(name)) {
       
   453         if (do_output) "./build " + name + " \"$OUTPUT\""
       
   454         else """ rm -f "$OUTPUT"; ./build """ + name
       
   455       }
       
   456       else {
       
   457         """
       
   458         . "$ISABELLE_HOME/lib/scripts/timestart.bash"
       
   459         """ +
       
   460           (if (do_output)
       
   461             """
       
   462             "$ISABELLE_PROCESS" -e "Build.build \"$ARGS_FILE\";" -q -w "$INPUT" "$OUTPUT"
       
   463             """
       
   464           else
       
   465             """
       
   466             rm -f "$OUTPUT"; "$ISABELLE_PROCESS" -e "Build.build \"$ARGS_FILE\";" -r -q "$INPUT"
       
   467             """) +
       
   468         """
       
   469         RC="$?"
       
   470 
       
   471         . "$ISABELLE_HOME/lib/scripts/timestop.bash"
       
   472 
       
   473         if [ "$RC" -eq 0 ]; then
       
   474           echo "Finished $TARGET ($TIMES_REPORT)" >&2
       
   475         fi
       
   476 
       
   477         exit "$RC"
       
   478         """
       
   479       }
       
   480 
       
   481     private val (thread, result) =
       
   482       Simple_Thread.future("build") { Isabelle_System.bash_env(info.dir.file, env, script) }
       
   483 
       
   484     def terminate: Unit = thread.interrupt
       
   485     def is_finished: Boolean = result.is_finished
       
   486 
       
   487     @volatile private var timeout = false
       
   488     private val time = info.options.seconds("timeout")
       
   489     private val timer: Option[Timer] =
       
   490       if (time.seconds > 0.0) {
       
   491         val t = new Timer("build", true)
       
   492         t.schedule(new TimerTask { def run = { terminate; timeout = true } }, time.ms)
       
   493         Some(t)
       
   494       }
       
   495       else None
       
   496 
       
   497     def join: (String, String, Int) = {
       
   498       val (out, err, rc) = result.join
       
   499       args_file.delete
       
   500       timer.map(_.cancel())
       
   501 
       
   502       val err1 =
       
   503         if (rc == 130)
       
   504           (if (err.isEmpty || err.endsWith("\n")) err else err + "\n") +
       
   505           (if (timeout) "*** Timeout\n" else "*** Interrupt\n")
       
   506         else err
       
   507       (out, err1, rc)
       
   508     }
       
   509   }
       
   510 
       
   511 
       
   512   /* log files and corresponding heaps */
       
   513 
       
   514   private val LOG = Path.explode("log")
       
   515   private def log(name: String): Path = LOG + Path.basic(name)
       
   516   private def log_gz(name: String): Path = log(name).ext("gz")
       
   517 
       
   518   private def sources_stamp(digests: List[SHA1.Digest]): String =
       
   519     digests.map(_.toString).sorted.mkString("sources: ", " ", "")
       
   520 
       
   521   private val no_heap: String = "heap: -"
       
   522 
       
   523   private def heap_stamp(heap: Option[Path]): String =
       
   524   {
       
   525     "heap: " +
       
   526       (heap match {
       
   527         case Some(path) =>
       
   528           val file = path.file
       
   529           if (file.isFile) file.length.toString + " " + file.lastModified.toString
       
   530           else "-"
       
   531         case None => "-"
       
   532       })
       
   533   }
       
   534 
       
   535   private def read_stamps(path: Path): Option[(String, String, String)] =
       
   536     if (path.is_file) {
       
   537       val stream = new GZIPInputStream (new BufferedInputStream(new FileInputStream(path.file)))
       
   538       val reader = new BufferedReader(new InputStreamReader(stream, UTF8.charset))
       
   539       val (s, h1, h2) =
       
   540         try { (reader.readLine, reader.readLine, reader.readLine) }
       
   541         finally { reader.close }
       
   542       if (s != null && s.startsWith("sources: ") &&
       
   543           h1 != null && h1.startsWith("heap: ") &&
       
   544           h2 != null && h2.startsWith("heap: ")) Some((s, h1, h2))
       
   545       else None
       
   546     }
       
   547     else None
       
   548 
       
   549 
       
   550   /* build */
       
   551 
       
   552   def build(
       
   553     progress: Build.Progress,
       
   554     options: Options,
       
   555     requirements: Boolean = false,
       
   556     all_sessions: Boolean = false,
       
   557     build_heap: Boolean = false,
       
   558     clean_build: Boolean = false,
       
   559     more_dirs: List[(Boolean, Path)] = Nil,
       
   560     session_groups: List[String] = Nil,
       
   561     max_jobs: Int = 1,
       
   562     list_files: Boolean = false,
       
   563     no_build: Boolean = false,
       
   564     system_mode: Boolean = false,
       
   565     verbose: Boolean = false,
       
   566     sessions: List[String] = Nil): Int =
       
   567   {
       
   568     val full_tree = find_sessions(options, more_dirs)
       
   569     val (selected, selected_tree) =
       
   570       full_tree.selection(requirements, all_sessions, session_groups, sessions)
       
   571 
       
   572     val deps = dependencies(progress, true, verbose, list_files, selected_tree)
       
   573     val queue = Queue(selected_tree)
       
   574 
       
   575     def make_stamp(name: String): String =
       
   576       sources_stamp(selected_tree(name).entry_digest :: deps.sources(name))
       
   577 
       
   578     val (input_dirs, output_dir, browser_info) =
       
   579       if (system_mode) {
       
   580         val output_dir = Path.explode("~~/heaps/$ML_IDENTIFIER")
       
   581         (List(output_dir), output_dir, Path.explode("~~/browser_info"))
       
   582       }
       
   583       else {
       
   584         val output_dir = Path.explode("$ISABELLE_OUTPUT")
       
   585         (output_dir :: Isabelle_System.find_logics_dirs(), output_dir,
       
   586          Path.explode("$ISABELLE_BROWSER_INFO"))
       
   587       }
       
   588 
       
   589     // prepare log dir
       
   590     (output_dir + LOG).file.mkdirs()
       
   591 
       
   592     // optional cleanup
       
   593     if (clean_build) {
       
   594       for (name <- full_tree.graph.all_succs(selected)) {
       
   595         val files =
       
   596           List(Path.basic(name), log(name), log_gz(name)).map(output_dir + _).filter(_.is_file)
       
   597         if (!files.isEmpty) progress.echo("Cleaning " + name + " ...")
       
   598         if (!files.forall(p => p.file.delete)) progress.echo(name + " FAILED to delete")
       
   599       }
       
   600     }
       
   601 
       
   602     // scheduler loop
       
   603     case class Result(current: Boolean, heap: String, rc: Int)
       
   604 
       
   605     def sleep(): Unit = Thread.sleep(500)
       
   606 
       
   607     @tailrec def loop(
       
   608       pending: Queue,
       
   609       running: Map[String, (String, Job)],
       
   610       results: Map[String, Result]): Map[String, Result] =
       
   611     {
       
   612       if (pending.is_empty) results
       
   613       else if (progress.stopped) {
       
   614         for ((_, (_, job)) <- running) job.terminate
       
   615         sleep(); loop(pending, running, results)
       
   616       }
       
   617       else
       
   618         running.find({ case (_, (_, job)) => job.is_finished }) match {
       
   619           case Some((name, (parent_heap, job))) =>
       
   620             //{{{ finish job
       
   621 
       
   622             val (out, err, rc) = job.join
       
   623             progress.echo(Library.trim_line(err))
       
   624 
       
   625             val heap =
       
   626               if (rc == 0) {
       
   627                 (output_dir + log(name)).file.delete
       
   628 
       
   629                 val sources = make_stamp(name)
       
   630                 val heap = heap_stamp(job.output_path)
       
   631                 File.write_gzip(output_dir + log_gz(name),
       
   632                   sources + "\n" + parent_heap + "\n" + heap + "\n" + out)
       
   633 
       
   634                 heap
       
   635               }
       
   636               else {
       
   637                 (output_dir + Path.basic(name)).file.delete
       
   638                 (output_dir + log_gz(name)).file.delete
       
   639 
       
   640                 File.write(output_dir + log(name), out)
       
   641                 progress.echo(name + " FAILED")
       
   642                 if (rc != 130) {
       
   643                   progress.echo("(see also " + (output_dir + log(name)).file.toString + ")")
       
   644                   val lines = split_lines(out)
       
   645                   val tail = lines.drop(lines.length - 20 max 0)
       
   646                   progress.echo("\n" + cat_lines(tail))
       
   647                 }
       
   648 
       
   649                 no_heap
       
   650               }
       
   651             loop(pending - name, running - name, results + (name -> Result(false, heap, rc)))
       
   652             //}}}
       
   653           case None if (running.size < (max_jobs max 1)) =>
       
   654             //{{{ check/start next job
       
   655             pending.dequeue(running.isDefinedAt(_)) match {
       
   656               case Some((name, info)) =>
       
   657                 val parent_result =
       
   658                   info.parent match {
       
   659                     case None => Result(true, no_heap, 0)
       
   660                     case Some(parent) => results(parent)
       
   661                   }
       
   662                 val output = output_dir + Path.basic(name)
       
   663                 val do_output = build_heap || queue.is_inner(name)
       
   664 
       
   665                 val (current, heap) =
       
   666                 {
       
   667                   input_dirs.find(dir => (dir + log_gz(name)).is_file) match {
       
   668                     case Some(dir) =>
       
   669                       read_stamps(dir + log_gz(name)) match {
       
   670                         case Some((s, h1, h2)) =>
       
   671                           val heap = heap_stamp(Some(dir + Path.basic(name)))
       
   672                           (s == make_stamp(name) && h1 == parent_result.heap && h2 == heap &&
       
   673                             !(do_output && heap == no_heap), heap)
       
   674                         case None => (false, no_heap)
       
   675                       }
       
   676                     case None => (false, no_heap)
       
   677                   }
       
   678                 }
       
   679                 val all_current = current && parent_result.current
       
   680 
       
   681                 if (all_current)
       
   682                   loop(pending - name, running, results + (name -> Result(true, heap, 0)))
       
   683                 else if (no_build) {
       
   684                   if (verbose) progress.echo("Skipping " + name + " ...")
       
   685                   loop(pending - name, running, results + (name -> Result(false, heap, 1)))
       
   686                 }
       
   687                 else if (parent_result.rc == 0) {
       
   688                   progress.echo((if (do_output) "Building " else "Running ") + name + " ...")
       
   689                   val job = new Job(name, info, output, do_output, verbose, browser_info)
       
   690                   loop(pending, running + (name -> (parent_result.heap, job)), results)
       
   691                 }
       
   692                 else {
       
   693                   progress.echo(name + " CANCELLED")
       
   694                   loop(pending - name, running, results + (name -> Result(false, heap, 1)))
       
   695                 }
       
   696               case None => sleep(); loop(pending, running, results)
       
   697             }
       
   698             ///}}}
       
   699           case None => sleep(); loop(pending, running, results)
       
   700         }
       
   701     }
       
   702 
       
   703     val results =
       
   704       if (deps.is_empty) {
       
   705         progress.echo("### Nothing to build")
       
   706         Map.empty
       
   707       }
       
   708       else loop(queue, Map.empty, Map.empty)
       
   709 
       
   710     val rc = (0 /: results)({ case (rc1, (_, res)) => rc1 max res.rc })
       
   711     if (rc != 0 && (verbose || !no_build)) {
       
   712       val unfinished =
       
   713         (for ((name, res) <- results.iterator if res.rc != 0) yield name).toList
       
   714           .sorted(scala.math.Ordering.String)  // FIXME scala-2.10.0-RC1
       
   715       progress.echo("Unfinished session(s): " + commas(unfinished))
       
   716     }
       
   717     rc
       
   718   }
       
   719 
       
   720 
       
   721   /* command line entry point */
       
   722 
       
   723   def main(args: Array[String])
       
   724   {
       
   725     Command_Line.tool {
       
   726       args.toList match {
       
   727         case
       
   728           Properties.Value.Boolean(requirements) ::
       
   729           Properties.Value.Boolean(all_sessions) ::
       
   730           Properties.Value.Boolean(build_heap) ::
       
   731           Properties.Value.Boolean(clean_build) ::
       
   732           Properties.Value.Int(max_jobs) ::
       
   733           Properties.Value.Boolean(list_files) ::
       
   734           Properties.Value.Boolean(no_build) ::
       
   735           Properties.Value.Boolean(system_mode) ::
       
   736           Properties.Value.Boolean(verbose) ::
       
   737           Command_Line.Chunks(select_dirs, include_dirs, session_groups, build_options, sessions) =>
       
   738             val options = (Options.init() /: build_options)(_ + _)
       
   739             val dirs =
       
   740               select_dirs.map(d => (true, Path.explode(d))) :::
       
   741               include_dirs.map(d => (false, Path.explode(d)))
       
   742             build(Build.Console_Progress, options, requirements, all_sessions, build_heap,
       
   743               clean_build, dirs, session_groups, max_jobs, list_files, no_build, system_mode,
       
   744               verbose, sessions)
       
   745         case _ => error("Bad arguments:\n" + cat_lines(args))
       
   746       }
       
   747     }
       
   748   }
       
   749 }
       
   750