# HG changeset patch # User wenzelm # Date 1656423245 -7200 # Node ID e4bbe0b9288db897b619e7742e2f96290a4194f9 # Parent 0dd3ac5fdbaabeb1ac4300177bec452ce654f858# Parent 809c37bfd8234198511ef34f55b9c92b7f743859 merged diff -r 0dd3ac5fdbaa -r e4bbe0b9288d Admin/jenkins/build/ci_build_benchmark.scala --- a/Admin/jenkins/build/ci_build_benchmark.scala Mon Jun 27 17:36:26 2022 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,16 +0,0 @@ -object profile extends isabelle.CI_Profile { - - import isabelle._ - - override def documents = false - override def threads = 6 - override def jobs = 1 - def include = Nil - def select = List(Path.explode("$ISABELLE_HOME/src/Benchmarks")) - - def pre_hook(args: List[String]) = Result.ok - def post_hook(results: Build.Results) = Result.ok - - def selection = Sessions.Selection(session_groups = List("timing")) - -} diff -r 0dd3ac5fdbaa -r e4bbe0b9288d etc/build.props --- a/etc/build.props Mon Jun 27 17:36:26 2022 +0200 +++ b/etc/build.props Tue Jun 28 15:34:05 2022 +0200 @@ -33,6 +33,7 @@ src/Pure/Admin/build_verit.scala \ src/Pure/Admin/build_zipperposition.scala \ src/Pure/Admin/check_sources.scala \ + src/Pure/Admin/ci_build_benchmark.scala \ src/Pure/Admin/ci_profile.scala \ src/Pure/Admin/isabelle_cronjob.scala \ src/Pure/Admin/isabelle_devel.scala \ diff -r 0dd3ac5fdbaa -r e4bbe0b9288d src/Pure/Admin/ci_build_benchmark.scala --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/Admin/ci_build_benchmark.scala Tue Jun 28 15:34:05 2022 +0200 @@ -0,0 +1,28 @@ +/* Title: Pure/Admin/ci_build_benchmark.scala + Author: Lars Hupel and Fabian Huch, TU Munich + +CI benchmark build profile. +*/ + +package isabelle + + +object CI_Build_Benchmark { + val isabelle_tool = + Isabelle_Tool("ci_build_benchmark", "builds Isabelle benchmarks + timing group", + Scala_Project.here, { args => + val getopts = Getopts(""" +Usage: isabelle ci_build_benchmark + + Builds Isabelle benchmark and timing sessions. + """) + getopts(args) + + val selection = Sessions.Selection(session_groups = List("timing")) + val profile = CI_Profile.Profile(threads = 6, jobs = 1, numa = false) + val config = CI_Profile.Build_Config(documents = false, + select = List(Path.explode("$ISABELLE_HOME/src/Benchmarks")), selection = selection) + + CI_Profile.build(profile, config) + }) +} diff -r 0dd3ac5fdbaa -r e4bbe0b9288d src/Pure/Admin/ci_profile.scala --- a/src/Pure/Admin/ci_profile.scala Mon Jun 27 17:36:26 2022 +0200 +++ b/src/Pure/Admin/ci_profile.scala Tue Jun 28 15:34:05 2022 +0200 @@ -1,5 +1,5 @@ /* Title: Pure/Admin/ci_profile.scala - Author: Lars Hupel + Author: Lars Hupel and Fabian Huch, TU Munich Build profile for continuous integration services. */ @@ -10,99 +10,152 @@ import java.time.{Instant, ZoneId} import java.time.format.DateTimeFormatter import java.util.{Properties => JProperties, Map => JMap} +import java.nio.file.Files -abstract class CI_Profile extends Isabelle_Tool.Body { +object CI_Profile { + + /* Result */ + case class Result(rc: Int) case object Result { def ok: Result = Result(Process_Result.RC.ok) def error: Result = Result(Process_Result.RC.error) } - private def build(options: Options): (Build.Results, Time) = { - val progress = new Console_Progress(verbose = true) - val start_time = Time.now() - val results = progress.interrupt_handler { - Build.build( - options + "system_heaps", - selection = selection, - progress = progress, - clean_build = clean, - verbose = true, - numa_shuffling = numa, - max_jobs = jobs, - dirs = include, - select_dirs = select) + + /* Profile */ + + case class Profile(threads: Int, jobs: Int, numa: Boolean) + + object Profile { + def from_host: Profile = { + Isabelle_System.hostname() match { + case "hpcisabelle" => Profile(8, 8, numa = true) + case "lxcisa1" => Profile(4, 10, numa = false) + case _ => Profile(2, 2, numa = false) + } } - val end_time = Time.now() - (results, end_time - start_time) } - private def load_properties(): JProperties = { - val props = new JProperties() - val file_name = Isabelle_System.getenv("ISABELLE_CI_PROPERTIES") + + /* Build_Config */ + + case class Build_Config( + documents: Boolean = true, + clean: Boolean = true, + include: List[Path] = Nil, + select: List[Path] = Nil, + pre_hook: () => Result = () => Result.ok, + post_hook: Build.Results => Result = _ => Result.ok, + selection: Sessions.Selection = Sessions.Selection.empty + ) + + + /* Status */ - if (file_name != "") { - val file = Path.explode(file_name).file - if (file.exists()) - props.load(new java.io.FileReader(file)) - props + sealed abstract class Status(val str: String) { + def merge(that: Status): Status = (this, that) match { + case (Ok, s) => s + case (Failed, _) => Failed + case (Skipped, Failed) => Failed + case (Skipped, _) => Skipped } - else - props + } + + object Status { + def merge(statuses: List[Status]): Status = + statuses.foldLeft(Ok: Status)(_ merge _) + + def from_results(results: Build.Results, session: String): Status = + if (results.cancelled(session)) Skipped + else if (results(session).ok) Ok + else Failed + } + + case object Ok extends Status("ok") + case object Skipped extends Status("skipped") + case object Failed extends Status("failed") + + + private def load_properties(): JProperties = { + val props = new JProperties + val file_name = Isabelle_System.getenv("ISABELLE_CI_PROPERTIES") + if (file_name.nonEmpty) { + val path = Path.explode(file_name) + if (path.is_file) props.load(Files.newBufferedReader(path.java_path)) + } + props } private def compute_timing(results: Build.Results, group: Option[String]): Timing = { - val timings = results.sessions.collect { - case session if group.forall(results.info(session).groups.contains(_)) => - results(session).timing - } + val timings = + results.sessions.collect { + case session if group.forall(results.info(session).groups.contains(_)) => + results(session).timing + } timings.foldLeft(Timing.zero)(_ + _) } - private def with_documents(options: Options): Options = { - if (documents) + private def with_documents(options: Options, config: Build_Config): Options = { + if (config.documents) { options .bool.update("browser_info", true) .string.update("document", "pdf") .string.update("document_variants", "document:outline=/proof,/ML") - else - options + } + else options } - - final def hg_id(path: Path): String = + def hg_id(path: Path): String = Mercurial.repository(path).id() - final def print_section(title: String): Unit = + def print_section(title: String): Unit = println(s"\n=== $title ===\n") + def build(profile: Profile, config: Build_Config): Unit = { + val isabelle_home = Path.explode(Isabelle_System.getenv_strict("ISABELLE_HOME")) + val isabelle_id = hg_id(isabelle_home) - final val isabelle_home = Path.explode(Isabelle_System.getenv_strict("ISABELLE_HOME")) - final val isabelle_id = hg_id(isabelle_home) - final val start_time = Instant.now().atZone(ZoneId.systemDefault).format(DateTimeFormatter.RFC_1123_DATE_TIME) + val start_time = + Instant.now().atZone(ZoneId.systemDefault).format(DateTimeFormatter.RFC_1123_DATE_TIME) - - override final def apply(args: List[String]): Unit = { print_section("CONFIGURATION") println(Build_Log.Settings.show()) val props = load_properties() - System.getProperties().asInstanceOf[JMap[AnyRef, AnyRef]].putAll(props) + System.getProperties.asInstanceOf[JMap[AnyRef, AnyRef]].putAll(props) val options = - with_documents(Options.init()) + with_documents(Options.init(), config) .int.update("parallel_proofs", 1) - .int.update("threads", threads) + .int.update("threads", profile.threads) - println(s"jobs = $jobs, threads = $threads, numa = $numa") + println(s"jobs = ${profile.jobs}, threads = ${profile.threads}, numa = ${profile.numa}") print_section("BUILD") println(s"Build started at $start_time") println(s"Isabelle id $isabelle_id") - val pre_result = pre_hook(args) + val pre_result = config.pre_hook() print_section("LOG") - val (results, elapsed_time) = build(options) + val (results, elapsed_time) = { + val progress = new Console_Progress(verbose = true) + val start_time = Time.now() + val results = progress.interrupt_handler { + Build.build( + options + "system_heaps", + selection = config.selection, + progress = progress, + clean_build = config.clean, + verbose = true, + numa_shuffling = profile.numa, + max_jobs = profile.jobs, + dirs = config.include, + select_dirs = config.select) + } + val end_time = Time.now() + (results, end_time - start_time) + } print_section("TIMING") @@ -117,46 +170,16 @@ print_section("FAILED SESSIONS") for (name <- results.sessions) { - if (results.cancelled(name)) { - println(s"Session $name: CANCELLED") - } + if (results.cancelled(name)) println(s"Session $name: CANCELLED") else { val result = results(name) - if (!result.ok) - println(s"Session $name: FAILED ${result.rc}") + if (!result.ok) println(s"Session $name: FAILED ${ result.rc }") } } } - val post_result = post_hook(results) - - System.exit(List(pre_result.rc, results.rc, post_result.rc).max) - } - - /* profile */ - - def threads: Int = Isabelle_System.hostname() match { - case "hpcisabelle" => 8 - case "lxcisa1" => 4 - case _ => 2 - } + val post_result = config.post_hook(results) - def jobs: Int = Isabelle_System.hostname() match { - case "hpcisabelle" => 8 - case "lxcisa1" => 10 - case _ => 2 + sys.exit(List(pre_result.rc, results.rc, post_result.rc).max) } - - def numa: Boolean = Isabelle_System.hostname() == "hpcisabelle" - - def documents: Boolean = true - def clean: Boolean = true - - def include: List[Path] - def select: List[Path] - - def pre_hook(args: List[String]): Result - def post_hook(results: Build.Results): Result - - def selection: Sessions.Selection -} +} \ No newline at end of file diff -r 0dd3ac5fdbaa -r e4bbe0b9288d src/Pure/PIDE/protocol.ML --- a/src/Pure/PIDE/protocol.ML Mon Jun 27 17:36:26 2022 +0200 +++ b/src/Pure/PIDE/protocol.ML Tue Jun 28 15:34:05 2022 +0200 @@ -18,13 +18,13 @@ raise Protocol_Command.STOP (Value.parse_int rc))); val _ = - Protocol_Command.define "Prover.options" + Protocol_Command.define_bytes "Prover.options" (fn [options_yxml] => - (Options.set_default (Options.decode (YXML.parse_body options_yxml)); + (Options.set_default (Options.decode (YXML.parse_body_bytes options_yxml)); Isabelle_Process.init_options_interactive ())); val _ = - Protocol_Command.define "Prover.init_session" + Protocol_Command.define_bytes "Prover.init_session" (fn [yxml] => Resources.init_session_yxml yxml); val _ = @@ -97,18 +97,18 @@ (fn [exec_id] => Execution.cancel (Document_ID.parse exec_id)); val _ = - Protocol_Command.define "Document.update" + Protocol_Command.define_bytes "Document.update" (Future.task_context "Document.update" (Future.new_group NONE) - (fn old_id_string :: new_id_string :: consolidate_yxml :: edits_yxml => + (fn old_id_bytes :: new_id_bytes :: consolidate_yxml :: edits_yxml => Document.change_state (fn state => let - val old_id = Document_ID.parse old_id_string; - val new_id = Document_ID.parse new_id_string; + val old_id = Document_ID.parse (Bytes.content old_id_bytes); + val new_id = Document_ID.parse (Bytes.content new_id_bytes); val consolidate = - YXML.parse_body consolidate_yxml |> + YXML.parse_body_bytes consolidate_yxml |> let open XML.Decode in list string end; val edits = - edits_yxml |> map (YXML.parse_body #> + edits_yxml |> map (YXML.parse_body_bytes #> let open XML.Decode in pair string (variant @@ -151,14 +151,14 @@ in Document.start_execution state' end))); val _ = - Protocol_Command.define "Document.remove_versions" + Protocol_Command.define_bytes "Document.remove_versions" (fn [versions_yxml] => Document.change_state (fn state => let val versions = - YXML.parse_body versions_yxml |> + YXML.parse_body_bytes versions_yxml |> let open XML.Decode in list int end; val state1 = Document.remove_versions versions state; - val _ = Output.protocol_message Markup.removed_versions [[XML.Text (versions_yxml)]]; + val _ = Output.protocol_message Markup.removed_versions [Bytes.contents_blob versions_yxml]; in state1 end)); val _ = diff -r 0dd3ac5fdbaa -r e4bbe0b9288d src/Pure/PIDE/resources.ML --- a/src/Pure/PIDE/resources.ML Mon Jun 27 17:36:26 2022 +0200 +++ b/src/Pure/PIDE/resources.ML Tue Jun 28 15:34:05 2022 +0200 @@ -16,7 +16,7 @@ scala_functions: (string * ((bool * bool) * Position.T)) list, global_theories: (string * string) list, loaded_theories: string list} -> unit - val init_session_yxml: string -> unit + val init_session_yxml: Bytes.T -> unit val init_session_env: unit -> unit val finish_session_base: unit -> unit val global_theory: string -> string option @@ -130,7 +130,7 @@ let val (session_positions, (session_directories, (bibtex_entries, (command_timings, (load_commands, (scala_functions, (global_theories, loaded_theories))))))) = - YXML.parse_body yxml |> + YXML.parse_body_bytes yxml |> let open XML.Decode in (pair (list (pair string properties)) (pair (list (pair string string)) @@ -155,7 +155,7 @@ (case getenv "ISABELLE_INIT_SESSION" of "" => () | name => - try File.read (Path.explode name) + try Bytes.read (Path.explode name) |> Option.app init_session_yxml); val _ = init_session_env (); diff -r 0dd3ac5fdbaa -r e4bbe0b9288d src/Pure/PIDE/yxml.ML --- a/src/Pure/PIDE/yxml.ML Mon Jun 27 17:36:26 2022 +0200 +++ b/src/Pure/PIDE/yxml.ML Tue Jun 28 15:34:05 2022 +0200 @@ -29,7 +29,9 @@ val output_markup: Markup.T -> string * string val output_markup_elem: Markup.T -> (string * string) * string val parse_body: string -> XML.body + val parse_body_bytes: Bytes.T -> XML.body val parse: string -> XML.tree + val parse_bytes: Bytes.T -> XML.tree val content_of: string -> string val is_wellformed: string -> bool end; @@ -119,6 +121,11 @@ Substring.tokens (fn c => c = X_char) #> map (Substring.fields (fn c => c = Y_char) #> map Substring.string); +val split_bytes = + Bytes.space_explode X + #> filter (fn "" => false | _ => true) + #> map (space_explode Y); + (* structural errors *) @@ -152,19 +159,25 @@ | parse_chunk ("" :: name :: atts) = push name (map parse_attrib atts) | parse_chunk txts = fold (add o XML.Text) txts; -in - -fun parse_body source = - (case fold parse_chunk (split_string source) [(("", []), [])] of +fun parse_body' chunks = + (case fold parse_chunk chunks [(("", []), [])] of [(("", _), result)] => rev result | ((name, _), _) :: _ => err_unbalanced name); -fun parse source = - (case parse_body source of +fun parse' chunks = + (case parse_body' chunks of [result] => result | [] => XML.Text "" | _ => err "multiple results"); +in + +val parse_body = parse_body' o split_string; +val parse_body_bytes = parse_body' o split_bytes; + +val parse = parse' o split_string; +val parse_bytes = parse' o split_bytes; + end; val content_of = parse_body #> XML.content_of; diff -r 0dd3ac5fdbaa -r e4bbe0b9288d src/Pure/System/isabelle_tool.scala --- a/src/Pure/System/isabelle_tool.scala Mon Jun 27 17:36:26 2022 +0200 +++ b/src/Pure/System/isabelle_tool.scala Tue Jun 28 15:34:05 2022 +0200 @@ -177,6 +177,7 @@ Build.isabelle_tool, Build_Docker.isabelle_tool, Build_Job.isabelle_tool, + CI_Build_Benchmark.isabelle_tool, Doc.isabelle_tool, Document_Build.isabelle_tool, Dump.isabelle_tool, diff -r 0dd3ac5fdbaa -r e4bbe0b9288d src/Pure/System/options.ML --- a/src/Pure/System/options.ML Mon Jun 27 17:36:26 2022 +0200 +++ b/src/Pure/System/options.ML Tue Jun 28 15:34:05 2022 +0200 @@ -211,8 +211,8 @@ (case getenv "ISABELLE_PROCESS_OPTIONS" of "" => () | name => - try File.read (Path.explode name) - |> Option.app (set_default o decode o YXML.parse_body)); + try Bytes.read (Path.explode name) + |> Option.app (set_default o decode o YXML.parse_body_bytes)); val _ = load_default (); val _ = ML_Print_Depth.set_print_depth (default_int "ML_print_depth"); diff -r 0dd3ac5fdbaa -r e4bbe0b9288d src/Pure/Tools/build.ML --- a/src/Pure/Tools/build.ML Mon Jun 27 17:36:26 2022 +0200 +++ b/src/Pure/Tools/build.ML Tue Jun 28 15:34:05 2022 +0200 @@ -66,12 +66,12 @@ (* build session *) val _ = - Protocol_Command.define "build_session" + Protocol_Command.define_bytes "build_session" (fn [resources_yxml, args_yxml] => let val _ = Resources.init_session_yxml resources_yxml; val (session_name, theories) = - YXML.parse_body args_yxml |> + YXML.parse_body_bytes args_yxml |> let open XML.Decode; val position = Position.of_properties o properties;