diff --git a/src/Pure/System/isabelle_system.scala b/src/Pure/System/isabelle_system.scala --- a/src/Pure/System/isabelle_system.scala +++ b/src/Pure/System/isabelle_system.scala @@ -1,443 +1,452 @@ /* Title: Pure/System/isabelle_system.scala Author: Makarius Fundamental Isabelle system environment: quasi-static module with optional init operation. */ package isabelle import java.lang.System import java.util.regex.Pattern import java.io.{InputStream, OutputStream, File => JFile, BufferedReader, InputStreamReader, BufferedWriter, OutputStreamWriter, IOException, FileInputStream, BufferedInputStream} import java.awt.{GraphicsEnvironment, Font} import java.awt.font.TextAttribute import javax.swing.ImageIcon import scala.io.Source import scala.util.matching.Regex object Isabelle_System { /** bootstrap information **/ def jdk_home(): String = { val java_home = System.getProperty("java.home") val home = new JFile(java_home) val parent = home.getParent if (home.getName == "jre" && parent != null && (new JFile(new JFile(parent, "bin"), "javac")).exists) parent else java_home } def cygwin_root(): String = { require(Platform.is_windows) val cygwin_root1 = System.getenv("CYGWIN_ROOT") val cygwin_root2 = System.getProperty("cygwin.root") if (cygwin_root1 != null && cygwin_root1 != "") cygwin_root1 else if (cygwin_root2 != null && cygwin_root2 != "") cygwin_root2 else error("Cannot determine Cygwin root directory") } /** implicit settings environment **/ @volatile private var _settings: Option[Map[String, String]] = None def settings(): Map[String, String] = { if (_settings.isEmpty) init() // unsynchronized check _settings.get } /* isabelle_home precedence: (1) this_isabelle_home as explicit argument (2) ISABELLE_HOME process environment variable (e.g. inherited from running isabelle tool) (3) isabelle.home system property (e.g. via JVM application boot process) */ def init(this_isabelle_home: String = null): Unit = synchronized { if (_settings.isEmpty) { import scala.collection.JavaConversions._ val settings = { val env0 = sys.env + ("ISABELLE_JDK_HOME" -> posix_path(jdk_home())) val user_home = System.getProperty("user.home") val env = if (user_home == null || env0.isDefinedAt("HOME")) env0 else env0 + ("HOME" -> user_home) val isabelle_home = if (this_isabelle_home != null) this_isabelle_home else env.get("ISABELLE_HOME") match { case None | Some("") => val path = System.getProperty("isabelle.home") if (path == null || path == "") error("Unknown Isabelle home directory") else path case Some(path) => path } File.with_tmp_file("settings") { dump => val shell_prefix = if (Platform.is_windows) List(cygwin_root() + "\\bin\\bash", "-l") else Nil val cmdline = shell_prefix ::: List(isabelle_home + "/bin/isabelle", "getenv", "-d", dump.toString) val (output, rc) = process_output(raw_execute(null, env, true, cmdline: _*)) if (rc != 0) error(output) val entries = (for (entry <- File.read(dump) split "\0" if entry != "") yield { val i = entry.indexOf('=') if (i <= 0) (entry -> "") else (entry.substring(0, i) -> entry.substring(i + 1)) }).toMap entries + ("PATH" -> entries("PATH_JVM")) - "PATH_JVM" } } _settings = Some(settings) } } /* getenv */ def getenv(name: String): String = settings.getOrElse(name, "") def getenv_strict(name: String): String = { val value = getenv(name) if (value != "") value else error("Undefined environment variable: " + name) } /** file-system operations **/ /* jvm_path */ private val Cygdrive = new Regex("/cygdrive/([a-zA-Z])($|/.*)") private val Named_Root = new Regex("//+([^/]*)(.*)") def jvm_path(posix_path: String): String = if (Platform.is_windows) { val result_path = new StringBuilder val rest = posix_path match { case Cygdrive(drive, rest) => result_path ++= (Library.uppercase(drive) + ":" + JFile.separator) rest case Named_Root(root, rest) => result_path ++= JFile.separator result_path ++= JFile.separator result_path ++= root rest case path if path.startsWith("/") => result_path ++= cygwin_root() path case path => path } for (p <- space_explode('/', rest) if p != "") { val len = result_path.length if (len > 0 && result_path(len - 1) != JFile.separatorChar) result_path += JFile.separatorChar result_path ++= p } result_path.toString } else posix_path /* posix_path */ def posix_path(jvm_path: String): String = if (Platform.is_windows) { val Platform_Root = new Regex("(?i)" + Pattern.quote(cygwin_root()) + """(?:\\+|\z)(.*)""") val Drive = new Regex("""([a-zA-Z]):\\*(.*)""") jvm_path.replace('/', '\\') match { case Platform_Root(rest) => "/" + rest.replace('\\', '/') case Drive(letter, rest) => "/cygdrive/" + Library.lowercase(letter) + (if (rest == "") "" else "/" + rest.replace('\\', '/')) case path => path.replace('\\', '/') } } else jvm_path def posix_path(file: JFile): String = posix_path(file.getPath) /* misc path specifications */ def standard_path(path: Path): String = path.expand.implode def platform_path(path: Path): String = jvm_path(standard_path(path)) def platform_file(path: Path): JFile = new JFile(platform_path(path)) def platform_file_url(raw_path: Path): String = { val path = raw_path.expand require(path.is_absolute) val s = platform_path(path).replaceAll(" ", "%20") if (!Platform.is_windows) "file://" + s else if (s.startsWith("\\\\")) "file:" + s.replace('\\', '/') else "file:///" + s.replace('\\', '/') } /* source files of Isabelle/ML bootstrap */ def source_file(path: Path): Option[Path] = { def check(p: Path): Option[Path] = if (p.is_file) Some(p) else None if (path.is_absolute || path.is_current) check(path) else { check(Path.explode("~~/src/Pure") + path) orElse (if (getenv("ML_SOURCES") == "") None else check(Path.explode("$ML_SOURCES") + path)) } } + /* mkdirs */ + + def mkdirs(path: Path) + { + path.file.mkdirs + if (!path.is_dir) error("Cannot create directory: " + quote(platform_path(path))) + } + + /** external processes **/ /* raw execute for bootstrapping */ private def raw_execute(cwd: JFile, env: Map[String, String], redirect: Boolean, args: String*) : Process = { val cmdline = new java.util.LinkedList[String] for (s <- args) cmdline.add(s) val proc = new ProcessBuilder(cmdline) if (cwd != null) proc.directory(cwd) if (env != null) { proc.environment.clear for ((x, y) <- env) proc.environment.put(x, y) } proc.redirectErrorStream(redirect) proc.start } private def process_output(proc: Process): (String, Int) = { proc.getOutputStream.close val output = File.read_stream(proc.getInputStream) val rc = try { proc.waitFor } finally { proc.getInputStream.close proc.getErrorStream.close proc.destroy Thread.interrupted } (output, rc) } /* plain execute */ def execute_env(cwd: JFile, env: Map[String, String], redirect: Boolean, args: String*): Process = { val cmdline = if (Platform.is_windows) List(cygwin_root() + "\\bin\\env.exe") ++ args else args val env1 = if (env == null) settings else settings ++ env raw_execute(cwd, env1, redirect, cmdline: _*) } def execute(redirect: Boolean, args: String*): Process = execute_env(null, null, redirect, args: _*) /* managed process */ class Managed_Process(cwd: JFile, env: Map[String, String], redirect: Boolean, args: String*) { private val params = List(standard_path(Path.explode("~~/lib/scripts/process")), "group", "-", "no_script") private val proc = execute_env(cwd, env, redirect, (params ++ args):_*) // channels val stdin: BufferedWriter = new BufferedWriter(new OutputStreamWriter(proc.getOutputStream, UTF8.charset)) val stdout: BufferedReader = new BufferedReader(new InputStreamReader(proc.getInputStream, UTF8.charset)) val stderr: BufferedReader = new BufferedReader(new InputStreamReader(proc.getErrorStream, UTF8.charset)) // signals private val pid = stdout.readLine private def kill(signal: String): Boolean = { execute(true, "kill", "-" + signal, "-" + pid).waitFor execute(true, "kill", "-0", "-" + pid).waitFor == 0 } private def multi_kill(signal: String): Boolean = { var running = true var count = 10 while (running && count > 0) { if (kill(signal)) { Thread.sleep(100) count -= 1 } else running = false } running } def interrupt() { multi_kill("INT") } def terminate() { multi_kill("INT") && multi_kill("TERM") && kill("KILL"); proc.destroy } // JVM shutdown hook private val shutdown_hook = new Thread { override def run = terminate() } try { Runtime.getRuntime.addShutdownHook(shutdown_hook) } catch { case _: IllegalStateException => } private def cleanup() = try { Runtime.getRuntime.removeShutdownHook(shutdown_hook) } catch { case _: IllegalStateException => } /* result */ def join: Int = { val rc = proc.waitFor; cleanup(); rc } } /* bash */ final case class Bash_Result(out_lines: List[String], err_lines: List[String], rc: Int) { def out: String = cat_lines(out_lines) def err: String = cat_lines(err_lines) def add_err(s: String): Bash_Result = copy(err_lines = err_lines ::: List(s)) } def bash_env(cwd: JFile, env: Map[String, String], script: String, out_progress: String => Unit = (_: String) => (), err_progress: String => Unit = (_: String) => ()): Bash_Result = { File.with_tmp_file("isabelle_script") { script_file => File.write(script_file, script) val proc = new Managed_Process(cwd, env, false, "bash", posix_path(script_file)) proc.stdin.close val (_, stdout) = Simple_Thread.future("bash_stdout") { File.read_lines(proc.stdout, out_progress) } val (_, stderr) = Simple_Thread.future("bash_stderr") { File.read_lines(proc.stderr, err_progress) } val rc = try { proc.join } catch { case e: InterruptedException => Thread.interrupted; proc.terminate; 130 } Bash_Result(stdout.join, stderr.join, rc) } } def bash(script: String): Bash_Result = bash_env(null, null, script) /* system tools */ def isabelle_tool(name: String, args: String*): (String, Int) = { Path.split(getenv_strict("ISABELLE_TOOLS")).find { dir => val file = (dir + Path.basic(name)).file try { file.isFile && file.canRead && file.canExecute && !name.endsWith("~") && !name.endsWith(".orig") } catch { case _: SecurityException => false } } match { case Some(dir) => val file = standard_path(dir + Path.basic(name)) process_output(execute(true, (List(file) ++ args): _*)) case None => ("Unknown Isabelle tool: " + name, 2) } } /** Isabelle resources **/ /* components */ def components(): List[Path] = Path.split(getenv_strict("ISABELLE_COMPONENTS")) /* logic images */ def find_logics_dirs(): List[Path] = { val ml_ident = Path.explode("$ML_IDENTIFIER").expand Path.split(getenv_strict("ISABELLE_PATH")).map(_ + ml_ident) } def find_logics(): List[String] = (for { dir <- find_logics_dirs() files = dir.file.listFiles() if files != null file <- files.toList if file.isFile } yield file.getName).sorted def default_logic(args: String*): String = { args.find(_ != "") match { case Some(logic) => logic case None => Isabelle_System.getenv_strict("ISABELLE_LOGIC") } } /* fonts */ def get_font(family: String = "IsabelleText", size: Int = 1, bold: Boolean = false): Font = new Font(family, if (bold) Font.BOLD else Font.PLAIN, size) def install_fonts() { val ge = GraphicsEnvironment.getLocalGraphicsEnvironment() for (font <- Path.split(getenv_strict("ISABELLE_FONTS"))) ge.registerFont(Font.createFont(Font.TRUETYPE_FONT, font.file)) } def install_fonts_jfx() { for (font <- Path.split(getenv_strict("ISABELLE_FONTS"))) { val stream = new BufferedInputStream(new FileInputStream(font.file)) try { javafx.scene.text.Font.loadFont(stream, 1.0) } finally { stream.close } } } /* icon */ def get_icon(): ImageIcon = new ImageIcon(platform_path(Path.explode("~~/lib/logo/isabelle.gif"))) } diff --git a/src/Pure/System/options.scala b/src/Pure/System/options.scala --- a/src/Pure/System/options.scala +++ b/src/Pure/System/options.scala @@ -1,427 +1,427 @@ /* Title: Pure/System/options.scala Author: Makarius Stand-alone options with external string representation. */ package isabelle import java.util.Calendar object Options { type Spec = (String, Option[String]) val empty: Options = new Options() /* representation */ sealed abstract class Type { def print: String = Library.lowercase(toString) } case object Bool extends Type case object Int extends Type case object Real extends Type case object String extends Type case object Unknown extends Type case class Opt(name: String, typ: Type, value: String, default_value: String, description: String, section: String) { private def print(default: Boolean): String = { val x = if (default) default_value else value "option " + name + " : " + typ.print + " = " + (if (typ == Options.String) quote(x) else x) + (if (description == "") "" else "\n -- " + quote(description)) } def print: String = print(false) def print_default: String = print(true) def title(strip: String): String = { val words = space_explode('_', name) val words1 = words match { case word :: rest if word == strip => rest case _ => words } words1.map(Library.capitalize).mkString(" ") } def unknown: Boolean = typ == Unknown } /* parsing */ private val SECTION = "section" private val OPTION = "option" private val OPTIONS = Path.explode("etc/options") private val PREFS_DIR = Path.explode("$ISABELLE_HOME_USER/etc") private val PREFS = PREFS_DIR + Path.basic("preferences") private val PREFS_BACKUP = PREFS_DIR + Path.basic("preferences~") lazy val options_syntax = Outer_Syntax.init() + ":" + "=" + "--" + (SECTION, Keyword.THY_HEADING2) + (OPTION, Keyword.THY_DECL) lazy val prefs_syntax = Outer_Syntax.init() + "=" object Parser extends Parse.Parser { val option_name = atom("option name", _.is_xname) val option_type = atom("option type", _.is_ident) val option_value = opt(token("-", tok => tok.is_sym_ident && tok.content == "-")) ~ atom("nat", _.is_nat) ^^ { case s ~ n => if (s.isDefined) "-" + n else n } | atom("option value", tok => tok.is_name || tok.is_float) val option_entry: Parser[Options => Options] = { command(SECTION) ~! text ^^ { case _ ~ a => (options: Options) => options.set_section(a) } | command(OPTION) ~! (option_name ~ keyword(":") ~ option_type ~ keyword("=") ~ option_value ~ (keyword("--") ~! text ^^ { case _ ~ x => x } | success(""))) ^^ { case _ ~ (a ~ _ ~ b ~ _ ~ c ~ d) => (options: Options) => options.declare(a, b, c, d) } } val prefs_entry: Parser[Options => Options] = { option_name ~ (keyword("=") ~! option_value) ^^ { case a ~ (_ ~ b) => (options: Options) => options.add_permissive(a, b) } } def parse_file(syntax: Outer_Syntax, parser: Parser[Options => Options], options: Options, file: Path): Options = { val toks = syntax.scan(File.read(file)) val ops = parse_all(rep(parser), Token.reader(toks, file.implode)) match { case Success(result, _) => result case bad => error(bad.toString) } try { (options.set_section("") /: ops) { case (opts, op) => op(opts) } } catch { case ERROR(msg) => error(msg + Position.here(file.position)) } } } def init_defaults(): Options = { var options = empty for { dir <- Isabelle_System.components() file = dir + OPTIONS if file.is_file } { options = Parser.parse_file(options_syntax, Parser.option_entry, options, file) } options } def init(): Options = init_defaults().load_prefs() /* encode */ val encode: XML.Encode.T[Options] = (options => options.encode) /* command line entry point */ def main(args: Array[String]) { Command_Line.tool { args.toList match { case export_file :: more_options => val options = (Options.init() /: more_options)(_ + _) if (export_file == "") java.lang.System.out.println(options.print) else File.write(Path.explode(export_file), YXML.string_of_body(options.encode)) 0 case _ => error("Bad arguments:\n" + cat_lines(args)) } } } } final class Options private( val options: Map[String, Options.Opt] = Map.empty, val section: String = "") { override def toString: String = options.iterator.mkString("Options (", ",", ")") def print: String = cat_lines(options.toList.sortBy(_._1).map(p => p._2.print)) def description(name: String): String = check_name(name).description /* check */ def check_name(name: String): Options.Opt = options.get(name) match { case Some(opt) if !opt.unknown => opt case _ => error("Unknown option " + quote(name)) } private def check_type(name: String, typ: Options.Type): Options.Opt = { val opt = check_name(name) if (opt.typ == typ) opt else error("Ill-typed option " + quote(name) + " : " + opt.typ.print + " vs. " + typ.print) } /* basic operations */ private def put[A](name: String, typ: Options.Type, value: String): Options = { val opt = check_type(name, typ) new Options(options + (name -> opt.copy(value = value)), section) } private def get[A](name: String, typ: Options.Type, parse: String => Option[A]): A = { val opt = check_type(name, typ) parse(opt.value) match { case Some(x) => x case None => error("Malformed value for option " + quote(name) + " : " + typ.print + " =\n" + quote(opt.value)) } } /* internal lookup and update */ class Bool_Access { def apply(name: String): Boolean = get(name, Options.Bool, Properties.Value.Boolean.unapply) def update(name: String, x: Boolean): Options = put(name, Options.Bool, Properties.Value.Boolean(x)) } val bool = new Bool_Access class Int_Access { def apply(name: String): Int = get(name, Options.Int, Properties.Value.Int.unapply) def update(name: String, x: Int): Options = put(name, Options.Int, Properties.Value.Int(x)) } val int = new Int_Access class Real_Access { def apply(name: String): Double = get(name, Options.Real, Properties.Value.Double.unapply) def update(name: String, x: Double): Options = put(name, Options.Real, Properties.Value.Double(x)) } val real = new Real_Access class String_Access { def apply(name: String): String = get(name, Options.String, s => Some(s)) def update(name: String, x: String): Options = put(name, Options.String, x) } val string = new String_Access class Seconds_Access { def apply(name: String): Time = Time.seconds(real(name)) } val seconds = new Seconds_Access /* external updates */ private def check_value(name: String): Options = { val opt = check_name(name) opt.typ match { case Options.Bool => bool(name); this case Options.Int => int(name); this case Options.Real => real(name); this case Options.String => string(name); this case Options.Unknown => this } } def declare(name: String, typ_name: String, value: String, description: String = ""): Options = { options.get(name) match { case Some(_) => error("Duplicate declaration of option " + quote(name)) case None => val typ = typ_name match { case "bool" => Options.Bool case "int" => Options.Int case "real" => Options.Real case "string" => Options.String case _ => error("Unknown type for option " + quote(name) + " : " + quote(typ_name)) } val opt = Options.Opt(name, typ, value, value, description, section) (new Options(options + (name -> opt), section)).check_value(name) } } def add_permissive(name: String, value: String): Options = { if (options.isDefinedAt(name)) this + (name, value) else new Options( options + (name -> Options.Opt(name, Options.Unknown, value, value, "", "")), section) } def + (name: String, value: String): Options = { val opt = check_name(name) (new Options(options + (name -> opt.copy(value = value)), section)).check_value(name) } def + (name: String, opt_value: Option[String]): Options = { val opt = check_name(name) opt_value match { case Some(value) => this + (name, value) case None if opt.typ == Options.Bool => this + (name, "true") case None => error("Missing value for option " + quote(name) + " : " + opt.typ.print) } } def + (str: String): Options = { str.indexOf('=') match { case -1 => this + (str, None) case i => this + (str.substring(0, i), str.substring(i + 1)) } } def ++ (specs: List[Options.Spec]): Options = (this /: specs)({ case (x, (y, z)) => x + (y, z) }) /* sections */ def set_section(new_section: String): Options = new Options(options, new_section) def sections: List[(String, List[Options.Opt])] = options.groupBy(_._2.section).toList.map({ case (a, opts) => (a, opts.toList.map(_._2)) }) /* encode */ def encode: XML.Body = { val opts = for ((name, opt) <- options.toList; if !opt.unknown) yield (name, opt.typ.print, opt.value) import XML.Encode.{string => str, _} list(triple(str, str, str))(opts) } /* user preferences */ def load_prefs(): Options = if (Options.PREFS.is_file) Options.Parser.parse_file( Options.prefs_syntax, Options.Parser.prefs_entry, this, Options.PREFS) else this def save_prefs() { val defaults = Options.init_defaults() val changed = (for { (name, opt2) <- options.iterator opt1 = defaults.options.get(name) if (opt1.isEmpty || opt1.get.value != opt2.value) } yield (name, opt2.value, if (opt1.isEmpty) " (* unknown *)" else "")).toList val prefs = changed.sortBy(_._1) .map({ case (x, y, z) => x + " = " + Outer_Syntax.quote_string(y) + z + "\n" }).mkString - Options.PREFS_DIR.file.mkdirs() + Isabelle_System.mkdirs(Options.PREFS_DIR) Options.PREFS.file renameTo Options.PREFS_BACKUP.file File.write(Options.PREFS, "(* generated by Isabelle " + Calendar.getInstance.getTime + " *)\n\n" + prefs) } } class Options_Variable { // owned by Swing thread @volatile private var options = Options.empty def value: Options = options def update(new_options: Options) { Swing_Thread.require() options = new_options } def + (name: String, x: String) { Swing_Thread.require() options = options + (name, x) } class Bool_Access { def apply(name: String): Boolean = options.bool(name) def update(name: String, x: Boolean) { Swing_Thread.require() options = options.bool.update(name, x) } } val bool = new Bool_Access class Int_Access { def apply(name: String): Int = options.int(name) def update(name: String, x: Int) { Swing_Thread.require() options = options.int.update(name, x) } } val int = new Int_Access class Real_Access { def apply(name: String): Double = options.real(name) def update(name: String, x: Double) { Swing_Thread.require() options = options.real.update(name, x) } } val real = new Real_Access class String_Access { def apply(name: String): String = options.string(name) def update(name: String, x: String) { Swing_Thread.require() options = options.string.update(name, x) } } val string = new String_Access class Seconds_Access { def apply(name: String): Time = options.seconds(name) } val seconds = new Seconds_Access } diff --git a/src/Pure/Tools/build.scala b/src/Pure/Tools/build.scala --- a/src/Pure/Tools/build.scala +++ b/src/Pure/Tools/build.scala @@ -1,791 +1,791 @@ /* Title: Pure/Tools/build.scala Author: Makarius Options: :folding=explicit:collapseFolds=1: Build and manage Isabelle sessions. */ package isabelle import java.util.{Timer, TimerTask} import java.io.{BufferedInputStream, FileInputStream, BufferedReader, InputStreamReader, IOException} import java.util.zip.GZIPInputStream import scala.collection.SortedSet import scala.annotation.tailrec object Build { /** progress context **/ class Progress { def echo(msg: String) {} def theory(session: String, theory: String) {} def stopped: Boolean = false } object Ignore_Progress extends Progress class Console_Progress(verbose: Boolean) extends Progress { override def echo(msg: String) { java.lang.System.out.println(msg) } override def theory(session: String, theory: String): Unit = if (verbose) echo(session + ": theory " + theory) } /** session information **/ // external version sealed case class Session_Entry( pos: Position.T, name: String, groups: List[String], path: String, parent: Option[String], description: String, options: List[Options.Spec], theories: List[(List[Options.Spec], List[String])], files: List[String]) // internal version sealed case class Session_Info( select: Boolean, pos: Position.T, groups: List[String], dir: Path, parent: Option[String], description: String, options: Options, theories: List[(Options, List[Path])], files: List[Path], entry_digest: SHA1.Digest) def is_pure(name: String): Boolean = name == "RAW" || name == "Pure" def session_info(options: Options, select: Boolean, dir: Path, entry: Session_Entry) : (String, Session_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(Path.explode(_))) }) val files = entry.files.map(Path.explode(_)) val entry_digest = SHA1.digest((name, entry.parent, entry.options, entry.theories).toString) val info = Session_Info(select, entry.pos, entry.groups, dir + Path.explode(entry.path), entry.parent, entry.description, session_options, theories, files, entry_digest) (name, info) } catch { case ERROR(msg) => error(msg + "\nThe error(s) above occurred in session entry " + quote(entry.name) + Position.here(entry.pos)) } /* session tree */ object Session_Tree { def apply(infos: Seq[(String, Session_Info)]): Session_Tree = { val graph1 = (Graph.string[Session_Info] /: infos) { case (graph, (name, info)) => if (graph.defined(name)) error("Duplicate session " + quote(name) + Position.here(info.pos)) else graph.new_node(name, info) } val graph2 = (graph1 /: graph1.entries) { case (graph, (name, (info, _))) => info.parent match { case None => graph case Some(parent) => if (!graph.defined(parent)) error("Bad parent session " + quote(parent) + " for " + quote(name) + Position.here(info.pos)) try { graph.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(info.pos)) } } } new Session_Tree(graph2) } } final class Session_Tree private(val graph: Graph[String, Session_Info]) extends PartialFunction[String, Session_Info] { def apply(name: String): Session_Info = graph.get_node(name) def isDefinedAt(name: String): Boolean = graph.defined(name) def selection(requirements: Boolean, all_sessions: Boolean, session_groups: List[String], sessions: List[String]): (List[String], Session_Tree) = { val bad_sessions = sessions.filterNot(isDefinedAt(_)) if (!bad_sessions.isEmpty) error("Undefined session(s): " + commas_quote(bad_sessions)) val pre_selected = { if (all_sessions) graph.keys.toList else { val select_group = session_groups.toSet val select = sessions.toSet (for { (name, (info, _)) <- graph.entries if info.select || select(name) || apply(name).groups.exists(select_group) } yield name).toList } } val selected = if (requirements) (graph.all_preds(pre_selected).toSet -- pre_selected).toList else pre_selected val tree1 = new Session_Tree(graph.restrict(graph.all_preds(selected).toSet)) (selected, tree1) } def topological_order: List[(String, Session_Info)] = graph.topological_order.map(name => (name, apply(name))) override def toString: String = graph.entries.map(_._1).toList.sorted.mkString(",") } /* parser */ private val SESSION = "session" private val IN = "in" private val DESCRIPTION = "description" private val OPTIONS = "options" private val THEORIES = "theories" private val FILES = "files" lazy val root_syntax = Outer_Syntax.init() + "(" + ")" + "+" + "," + "=" + "[" + "]" + (SESSION, Keyword.THY_DECL) + IN + DESCRIPTION + OPTIONS + THEORIES + FILES private object Parser extends Parse.Parser { def session_entry(pos: Position.T): Parser[Session_Entry] = { val session_name = atom("session name", _.is_name) val option = name ~ opt(keyword("=") ~! name ^^ { case _ ~ x => x }) ^^ { case x ~ y => (x, y) } val options = keyword("[") ~> rep1sep(option, keyword(",")) <~ keyword("]") val theories = keyword(THEORIES) ~! ((options | success(Nil)) ~ rep(theory_name)) ^^ { case _ ~ (x ~ y) => (x, y) } command(SESSION) ~! (session_name ~ ((keyword("(") ~! (rep1(name) <~ keyword(")")) ^^ { case _ ~ x => x }) | success(Nil)) ~ ((keyword(IN) ~! path ^^ { case _ ~ x => x }) | success(".")) ~ (keyword("=") ~! (opt(session_name ~! keyword("+") ^^ { case x ~ _ => x }) ~ ((keyword(DESCRIPTION) ~! text ^^ { case _ ~ x => x }) | success("")) ~ ((keyword(OPTIONS) ~! options ^^ { case _ ~ x => x }) | success(Nil)) ~ rep1(theories) ~ ((keyword(FILES) ~! rep1(path) ^^ { case _ ~ x => x }) | success(Nil))))) ^^ { case _ ~ (a ~ b ~ c ~ (_ ~ (d ~ e ~ f ~ g ~ h))) => Session_Entry(pos, a, b, c, d, e, f, g, h) } } def parse_entries(root: Path): List[Session_Entry] = { val toks = root_syntax.scan(File.read(root)) parse_all(rep(session_entry(root.position)), Token.reader(toks, root.implode)) match { case Success(result, _) => result case bad => error(bad.toString) } } } /* find sessions within certain directories */ private val ROOT = Path.explode("ROOT") private val ROOTS = Path.explode("ROOTS") 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 find_sessions(options: Options, more_dirs: List[(Boolean, Path)]): Session_Tree = { def find_dir(select: Boolean, dir: Path): List[(String, Session_Info)] = find_root(select, dir) ::: find_roots(select, dir) def find_root(select: Boolean, dir: Path): List[(String, Session_Info)] = { val root = dir + ROOT if (root.is_file) Parser.parse_entries(root).map(session_info(options, select, dir, _)) else Nil } def find_roots(select: Boolean, dir: Path): List[(String, Session_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 <- find_dir(select, dir1) } yield info } else Nil } val default_dirs = Isabelle_System.components().filter(is_session_dir(_)).map((false, _)) more_dirs foreach { case (_, dir) => check_session_dir(dir) } Session_Tree( for { (select, dir) <- default_dirs ::: more_dirs info <- find_dir(select, dir) } yield info) } /** build **/ /* queue */ object Queue { def apply(tree: Session_Tree): Queue = { val graph = tree.graph def outdegree(name: String): Int = graph.imm_succs(name).size def timeout(name: String): Double = tree(name).options.real("timeout") object Ordering extends scala.math.Ordering[String] { def compare(name1: String, name2: String): Int = outdegree(name2) compare outdegree(name1) match { case 0 => timeout(name2) compare timeout(name1) match { case 0 => name1 compare name2 case ord => ord } case ord => ord } } new Queue(graph, SortedSet(graph.keys.toList: _*)(Ordering)) } } final class Queue private(graph: Graph[String, Session_Info], order: SortedSet[String]) { def is_inner(name: String): Boolean = !graph.is_maximal(name) def is_empty: Boolean = graph.is_empty def - (name: String): Queue = new Queue(graph.del_node(name), order - name) def dequeue(skip: String => Boolean): Option[(String, Session_Info)] = { val it = order.iterator.dropWhile(name => skip(name) || !graph.is_minimal(name)) if (it.hasNext) { val name = it.next; Some((name, graph.get_node(name))) } else None } } /* source dependencies and static content */ sealed case class Session_Content( loaded_theories: Set[String], syntax: Outer_Syntax, sources: List[(Path, SHA1.Digest)], errors: List[String]) { def check_errors: Session_Content = { if (errors.isEmpty) this else error(cat_lines(errors)) } } sealed case class Deps(deps: Map[String, Session_Content]) { def is_empty: Boolean = deps.isEmpty def apply(name: String): Session_Content = deps(name) def sources(name: String): List[SHA1.Digest] = deps(name).sources.map(_._2) } def dependencies(progress: Build.Progress, inlined_files: Boolean, verbose: Boolean, list_files: Boolean, tree: Session_Tree): Deps = Deps((Map.empty[String, Session_Content] /: tree.topological_order)( { case (deps, (name, info)) => val (preloaded, parent_syntax, parent_errors) = info.parent match { case None => (Set.empty[String], Outer_Syntax.init_pure(), Nil) case Some(parent_name) => val parent = deps(parent_name) (parent.loaded_theories, parent.syntax, parent.errors) } val thy_info = new Thy_Info(new Thy_Load(preloaded, parent_syntax)) if (verbose || list_files) { val groups = if (info.groups.isEmpty) "" else info.groups.mkString(" (", " ", ")") progress.echo("Session " + name + groups) } val thy_deps = thy_info.dependencies(inlined_files, info.theories.map(_._2).flatten. map(thy => Thy_Load.path_node_name(info.dir + Thy_Load.thy_path(thy)))) val loaded_theories = thy_deps.loaded_theories val syntax = thy_deps.make_syntax val all_files = (thy_deps.deps.map({ case dep => val thy = Path.explode(dep.name.node) val uses = dep.join_header.uses.map(p => Path.explode(dep.name.dir) + Path.explode(p._1)) thy :: uses }).flatten ::: info.files.map(file => info.dir + file)).map(_.expand) if (list_files) progress.echo(cat_lines(all_files.map(_.implode).sorted.map(" " + _))) val sources = try { all_files.map(p => (p, SHA1.digest(p.file))) } catch { case ERROR(msg) => error(msg + "\nThe error(s) above occurred in session " + quote(name) + Position.here(info.pos)) } val errors = parent_errors ::: thy_deps.deps.map(dep => dep.join_header.errors).flatten deps + (name -> Session_Content(loaded_theories, syntax, sources, errors)) })) def session_content(inlined_files: Boolean, dirs: List[Path], session: String): Session_Content = { val options = Options.init() val (_, tree) = find_sessions(options, dirs.map((false, _))).selection(false, false, Nil, List(session)) dependencies(Build.Ignore_Progress, inlined_files, false, false, tree)(session) } def outer_syntax(session: String): Outer_Syntax = session_content(false, Nil, session).check_errors.syntax /* jobs */ private class Job(progress: Build.Progress, name: String, val info: Session_Info, output: Path, do_output: Boolean, verbose: Boolean, browser_info: Path) { // global browser info dir if (info.options.bool("browser_info") && !(browser_info + Path.explode("index.html")).is_file) { - browser_info.file.mkdirs() + Isabelle_System.mkdirs(browser_info) File.copy(Path.explode("~~/lib/logo/isabelle.gif"), browser_info + Path.explode("isabelle.gif")) File.write(browser_info + Path.explode("index.html"), File.read(Path.explode("~~/lib/html/library_index_header.template")) + File.read(Path.explode("~~/lib/html/library_index_content.template")) + File.read(Path.explode("~~/lib/html/library_index_footer.template"))) } def output_path: Option[Path] = if (do_output) Some(output) else None private val parent = info.parent.getOrElse("") private val args_file = File.tmp_file("args") File.write(args_file, YXML.string_of_body( if (is_pure(name)) Options.encode(info.options) else { import XML.Encode._ pair(bool, pair(Options.encode, pair(bool, pair(Path.encode, pair(string, pair(string, list(pair(Options.encode, list(Path.encode)))))))))( (do_output, (info.options, (verbose, (browser_info, (parent, (name, info.theories))))))) })) private val env = Map("INPUT" -> parent, "TARGET" -> name, "OUTPUT" -> Isabelle_System.standard_path(output), (if (is_pure(name)) "ISABELLE_PROCESS_OPTIONS" else "ARGS_FILE") -> Isabelle_System.posix_path(args_file)) private val script = if (is_pure(name)) { if (do_output) "./build " + name + " \"$OUTPUT\"" else """ rm -f "$OUTPUT"; ./build """ + name } else { """ . "$ISABELLE_HOME/lib/scripts/timestart.bash" """ + (if (do_output) """ "$ISABELLE_PROCESS" -e "Build.build \"$ARGS_FILE\";" -q -w "$INPUT" "$OUTPUT" """ else """ rm -f "$OUTPUT"; "$ISABELLE_PROCESS" -e "Build.build \"$ARGS_FILE\";" -r -q "$INPUT" """) + """ RC="$?" . "$ISABELLE_HOME/lib/scripts/timestop.bash" if [ "$RC" -eq 0 ]; then echo "Finished $TARGET ($TIMES_REPORT)" >&2 fi exit "$RC" """ } private val (thread, result) = Simple_Thread.future("build") { Isabelle_System.bash_env(info.dir.file, env, script, out_progress = (line: String) => Library.try_unprefix("\floading_theory = ", line) match { case Some(theory) => progress.theory(name, theory) case None => }) } def terminate: Unit = thread.interrupt def is_finished: Boolean = result.is_finished @volatile private var timeout = false private val time = info.options.seconds("timeout") private val timer: Option[Timer] = if (time.seconds > 0.0) { val t = new Timer("build", true) t.schedule(new TimerTask { def run = { terminate; timeout = true } }, time.ms) Some(t) } else None def join: Isabelle_System.Bash_Result = { val res = result.join args_file.delete timer.map(_.cancel()) if (res.rc == 130) res.add_err(if (timeout) "*** Timeout" else "*** Interrupt") else res } } /* log files */ private val LOG = Path.explode("log") private def log(name: String): Path = LOG + Path.basic(name) private def log_gz(name: String): Path = log(name).ext("gz") private val SESSION_PARENT_PATH = "\fSession.parent_path = " sealed case class Log_Info(stats: List[Properties.T], timing: Properties.T) def parse_log(text: String): Log_Info = { val lines = split_lines(text) val stats = Properties.parse_lines("\fML_statistics = ", lines) val timing = Properties.find_parse_line("\fTiming = ", lines) getOrElse Nil Log_Info(stats, timing) } /* sources and heaps */ private def sources_stamp(digests: List[SHA1.Digest]): String = digests.map(_.toString).sorted.mkString("sources: ", " ", "") private val no_heap: String = "heap: -" private def heap_stamp(heap: Option[Path]): String = { "heap: " + (heap match { case Some(path) => val file = path.file if (file.isFile) file.length.toString + " " + file.lastModified.toString else "-" case None => "-" }) } private def read_stamps(path: Path): Option[(String, String, String)] = if (path.is_file) { val stream = new GZIPInputStream (new BufferedInputStream(new FileInputStream(path.file))) val reader = new BufferedReader(new InputStreamReader(stream, UTF8.charset)) val (s, h1, h2) = try { (reader.readLine, reader.readLine, reader.readLine) } finally { reader.close } if (s != null && s.startsWith("sources: ") && h1 != null && h1.startsWith("heap: ") && h2 != null && h2.startsWith("heap: ")) Some((s, h1, h2)) else None } else None /* build */ def build( progress: Build.Progress, options: Options, requirements: Boolean = false, all_sessions: Boolean = false, build_heap: Boolean = false, clean_build: Boolean = false, more_dirs: List[(Boolean, Path)] = Nil, session_groups: List[String] = Nil, max_jobs: Int = 1, list_files: Boolean = false, no_build: Boolean = false, system_mode: Boolean = false, verbose: Boolean = false, sessions: List[String] = Nil): Int = { val full_tree = find_sessions(options, more_dirs) val (selected, selected_tree) = full_tree.selection(requirements, all_sessions, session_groups, sessions) val deps = dependencies(progress, true, verbose, list_files, selected_tree) val queue = Queue(selected_tree) def make_stamp(name: String): String = sources_stamp(selected_tree(name).entry_digest :: deps.sources(name)) val (input_dirs, output_dir, browser_info) = if (system_mode) { val output_dir = Path.explode("~~/heaps/$ML_IDENTIFIER") (List(output_dir), output_dir, Path.explode("~~/browser_info")) } else { val output_dir = Path.explode("$ISABELLE_OUTPUT") (output_dir :: Isabelle_System.find_logics_dirs(), output_dir, Path.explode("$ISABELLE_BROWSER_INFO")) } // prepare log dir - (output_dir + LOG).file.mkdirs() + Isabelle_System.mkdirs(output_dir + LOG) // optional cleanup if (clean_build) { for (name <- full_tree.graph.all_succs(selected)) { val files = List(Path.basic(name), log(name), log_gz(name)).map(output_dir + _).filter(_.is_file) if (!files.isEmpty) progress.echo("Cleaning " + name + " ...") if (!files.forall(p => p.file.delete)) progress.echo(name + " FAILED to delete") } } // scheduler loop case class Result(current: Boolean, parent_path: Option[String], heap: String, rc: Int) def sleep(): Unit = Thread.sleep(500) @tailrec def loop( pending: Queue, running: Map[String, (String, Job)], results: Map[String, Result]): Map[String, Result] = { if (pending.is_empty) results else if (progress.stopped) { for ((_, (_, job)) <- running) job.terminate sleep(); loop(pending, running, results) } else running.find({ case (_, (_, job)) => job.is_finished }) match { case Some((name, (parent_heap, job))) => //{{{ finish job val res = job.join progress.echo(res.err) val (parent_path, heap) = if (res.rc == 0) { (output_dir + log(name)).file.delete val sources = make_stamp(name) val heap = heap_stamp(job.output_path) File.write_gzip(output_dir + log_gz(name), sources + "\n" + parent_heap + "\n" + heap + "\n" + res.out) val parent_path = if (job.info.options.bool("browser_info")) res.out_lines.find(_.startsWith(SESSION_PARENT_PATH)).map(line => line.substring(SESSION_PARENT_PATH.length)) else None (parent_path, heap) } else { (output_dir + Path.basic(name)).file.delete (output_dir + log_gz(name)).file.delete File.write(output_dir + log(name), res.out) progress.echo(name + " FAILED") if (res.rc != 130) { progress.echo("(see also " + (output_dir + log(name)).file.toString + ")") val lines = res.out_lines.filterNot(_.startsWith("\f")) val tail = lines.drop(lines.length - 20 max 0) progress.echo("\n" + cat_lines(tail)) } (None, no_heap) } loop(pending - name, running - name, results + (name -> Result(false, parent_path, heap, res.rc))) //}}} case None if (running.size < (max_jobs max 1)) => //{{{ check/start next job pending.dequeue(running.isDefinedAt(_)) match { case Some((name, info)) => val parent_result = info.parent match { case None => Result(true, None, no_heap, 0) case Some(parent) => results(parent) } val output = output_dir + Path.basic(name) val do_output = build_heap || queue.is_inner(name) val (current, heap) = { input_dirs.find(dir => (dir + log_gz(name)).is_file) match { case Some(dir) => read_stamps(dir + log_gz(name)) match { case Some((s, h1, h2)) => val heap = heap_stamp(Some(dir + Path.basic(name))) (s == make_stamp(name) && h1 == parent_result.heap && h2 == heap && !(do_output && heap == no_heap), heap) case None => (false, no_heap) } case None => (false, no_heap) } } val all_current = current && parent_result.current if (all_current) loop(pending - name, running, results + (name -> Result(true, None, heap, 0))) else if (no_build) { if (verbose) progress.echo("Skipping " + name + " ...") loop(pending - name, running, results + (name -> Result(false, None, heap, 1))) } else if (parent_result.rc == 0) { progress.echo((if (do_output) "Building " else "Running ") + name + " ...") val job = new Job(progress, name, info, output, do_output, verbose, browser_info) loop(pending, running + (name -> (parent_result.heap, job)), results) } else { progress.echo(name + " CANCELLED") loop(pending - name, running, results + (name -> Result(false, None, heap, 1))) } case None => sleep(); loop(pending, running, results) } ///}}} case None => sleep(); loop(pending, running, results) } } val results = if (deps.is_empty) { progress.echo("### Nothing to build") Map.empty[String, Result] } else loop(queue, Map.empty, Map.empty) val session_entries = (for ((name, res) <- results.iterator if res.parent_path.isDefined) yield (res.parent_path.get, name)).toList.groupBy(_._1).map( { case (p, es) => (p, es.map(_._2).sorted) }) for ((p, names) <- session_entries) Present.update_index(browser_info + Path.explode(p), names) val rc = (0 /: results)({ case (rc1, (_, res)) => rc1 max res.rc }) if (rc != 0 && (verbose || !no_build)) { val unfinished = (for ((name, res) <- results.iterator if res.rc != 0) yield name).toList .sorted(scala.math.Ordering.String) // FIXME scala-2.10.0-RC1 progress.echo("Unfinished session(s): " + commas(unfinished)) } rc } /* command line entry point */ def main(args: Array[String]) { Command_Line.tool { args.toList match { case Properties.Value.Boolean(requirements) :: Properties.Value.Boolean(all_sessions) :: Properties.Value.Boolean(build_heap) :: Properties.Value.Boolean(clean_build) :: Properties.Value.Int(max_jobs) :: Properties.Value.Boolean(list_files) :: Properties.Value.Boolean(no_build) :: Properties.Value.Boolean(system_mode) :: Properties.Value.Boolean(verbose) :: Command_Line.Chunks(select_dirs, include_dirs, session_groups, build_options, sessions) => val options = (Options.init() /: build_options)(_ + _) val dirs = select_dirs.map(d => (true, Path.explode(d))) ::: include_dirs.map(d => (false, Path.explode(d))) build(new Build.Console_Progress(verbose), options, requirements, all_sessions, build_heap, clean_build, dirs, session_groups, max_jobs, list_files, no_build, system_mode, verbose, sessions) case _ => error("Bad arguments:\n" + cat_lines(args)) } } } }