diff --git a/src/Pure/Admin/build_history.scala b/src/Pure/Admin/build_history.scala --- a/src/Pure/Admin/build_history.scala +++ b/src/Pure/Admin/build_history.scala @@ -1,448 +1,456 @@ /* Title: Pure/Admin/build_history.scala Author: Makarius Build other history versions. */ package isabelle import java.io.{File => JFile} import java.time.format.DateTimeFormatter import java.util.Locale object Build_History { /* log files */ val engine = "build_history" val log_prefix = engine + "_" val META_INFO_MARKER = "\fmeta_info = " /* augment settings */ def augment_settings( other_isabelle: Other_Isabelle, threads: Int, - arch_64: Boolean = false, - heap: Int = default_heap, - max_heap: Option[Int] = None, + arch_64: Boolean, + heap: Int, + max_heap: Option[Int], + init_settings: List[String], more_settings: List[String]): String = { val (ml_platform, ml_settings) = { val windows_32 = "x86-windows" val windows_64 = "x86_64-windows" val platform_32 = other_isabelle("getenv -b ISABELLE_PLATFORM32").check.out val platform_64 = other_isabelle("getenv -b ISABELLE_PLATFORM64").check.out val platform_family = other_isabelle("getenv -b ISABELLE_PLATFORM_FAMILY").check.out val polyml_home = try { Path.explode(other_isabelle("getenv -b ML_HOME").check.out).dir } catch { case ERROR(msg) => error("Bad ML_HOME: " + msg) } def ml_home(platform: String): Path = polyml_home + Path.explode(platform) def err(platform: String): Nothing = error("Platform " + platform + " unavailable on this machine") def check_dir(platform: String): Boolean = platform != "" && ml_home(platform).is_dir val ml_platform = if (Platform.is_windows && arch_64) { if (check_dir(windows_64)) windows_64 else err(windows_64) } else if (Platform.is_windows && !arch_64) { if (check_dir(windows_32)) windows_32 else platform_32 // x86-cygwin } else { val (platform, platform_name) = if (arch_64) (platform_64, "x86_64-" + platform_family) else (platform_32, "x86-" + platform_family) if (check_dir(platform)) platform else err(platform_name) } val ml_options = "--minheap " + heap + (if (max_heap.isDefined) " --maxheap " + max_heap.get else "") + " --gcthreads " + threads + (if (ml_platform.endsWith("-windows")) " --codepage utf8" else "") (ml_platform, List( "ML_HOME=" + File.bash_path(ml_home(ml_platform)), "ML_PLATFORM=" + quote(ml_platform), "ML_OPTIONS=" + quote(ml_options))) } val thread_settings = List( "ISABELLE_JAVA_SYSTEM_OPTIONS=\"$ISABELLE_JAVA_SYSTEM_OPTIONS -Disabelle.threads=" + threads + "\"", "ISABELLE_BUILD_OPTIONS=\"threads=" + threads + "\"") val settings = + (if (init_settings.isEmpty) Nil else List(init_settings)) ::: List(ml_settings, thread_settings) ::: - (if (more_settings.isEmpty) Nil else List(more_settings)) + (if (more_settings.isEmpty) Nil else List(more_settings)) File.append(other_isabelle.etc_settings, "\n" + cat_lines(settings.map(terminate_lines(_)))) ml_platform } /** build_history **/ private val default_rev = "tip" private val default_multicore = (1, 1) private val default_heap = 1500 private val default_isabelle_identifier = "build_history" def build_history( hg: Mercurial.Repository, progress: Progress = No_Progress, rev: String = default_rev, isabelle_identifier: String = default_isabelle_identifier, components_base: String = "", fresh: Boolean = false, nonfree: Boolean = false, multicore_base: Boolean = false, multicore_list: List[(Int, Int)] = List(default_multicore), arch_64: Boolean = false, heap: Int = default_heap, max_heap: Option[Int] = None, + init_settings: List[String] = Nil, more_settings: List[String] = Nil, verbose: Boolean = false, build_tags: List[String] = Nil, build_args: List[String] = Nil): List[(Process_Result, Path)] = { /* sanity checks */ if (File.eq(Path.explode("~~"), hg.root)) error("Repository coincides with ISABELLE_HOME=" + Path.explode("~~").expand) for ((threads, _) <- multicore_list if threads < 1) error("Bad threads value < 1: " + threads) for ((_, processes) <- multicore_list if processes < 1) error("Bad processes value < 1: " + processes) if (heap < 100) error("Bad heap value < 100: " + heap) if (max_heap.isDefined && max_heap.get < heap) error("Bad max_heap value < heap: " + max_heap.get) System.getenv("ISABELLE_SETTINGS_PRESENT") match { case null | "" => case _ => error("Cannot run build_history within existing Isabelle settings environment") } /* init repository */ hg.update(rev = rev, clean = true) progress.echo_if(verbose, hg.log(rev, options = "-l1")) val isabelle_version = hg.id(rev) val other_isabelle = new Other_Isabelle(progress, hg.root, isabelle_identifier) /* main */ val build_host = Isabelle_System.hostname() val build_history_date = Date.now() val build_group_id = build_host + ":" + build_history_date.time.ms var first_build = true for ((threads, processes) <- multicore_list) yield { /* init settings */ other_isabelle.init_settings(components_base, nonfree) other_isabelle.resolve_components(verbose) val ml_platform = - augment_settings(other_isabelle, threads, arch_64, heap, max_heap, more_settings) + augment_settings(other_isabelle, threads, arch_64, heap, max_heap, + init_settings, more_settings) val isabelle_output = Path.explode(other_isabelle("getenv -b ISABELLE_OUTPUT").check.out) val isabelle_output_log = isabelle_output + Path.explode("log") val isabelle_base_log = isabelle_output + Path.explode("../base_log") if (first_build) { other_isabelle.resolve_components(verbose) if (fresh) Isabelle_System.rm_tree(other_isabelle.isabelle_home + Path.explode("lib/classes")) other_isabelle.bash( "env PATH=\"" + File.bash_path(Path.explode("~~/lib/dummy_stty").expand) + ":$PATH\" " + "bin/isabelle jedit -b", redirect = true, echo = verbose).check Isabelle_System.rm_tree(isabelle_base_log) } Isabelle_System.rm_tree(isabelle_output) Isabelle_System.mkdirs(isabelle_output) /* build */ if (multicore_base && !first_build && isabelle_base_log.is_dir) Isabelle_System.copy_dir(isabelle_base_log, isabelle_output_log) val build_start = Date.now() val build_args1 = List("-v", "-j" + processes) ::: build_args val build_result = other_isabelle("build " + Bash.strings(build_args1), redirect = true, echo = verbose) val build_end = Date.now() val log_path = other_isabelle.isabelle_home_user + Build_Log.log_subdir(build_history_date) + Build_Log.log_filename(Build_History.engine, build_history_date, List(build_host, ml_platform, "M" + threads) ::: build_tags) val build_info: Build_Log.Build_Info = Build_Log.Log_File(log_path.base.implode, build_result.out_lines). parse_build_info(ml_statistics = true) /* output log */ val store = Sessions.store() val meta_info = Properties.lines_nonempty(Build_Log.Prop.build_tags.name, build_tags) ::: Properties.lines_nonempty(Build_Log.Prop.build_args.name, build_args1) ::: List( Build_Log.Prop.build_group_id.name -> build_group_id, Build_Log.Prop.build_id.name -> (build_host + ":" + build_start.time.ms), Build_Log.Prop.build_engine.name -> Build_History.engine, Build_Log.Prop.build_host.name -> build_host, Build_Log.Prop.build_start.name -> Build_Log.print_date(build_start), Build_Log.Prop.build_end.name -> Build_Log.print_date(build_end), Build_Log.Prop.isabelle_version.name -> isabelle_version) val ml_statistics = build_info.finished_sessions.flatMap(session_name => { val database = isabelle_output + store.database(session_name) val log_gz = isabelle_output + store.log_gz(session_name) val properties = if (database.is_file) { using(SQLite.open_database(database))(db => store.read_build_log(db, session_name, ml_statistics = true)).ml_statistics } else if (log_gz.is_file) { Build_Log.Log_File(log_gz).parse_session_info(ml_statistics = true).ml_statistics } else Nil properties.map(props => (Build_Log.SESSION_NAME -> session_name) :: props) }) val heap_sizes = build_info.finished_sessions.flatMap(session_name => { val heap = isabelle_output + Path.explode(session_name) if (heap.is_file) Some("Heap " + session_name + " (" + Value.Long(heap.file.length) + " bytes)") else None }) Isabelle_System.mkdirs(log_path.dir) File.write_xz(log_path.ext("xz"), terminate_lines( Build_Log.Log_File.print_props(META_INFO_MARKER, meta_info) :: build_result.out_lines ::: ml_statistics.map(Build_Log.Log_File.print_props(Build_Log.ML_STATISTICS_MARKER, _)) ::: heap_sizes), XZ.options(6)) /* next build */ if (multicore_base && first_build && isabelle_output_log.is_dir) Isabelle_System.copy_dir(isabelle_output_log, isabelle_base_log) Isabelle_System.rm_tree(isabelle_output) first_build = false (build_result, log_path.ext("xz")) } } /* command line entry point */ private object Multicore { private val Pat1 = """^(\d+)$""".r private val Pat2 = """^(\d+)x(\d+)$""".r def parse(s: String): (Int, Int) = s match { case Pat1(Value.Int(x)) => (x, 1) case Pat2(Value.Int(x), Value.Int(y)) => (x, y) case _ => error("Bad multicore configuration: " + quote(s)) } } def main(args: Array[String]) { Command_Line.tool0 { var multicore_base = false var components_base = "" var heap: Option[Int] = None var max_heap: Option[Int] = None var multicore_list = List(default_multicore) var isabelle_identifier = default_isabelle_identifier var more_settings: List[String] = Nil var fresh = false + var init_settings: List[String] = Nil var arch_64 = false var nonfree = false var rev = default_rev var build_tags = List.empty[String] var verbose = false val getopts = Getopts(""" Usage: isabelle build_history [OPTIONS] REPOSITORY [ARGS ...] Options are: -B first multicore build serves as base for scheduling information -C DIR base directory for Isabelle components (default: $ISABELLE_HOME_USER/../contrib) - -H SIZE minimal ML heap in MB (default: """ + default_heap + """ for x86, """ + default_heap * 2 + """ for x86_64) + -H SIZE minimal ML heap in MB (default: """ + default_heap + """ for x86, """ + + default_heap * 2 + """ for x86_64) -M MULTICORE multicore configurations (see below) -N NAME alternative ISABELLE_IDENTIFIER (default: """ + default_isabelle_identifier + """) -U SIZE maximal ML heap in MB (default: unbounded) -e TEXT additional text for generated etc/settings -f fresh build of Isabelle/Scala components (recommended) + -i TEXT initial text for generated etc/settings -m ARCH processor architecture (32=x86, 64=x86_64, default: x86) -n include nonfree components -r REV update to revision (default: """ + default_rev + """) -t TAG free-form build tag (multiple occurrences possible) -v verbose Build Isabelle sessions from the history of another REPOSITORY clone, passing ARGS directly to its isabelle build tool. Each MULTICORE configuration consists of one or two numbers (default 1): THREADS or THREADSxPROCESSES, e.g. -M 1,2,4 or -M 1x4,2x2,4. """, "B" -> (_ => multicore_base = true), "C:" -> (arg => components_base = arg), "H:" -> (arg => heap = Some(Value.Int.parse(arg))), "M:" -> (arg => multicore_list = space_explode(',', arg).map(Multicore.parse(_))), "N:" -> (arg => isabelle_identifier = arg), "U:" -> (arg => max_heap = Some(Value.Int.parse(arg))), "e:" -> (arg => more_settings = more_settings ::: List(arg)), "f" -> (_ => fresh = true), + "i:" -> (arg => init_settings = init_settings ::: List(arg)), "m:" -> { case "32" | "x86" => arch_64 = false case "64" | "x86_64" => arch_64 = true case bad => error("Bad processor architecture: " + quote(bad)) }, "n" -> (_ => nonfree = true), "r:" -> (arg => rev = arg), "t:" -> (arg => build_tags = build_tags ::: List(arg)), "v" -> (_ => verbose = true)) val more_args = getopts(args) val (root, build_args) = more_args match { case root :: build_args => (root, build_args) case _ => getopts.usage() } val hg = Mercurial.repository(Path.explode(root)) val progress = new Console_Progress(stderr = true) val results = build_history(hg, progress = progress, rev = rev, isabelle_identifier = isabelle_identifier, components_base = components_base, fresh = fresh, nonfree = nonfree, multicore_base = multicore_base, multicore_list = multicore_list, arch_64 = arch_64, heap = heap.getOrElse(if (arch_64) default_heap * 2 else default_heap), - max_heap = max_heap, more_settings = more_settings, verbose = verbose, - build_tags = build_tags, build_args = build_args) + max_heap = max_heap, init_settings = init_settings, more_settings = more_settings, + verbose = verbose, build_tags = build_tags, build_args = build_args) for ((_, log_path) <- results) Output.writeln(log_path.implode, stdout = true) val rc = (0 /: results) { case (rc, (res, _)) => rc max res.rc } if (rc != 0) sys.exit(rc) } } /** remote build_history -- via command-line **/ def remote_build_history( ssh: SSH.Session, isabelle_repos_self: Path, isabelle_repos_other: Path, isabelle_repos_source: String = "http://isabelle.in.tum.de/repos/isabelle", self_update: Boolean = false, push_isabelle_home: Boolean = false, progress: Progress = No_Progress, options: String = "", args: String = ""): (List[(String, Bytes)], Process_Result) = { val isabelle_admin = isabelle_repos_self + Path.explode("Admin") /* prepare repository clones */ val isabelle_hg = Mercurial.setup_repository(isabelle_repos_source, isabelle_repos_self, ssh = Some(ssh)) val rev = if (self_update) { val rev = if (push_isabelle_home) { val isabelle_home_hg = Mercurial.repository(Path.explode("~~")) val rev = isabelle_home_hg.id() isabelle_home_hg.push(isabelle_hg.root_url, rev = rev, force = true) rev } else { isabelle_hg.pull() isabelle_hg.id() } isabelle_hg.update(rev = rev, clean = true) ssh.execute( ssh.bash_path(isabelle_repos_self + Path.explode("bin/isabelle")) + " components -a").check ssh.execute(ssh.bash_path(isabelle_admin + Path.explode("build")) + " jars_fresh").check rev } else "tip" val other_hg = Mercurial.setup_repository( ssh.bash_path(isabelle_repos_self), isabelle_repos_other, ssh = Some(ssh)) other_hg.pull(isabelle_hg.root.implode) other_hg.update(rev = rev, clean = true) /* Admin/build_history */ val process_result = ssh.execute( ssh.bash_path(isabelle_admin + Path.explode("build_history")) + " " + options + " " + ssh.bash_path(isabelle_repos_other) + " " + args, progress_stdout = progress.echo(_), progress_stderr = progress.echo(_), strict = false) val result = for (line <- process_result.out_lines) yield { val log = Path.explode(line) val bytes = ssh.read_bytes(log) ssh.rm(log) (log.base.implode, bytes) } (result, process_result) } } diff --git a/src/Pure/Admin/isabelle_cronjob.scala b/src/Pure/Admin/isabelle_cronjob.scala --- a/src/Pure/Admin/isabelle_cronjob.scala +++ b/src/Pure/Admin/isabelle_cronjob.scala @@ -1,435 +1,436 @@ /* Title: Pure/Admin/isabelle_cronjob.scala Author: Makarius Main entry point for administrative cronjob at TUM. */ package isabelle import scala.annotation.tailrec import scala.collection.mutable object Isabelle_Cronjob { /* file-system state: owned by main cronjob */ val main_dir = Path.explode("~/cronjob") val main_state_file = main_dir + Path.explode("run/main.state") val current_log = main_dir + Path.explode("run/main.log") // owned by log service val cumulative_log = main_dir + Path.explode("log/main.log") // owned by log service val isabelle_repos = main_dir + Path.explode("isabelle") val isabelle_repos_test = main_dir + Path.explode("isabelle-test") val afp_repos = main_dir + Path.explode("AFP") val isabelle_dev_source = "http://isabelle.in.tum.de/repos/isabelle" val isabelle_release_source = "http://bitbucket.org/isabelle_project/isabelle-release" val afp_source = "https://bitbucket.org/isa-afp/afp-devel" val jenkins_jobs = "identify" :: Jenkins.build_log_jobs /** particular tasks **/ /* identify Isabelle + AFP repository snapshots and build release */ private val build_release = Logger_Task("build_release", logger => { Isabelle_Devel.make_index() val rev = Mercurial.repository(isabelle_repos).id() val afp_rev = Mercurial.setup_repository(afp_source, afp_repos).id() File.write(logger.log_dir + Build_Log.log_filename("isabelle_identify", logger.start_date), Build_Log.Identify.content(logger.start_date, Some(rev), Some(afp_rev))) Isabelle_Devel.release_snapshot(rev = rev, afp_rev = afp_rev, parallel_jobs = 4, remote_mac = "macbroy31") }) /* integrity test of build_history vs. build_history_base */ private val build_history_base = Logger_Task("build_history_base", logger => { val hg = Mercurial.setup_repository( File.standard_path(isabelle_repos), isabelle_repos_test) for { (result, log_path) <- Build_History.build_history( hg, rev = "build_history_base", fresh = true, build_args = List("HOL")) } { result.check File.move(log_path, logger.log_dir + log_path.base) } }) /* remote build_history */ sealed case class Item(known: Boolean, isabelle_version: String, pull_date: Date) { def unknown: Boolean = !known } def recent_items(db: SQL.Database, days: Int, rev: String, sql: SQL.Source): List[Item] = { val select = Build_Log.Data.select_recent_isabelle_versions(days = days, rev = rev, sql = "WHERE " + sql) db.using_statement(select)(stmt => stmt.execute_query().iterator(res => { val known = res.bool(Build_Log.Data.known) val isabelle_version = res.string(Build_Log.Prop.isabelle_version) val pull_date = res.date(Build_Log.Data.pull_date) Item(known, isabelle_version, pull_date) }).toList) } def unknown_runs(items: List[Item]): List[List[Item]] = { val (run, rest) = Library.take_prefix[Item](_.unknown, items.dropWhile(_.known)) if (run.nonEmpty) run :: unknown_runs(rest) else Nil } sealed case class Remote_Build( description: String, host: String, user: String = "", port: Int = 0, shared_home: Boolean = true, historic: Boolean = false, history: Int = 0, history_base: String = "build_history_base", options: String = "", args: String = "", detect: SQL.Source = "") { def sql: SQL.Source = Build_Log.Prop.build_engine + " = " + SQL.string(Build_History.engine) + " AND " + Build_Log.Prop.build_host + " = " + SQL.string(host) + (if (detect == "") "" else " AND " + SQL.enclose(detect)) def profile: Build_Status.Profile = Build_Status.Profile(description, history, sql) def history_base_filter(hg: Mercurial.Repository): Set[String] = { val rev0 = hg.id(history_base) val graph = hg.graph() (rev0 :: graph.all_succs(List(rev0))).toSet } def pick(options: Options, rev: String = "", filter: String => Boolean = (_: String) => true) : Option[String] = { val store = Build_Log.store(options) using(store.open_database())(db => { def pick_days(days: Int): Option[String] = { val items = recent_items(db, days = days, rev = rev, sql = sql). filter(item => filter(item.isabelle_version)) def runs = unknown_runs(items) val known_rev = rev != "" && items.exists(item => item.known && item.isabelle_version == rev) if (historic || known_rev) { val longest_run = (List.empty[Item] /: runs)({ case (item1, item2) => if (item1.length >= item2.length) item1 else item2 }) if (longest_run.isEmpty) None else Some(longest_run(longest_run.length / 2).isabelle_version) } else if (rev != "") Some(rev) else runs.flatten.headOption.map(_.isabelle_version) } pick_days(options.int("build_log_history") max history) orElse pick_days(200) orElse pick_days(2000) }) } } val remote_builds_old: List[Remote_Build] = List( Remote_Build("Poly/ML test", "lxbroy8", - options = "-m32 -B -M1x2,2 -t polyml-test -e 'init_component /home/isabelle/contrib/polyml-5.7-20170217'", + options = "-m32 -B -M1x2,2 -t polyml-test -i 'init_component /home/isabelle/contrib/polyml-5.7-20170217'", args = "-N -g timing", detect = Build_Log.Prop.build_tags + " = " + SQL.string("polyml-test")), Remote_Build("Mac OS X 10.8 Mountain Lion", "macbroy30", options = "-m32 -M2", args = "-a", detect = Build_Log.Prop.build_start + " < date '2017-03-03'")) val remote_builds: List[List[Remote_Build]] = { List( List(Remote_Build("Poly/ML 5.7 Linux", "lxbroy8", history_base = "37074e22e8be", - options = "-m32 -B -M1x2,2 -t polyml-5.7 -e 'init_component /home/isabelle/contrib/polyml-5.7'", + options = "-m32 -B -M1x2,2 -t polyml-5.7 -i 'init_component /home/isabelle/contrib/polyml-5.7'", args = "-N -g timing", - detect = Build_Log.Prop.build_tags + " = " + SQL.string("polyml-5.7"))), + detect = Build_Log.Prop.build_tags + " = " + SQL.string("polyml-5.7") + " AND " + + Build_Log.Settings.ML_OPTIONS + " <> " + SQL.string("-H 500"))), List(Remote_Build("Linux A", "lxbroy9", options = "-m32 -B -M1x2,2", args = "-N -g timing")), List(Remote_Build("Linux B", "lxbroy10", historic = true, history = 90, options = "-m32 -B -M1x4,2,4,6", args = "-N -g timing")), List( Remote_Build("Mac OS X 10.9 Mavericks", "macbroy2", options = "-m32 -M8", args = "-a", detect = Build_Log.Prop.build_tags + " IS NULL"), Remote_Build("Mac OS X 10.9 Mavericks, quick_and_dirty", "macbroy2", options = "-m32 -M8 -t quick_and_dirty", args = "-a -o quick_and_dirty", detect = Build_Log.Prop.build_tags + " = " + SQL.string("quick_and_dirty")), Remote_Build("Mac OS X 10.9 Mavericks, skip_proofs", "macbroy2", options = "-m32 -M8 -t skip_proofs", args = "-a -o skip_proofs", detect = Build_Log.Prop.build_tags + " = " + SQL.string("skip_proofs"))), List( Remote_Build("Mac OS X 10.12 Sierra", "macbroy30", options = "-m32 -M2", args = "-a", detect = Build_Log.Prop.build_start + " > date '2017-03-03'")), List(Remote_Build("Mac OS X 10.10 Yosemite", "macbroy31", options = "-m32 -M2", args = "-a")), List( Remote_Build("Windows", "vmnipkow9", history = 90, shared_home = false, options = "-m32 -M4", args = "-a", detect = Build_Log.Settings.ML_PLATFORM + " = " + SQL.string("x86-windows")), Remote_Build("Windows", "vmnipkow9", history = 90, shared_home = false, options = "-m64 -M4", args = "-a", detect = Build_Log.Settings.ML_PLATFORM + " = " + SQL.string("x86_64-windows")))) } private def remote_build_history(rev: String, r: Remote_Build): Logger_Task = { val task_name = "build_history-" + r.host Logger_Task(task_name, logger => { using(logger.ssh_context.open_session(host = r.host, user = r.user, port = r.port))( ssh => { val self_update = !r.shared_home val push_isabelle_home = self_update && Mercurial.is_repository(Path.explode("~~")) val (results, _) = Build_History.remote_build_history(ssh, isabelle_repos, isabelle_repos.ext(r.host), isabelle_repos_source = isabelle_dev_source, self_update = self_update, push_isabelle_home = push_isabelle_home, options = "-r " + Bash.string(rev) + " -N " + Bash.string(task_name) + " -f " + r.options, args = "-o timeout=10800 " + r.args) for ((log_name, bytes) <- results) { logger.log(Date.now(), log_name) Bytes.write(logger.log_dir + Path.explode(log_name), bytes) } }) }) } val build_status_profiles: List[Build_Status.Profile] = (remote_builds_old :: remote_builds).flatten.map(_.profile) /** task logging **/ sealed case class Logger_Task(name: String = "", body: Logger => Unit) class Log_Service private[Isabelle_Cronjob](progress: Progress, val ssh_context: SSH.Context) { current_log.file.delete private val thread: Consumer_Thread[String] = Consumer_Thread.fork("cronjob: logger", daemon = true)( consume = (text: String) => { // critical File.append(current_log, text + "\n") File.append(cumulative_log, text + "\n") progress.echo(text) true }) def shutdown() { thread.shutdown() } val hostname = Isabelle_System.hostname() def log(date: Date, task_name: String, msg: String): Unit = if (task_name != "") thread.send( "[" + Build_Log.print_date(date) + ", " + hostname + ", " + task_name + "]: " + msg) def start_logger(start_date: Date, task_name: String): Logger = new Logger(this, start_date, task_name) def run_task(start_date: Date, task: Logger_Task) { val logger = start_logger(start_date, task.name) val res = Exn.capture { task.body(logger) } val end_date = Date.now() val err = res match { case Exn.Res(_) => None case Exn.Exn(exn) => System.err.println("Exception trace for " + quote(task.name) + ":") exn.printStackTrace() val first_line = Library.split_lines(Exn.message(exn)).headOption getOrElse "exception" Some(first_line) } logger.log_end(end_date, err) } def fork_task(start_date: Date, task: Logger_Task): Task = new Task(task.name, run_task(start_date, task)) } class Logger private[Isabelle_Cronjob]( val log_service: Log_Service, val start_date: Date, val task_name: String) { def ssh_context: SSH.Context = log_service.ssh_context def options: Options = ssh_context.options def log(date: Date, msg: String): Unit = log_service.log(date, task_name, msg) def log_end(end_date: Date, err: Option[String]) { val elapsed_time = end_date.time - start_date.time val msg = (if (err.isEmpty) "finished" else "ERROR " + err.get) + (if (elapsed_time.seconds < 3.0) "" else " (" + elapsed_time.message_hms + " elapsed time)") log(end_date, msg) } val log_dir: Path = main_dir + Build_Log.log_subdir(start_date) Isabelle_System.mkdirs(log_dir) log(start_date, "started") } class Task private[Isabelle_Cronjob](name: String, body: => Unit) { private val future: Future[Unit] = Future.thread("cronjob: " + name) { body } def is_finished: Boolean = future.is_finished } /** cronjob **/ def cronjob(progress: Progress, exclude_task: Set[String]) { /* soft lock */ val still_running = try { Some(File.read(main_state_file)) } catch { case ERROR(_) => None } still_running match { case None | Some("") => case Some(running) => error("Isabelle cronjob appears to be still running: " + running) } /* log service */ val log_service = new Log_Service(progress, SSH.init_context(Options.init())) def run(start_date: Date, task: Logger_Task) { log_service.run_task(start_date, task) } def run_now(task: Logger_Task) { run(Date.now(), task) } /* structured tasks */ def SEQ(tasks: List[Logger_Task]): Logger_Task = Logger_Task(body = _ => for (task <- tasks.iterator if !exclude_task(task.name) || task.name == "") run_now(task)) def PAR(tasks: List[Logger_Task]): Logger_Task = Logger_Task(body = _ => { @tailrec def join(running: List[Task]) { running.partition(_.is_finished) match { case (Nil, Nil) => case (Nil, _ :: _) => Thread.sleep(500); join(running) case (_ :: _, remaining) => join(remaining) } } val start_date = Date.now() val running = for (task <- tasks if !exclude_task(task.name)) yield log_service.fork_task(start_date, task) join(running) }) /* main */ val main_start_date = Date.now() File.write(main_state_file, main_start_date + " " + log_service.hostname) val hg = Mercurial.repository(isabelle_repos) val rev = hg.id() run(main_start_date, Logger_Task("isabelle_cronjob", logger => run_now( SEQ(List(build_release, build_history_base, PAR(remote_builds.map(seq => SEQ(seq.flatMap(r => r.pick(logger.options, rev, r.history_base_filter(hg)). map(remote_build_history(_, r)))))), Logger_Task("jenkins_logs", _ => Jenkins.download_logs(jenkins_jobs, main_dir)), Logger_Task("build_log_database", logger => Isabelle_Devel.build_log_database(logger.options)), Logger_Task("build_status", logger => Isabelle_Devel.build_status(logger.options))))))) log_service.shutdown() main_state_file.file.delete } /** command line entry point **/ def main(args: Array[String]) { Command_Line.tool0 { var force = false var verbose = false var exclude_task = Set.empty[String] val getopts = Getopts(""" Usage: Admin/cronjob/main [OPTIONS] Options are: -f apply force to do anything -v verbose -x NAME exclude tasks with this name """, "f" -> (_ => force = true), "v" -> (_ => verbose = true), "x:" -> (arg => exclude_task += arg)) val more_args = getopts(args) if (more_args.nonEmpty) getopts.usage() val progress = if (verbose) new Console_Progress() else No_Progress if (force) cronjob(progress, exclude_task) else error("Need to apply force to do anything") } } }