# HG changeset patch # User wenzelm # Date 1648848379 -7200 # Node ID cd9f2d3820144e193775ce64641ebf01ca61cc60 # Parent 2c3eadf5ca6f03092cc9521f77812fd13b59287c# Parent 42267c6502052fc4fee30aabf8531f81aacb681b merged diff -r 2c3eadf5ca6f -r cd9f2d382014 src/HOL/SPARK/Tools/spark.scala --- a/src/HOL/SPARK/Tools/spark.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/HOL/SPARK/Tools/spark.scala Fri Apr 01 23:26:19 2022 +0200 @@ -9,15 +9,12 @@ import isabelle._ -object SPARK -{ - class Load_Command1 extends Command_Span.Load_Command("spark_vcg", Scala_Project.here) - { +object SPARK { + class Load_Command1 extends Command_Span.Load_Command("spark_vcg", Scala_Project.here) { override val extensions: List[String] = List("vcg", "fdl", "rls") } - class Load_Command2 extends Command_Span.Load_Command("spark_siv", Scala_Project.here) - { + class Load_Command2 extends Command_Span.Load_Command("spark_siv", Scala_Project.here) { override val extensions: List[String] = List("siv", "fdl", "rls") } } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/HOL/Tools/ATP/system_on_tptp.scala --- a/src/HOL/Tools/ATP/system_on_tptp.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/HOL/Tools/ATP/system_on_tptp.scala Fri Apr 01 23:26:19 2022 +0200 @@ -11,8 +11,7 @@ import java.net.URL -object SystemOnTPTP -{ +object SystemOnTPTP { /* requests */ def get_url(options: Options): URL = Url(options.string("SystemOnTPTP")) @@ -20,8 +19,8 @@ def post_request( url: URL, parameters: List[(String, Any)], - timeout: Time = HTTP.Client.default_timeout): HTTP.Content = - { + timeout: Time = HTTP.Client.default_timeout + ): HTTP.Content = { try { HTTP.Client.post(url, ("NoHTML" -> 1) :: parameters, @@ -40,8 +39,7 @@ "ListStatus" -> "READY", "QuietFlag" -> "-q0")) - object List_Systems extends Scala.Fun_String("SystemOnTPTP.list_systems", thread = true) - { + object List_Systems extends Scala.Fun_String("SystemOnTPTP.list_systems", thread = true) { val here = Scala_Project.here def apply(url: String): String = list_systems(Url(url)).text } @@ -54,8 +52,8 @@ problem: String, extra: String = "", quiet: String = "01", - timeout: Time = Time.seconds(60)): HTTP.Content = - { + timeout: Time = Time.seconds(60) + ): HTTP.Content = { val paramaters = List( "SubmitButton" -> "RunSelectedSystems", @@ -69,11 +67,9 @@ post_request(url, paramaters, timeout = timeout + Time.seconds(15)) } - object Run_System extends Scala.Fun_Strings("SystemOnTPTP.run_system", thread = true) - { + object Run_System extends Scala.Fun_Strings("SystemOnTPTP.run_system", thread = true) { val here = Scala_Project.here - def apply(args: List[String]): List[String] = - { + def apply(args: List[String]): List[String] = { val List(url, system, problem_path, extra, Value.Int(timeout)) = args val problem = File.read(Path.explode(problem_path)) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/HOL/Tools/Mirabelle/mirabelle.scala --- a/src/HOL/Tools/Mirabelle/mirabelle.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/HOL/Tools/Mirabelle/mirabelle.scala Fri Apr 01 23:26:19 2022 +0200 @@ -9,12 +9,10 @@ import isabelle._ -object Mirabelle -{ +object Mirabelle { /* actions and options */ - def action_names(): List[String] = - { + def action_names(): List[String] = { val Pattern = """Mirabelle action: "(\w+)".*""".r (for { file <- @@ -25,8 +23,7 @@ } yield name).sorted } - def sledgehammer_options(): List[String] = - { + def sledgehammer_options(): List[String] = { val Pattern = """val .*K *= *"(.*)" *\(\*(.*)\*\)""".r split_lines(File.read(Path.explode("~~/src/HOL/Tools/Mirabelle/mirabelle_sledgehammer.ML"))). flatMap(line => line match { case Pattern(a, b) => Some (a + b) case _ => None }) @@ -46,8 +43,8 @@ select_dirs: List[Path] = Nil, numa_shuffling: Boolean = false, max_jobs: Int = 1, - verbose: Boolean = false): Build.Results = - { + verbose: Boolean = false + ): Build.Results = { require(!selection.requirements) Isabelle_System.make_directory(output_dir) @@ -70,8 +67,7 @@ val store = Sessions.store(build_options) - def session_setup(session_name: String, session: Session): Unit = - { + def session_setup(session_name: String, session: Session): Unit = { session.all_messages += Session.Consumer[Prover.Message]("mirabelle_export") { case msg: Prover.Protocol_Output => @@ -113,33 +109,33 @@ /* Isabelle tool wrapper */ val isabelle_tool = Isabelle_Tool("mirabelle", "testing tool for automated proof tools", - Scala_Project.here, args => - { - val build_options = Word.explode(Isabelle_System.getenv("ISABELLE_BUILD_OPTIONS")) + Scala_Project.here, + { args => + val build_options = Word.explode(Isabelle_System.getenv("ISABELLE_BUILD_OPTIONS")) - var options = Options.init(opts = build_options) - val mirabelle_dry_run = options.check_name("mirabelle_dry_run") - val mirabelle_max_calls = options.check_name("mirabelle_max_calls") - val mirabelle_randomize = options.check_name("mirabelle_randomize") - val mirabelle_stride = options.check_name("mirabelle_stride") - val mirabelle_timeout = options.check_name("mirabelle_timeout") - val mirabelle_output_dir = options.check_name("mirabelle_output_dir") + var options = Options.init(opts = build_options) + val mirabelle_dry_run = options.check_name("mirabelle_dry_run") + val mirabelle_max_calls = options.check_name("mirabelle_max_calls") + val mirabelle_randomize = options.check_name("mirabelle_randomize") + val mirabelle_stride = options.check_name("mirabelle_stride") + val mirabelle_timeout = options.check_name("mirabelle_timeout") + val mirabelle_output_dir = options.check_name("mirabelle_output_dir") - var actions: List[String] = Nil - var base_sessions: List[String] = Nil - var select_dirs: List[Path] = Nil - var numa_shuffling = false - var output_dir = Path.explode(mirabelle_output_dir.default_value) - var theories: List[String] = Nil - var exclude_session_groups: List[String] = Nil - var all_sessions = false - var dirs: List[Path] = Nil - var session_groups: List[String] = Nil - var max_jobs = 1 - var verbose = false - var exclude_sessions: List[String] = Nil + var actions: List[String] = Nil + var base_sessions: List[String] = Nil + var select_dirs: List[Path] = Nil + var numa_shuffling = false + var output_dir = Path.explode(mirabelle_output_dir.default_value) + var theories: List[String] = Nil + var exclude_session_groups: List[String] = Nil + var all_sessions = false + var dirs: List[Path] = Nil + var session_groups: List[String] = Nil + var max_jobs = 1 + var verbose = false + var exclude_sessions: List[String] = Nil - val getopts = Getopts(""" + val getopts = Getopts(""" Usage: isabelle mirabelle [OPTIONS] [SESSIONS ...] Options are: @@ -175,69 +171,69 @@ Available actions are:""" + action_names().mkString("\n ", "\n ", "") + """ For the ACTION "sledgehammer", the usual sledgehammer as well as the following mirabelle-specific OPTIONs are available:""" + - sledgehammer_options().mkString("\n ", "\n ", "\n"), - "A:" -> (arg => actions = actions ::: List(arg)), - "B:" -> (arg => base_sessions = base_sessions ::: List(arg)), - "D:" -> (arg => select_dirs = select_dirs ::: List(Path.explode(arg))), - "N" -> (_ => numa_shuffling = true), - "O:" -> (arg => output_dir = Path.explode(arg)), - "T:" -> (arg => theories = theories ::: List(arg)), - "X:" -> (arg => exclude_session_groups = exclude_session_groups ::: List(arg)), - "a" -> (_ => all_sessions = true), - "d:" -> (arg => dirs = dirs ::: List(Path.explode(arg))), - "g:" -> (arg => session_groups = session_groups ::: List(arg)), - "j:" -> (arg => max_jobs = Value.Int.parse(arg)), - "m:" -> (arg => options = options + ("mirabelle_max_calls=" + arg)), - "o:" -> (arg => options = options + arg), - "r:" -> (arg => options = options + ("mirabelle_randomize=" + arg)), - "s:" -> (arg => options = options + ("mirabelle_stride=" + arg)), - "t:" -> (arg => options = options + ("mirabelle_timeout=" + arg)), - "v" -> (_ => verbose = true), - "x:" -> (arg => exclude_sessions = exclude_sessions ::: List(arg)), - "y" -> (arg => options = options + ("mirabelle_dry_run=true"))) - - val sessions = getopts(args) - if (actions.isEmpty) getopts.usage() + sledgehammer_options().mkString("\n ", "\n ", "\n"), + "A:" -> (arg => actions = actions ::: List(arg)), + "B:" -> (arg => base_sessions = base_sessions ::: List(arg)), + "D:" -> (arg => select_dirs = select_dirs ::: List(Path.explode(arg))), + "N" -> (_ => numa_shuffling = true), + "O:" -> (arg => output_dir = Path.explode(arg)), + "T:" -> (arg => theories = theories ::: List(arg)), + "X:" -> (arg => exclude_session_groups = exclude_session_groups ::: List(arg)), + "a" -> (_ => all_sessions = true), + "d:" -> (arg => dirs = dirs ::: List(Path.explode(arg))), + "g:" -> (arg => session_groups = session_groups ::: List(arg)), + "j:" -> (arg => max_jobs = Value.Int.parse(arg)), + "m:" -> (arg => options = options + ("mirabelle_max_calls=" + arg)), + "o:" -> (arg => options = options + arg), + "r:" -> (arg => options = options + ("mirabelle_randomize=" + arg)), + "s:" -> (arg => options = options + ("mirabelle_stride=" + arg)), + "t:" -> (arg => options = options + ("mirabelle_timeout=" + arg)), + "v" -> (_ => verbose = true), + "x:" -> (arg => exclude_sessions = exclude_sessions ::: List(arg)), + "y" -> (arg => options = options + ("mirabelle_dry_run=true"))) - val progress = new Console_Progress(verbose = verbose) - - val start_date = Date.now() + val sessions = getopts(args) + if (actions.isEmpty) getopts.usage() - if (verbose) { - progress.echo("Started at " + Build_Log.print_date(start_date)) - } + val progress = new Console_Progress(verbose = verbose) - val results = - progress.interrupt_handler { - mirabelle(options, actions, output_dir.absolute, - theories = theories, - selection = Sessions.Selection( - all_sessions = all_sessions, - base_sessions = base_sessions, - exclude_session_groups = exclude_session_groups, - exclude_sessions = exclude_sessions, - session_groups = session_groups, - sessions = sessions), - progress = progress, - dirs = dirs, - select_dirs = select_dirs, - numa_shuffling = NUMA.enabled_warning(progress, numa_shuffling), - max_jobs = max_jobs, - verbose = verbose) + val start_date = Date.now() + + if (verbose) { + progress.echo("Started at " + Build_Log.print_date(start_date)) } - val end_date = Date.now() - val elapsed_time = end_date.time - start_date.time - - if (verbose) { - progress.echo("\nFinished at " + Build_Log.print_date(end_date)) - } + val results = + progress.interrupt_handler { + mirabelle(options, actions, output_dir.absolute, + theories = theories, + selection = Sessions.Selection( + all_sessions = all_sessions, + base_sessions = base_sessions, + exclude_session_groups = exclude_session_groups, + exclude_sessions = exclude_sessions, + session_groups = session_groups, + sessions = sessions), + progress = progress, + dirs = dirs, + select_dirs = select_dirs, + numa_shuffling = NUMA.enabled_warning(progress, numa_shuffling), + max_jobs = max_jobs, + verbose = verbose) + } - val total_timing = - results.sessions.iterator.map(a => results(a).timing).foldLeft(Timing.zero)(_ + _). - copy(elapsed = elapsed_time) - progress.echo(total_timing.message_resources) + val end_date = Date.now() + val elapsed_time = end_date.time - start_date.time + + if (verbose) { + progress.echo("\nFinished at " + Build_Log.print_date(end_date)) + } - sys.exit(results.rc) - }) + val total_timing = + results.sessions.iterator.map(a => results(a).timing).foldLeft(Timing.zero)(_ + _). + copy(elapsed = elapsed_time) + progress.echo(total_timing.message_resources) + + sys.exit(results.rc) + }) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/HOL/Tools/Nitpick/kodkod.scala --- a/src/HOL/Tools/Nitpick/kodkod.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/HOL/Tools/Nitpick/kodkod.scala Fri Apr 01 23:26:19 2022 +0200 @@ -14,18 +14,15 @@ import isabelle.kodkodi.{Context, KodkodiLexer, KodkodiParser} -object Kodkod -{ +object Kodkod { /** result **/ - sealed case class Result(rc: Int, out: String, err: String) - { + sealed case class Result(rc: Int, out: String, err: String) { def ok: Boolean = rc == Process_Result.RC.ok def check: String = if (ok) out else error(if (err.isEmpty) "Error" else err) - def encode: XML.Body = - { + def encode: XML.Body = { import XML.Encode._ triple(int, string, string)((rc, out, err)) } @@ -41,8 +38,8 @@ max_solutions: Int = Integer.MAX_VALUE, cleanup_inst: Boolean = false, timeout: Time = Time.zero, - max_threads: Int = 0): Result = - { + max_threads: Int = 0 + ): Result = { /* executor */ val pool_size = if (max_threads == 0) Isabelle_Thread.max_threads() else max_threads @@ -60,8 +57,7 @@ class Exit extends Exception("EXIT") - class Exec_Context extends Context - { + class Exec_Context extends Context { private var rc = Process_Result.RC.ok private val out = new StringBuilder private val err = new StringBuilder @@ -145,8 +141,7 @@ "solver: \"MiniSat\"\n" + File.read(Path.explode("$KODKODI/examples/weber3.kki"))).check - class Handler extends Session.Protocol_Handler - { + class Handler extends Session.Protocol_Handler { override def init(session: Session): Unit = warmup() } @@ -154,13 +149,10 @@ /** scala function **/ - object Fun extends Scala.Fun_String("kodkod", thread = true) - { + object Fun extends Scala.Fun_String("kodkod", thread = true) { val here = Scala_Project.here - def apply(args: String): String = - { - val (timeout, (solve_all, (max_solutions, (max_threads, kki)))) = - { + def apply(args: String): String = { + val (timeout, (solve_all, (max_solutions, (max_threads, kki)))) = { import XML.Decode._ pair(int, pair(bool, pair(int, pair(int, string))))(YXML.parse_body(args)) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/afp.scala --- a/src/Pure/Admin/afp.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/afp.scala Fri Apr 01 23:26:19 2022 +0200 @@ -11,8 +11,7 @@ import scala.collection.immutable.SortedMap -object AFP -{ +object AFP { val groups: Map[String, String] = Map("large" -> "full 64-bit memory model or word arithmetic required", "slow" -> "CPU time much higher than 60min (on mid-range hardware)", @@ -30,16 +29,14 @@ /* entries */ - def parse_date(s: String): Date = - { + def parse_date(s: String): Date = { val t = Date.Formatter.pattern("uuuu-MM-dd").parse(s) Date(LocalDate.from(t).atStartOfDay(Date.timezone_berlin)) } def trim_mail(s: String): String = s.replaceAll("<[^>]*>", "").trim - sealed case class Entry(name: String, metadata: Properties.T, sessions: List[String]) - { + sealed case class Entry(name: String, metadata: Properties.T, sessions: List[String]) { def get(prop: String): Option[String] = Properties.get(metadata, prop) def get_string(prop: String): String = get(prop).getOrElse("") def get_strings(prop: String): List[String] = @@ -66,8 +63,7 @@ } } -class AFP private(options: Options, val base_dir: Path) -{ +class AFP private(options: Options, val base_dir: Path) { override def toString: String = base_dir.expand.toString val main_dir: Path = base_dir + Path.explode("thys") @@ -75,8 +71,7 @@ /* metadata */ - private val entry_metadata: Map[String, Properties.T] = - { + private val entry_metadata: Map[String, Properties.T] = { val metadata_file = base_dir + Path.explode("metadata/metadata") var result = Map.empty[String, Properties.T] @@ -88,15 +83,13 @@ val Extra_Line = """^\s+(.*)$""".r val Blank_Line = """^\s*$""".r - def flush(): Unit = - { + def flush(): Unit = { if (section != "") result += (section -> props.reverse.filter(p => p._2.nonEmpty)) section = "" props = Nil } - for ((line, i) <- split_lines(File.read(metadata_file)).zipWithIndex) - { + for ((line, i) <- split_lines(File.read(metadata_file)).zipWithIndex) { def err(msg: String): Nothing = error(msg + Position.here(Position.Line_File(i + 1, metadata_file.expand.implode))) @@ -122,8 +115,7 @@ /* entries */ - val entries_map: SortedMap[String, AFP.Entry] = - { + val entries_map: SortedMap[String, AFP.Entry] = { val entries = for (name <- Sessions.parse_roots(main_dir + Sessions.ROOTS)) yield { val metadata = @@ -167,8 +159,7 @@ private def sessions_deps(entry: AFP.Entry): List[String] = entry.sessions.flatMap(sessions_structure.imports_graph.imm_preds).distinct.sorted - lazy val entries_graph: Graph[String, Unit] = - { + lazy val entries_graph: Graph[String, Unit] = { val session_entries = entries.foldLeft(Map.empty[String, String]) { case (m1, e) => e.sessions.foldLeft(m1) { case (m2, s) => m2 + (s -> e.name) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/build_csdp.scala --- a/src/Pure/Admin/build_csdp.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/build_csdp.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,16 +7,14 @@ package isabelle -object Build_CSDP -{ +object Build_CSDP { // Note: version 6.2.0 does not quite work for the "sos" proof method val default_download_url = "https://github.com/coin-or/Csdp/archive/releases/6.1.1.tar.gz" /* flags */ - sealed case class Flags(platform: String, CFLAGS: String = "", LIBS: String = "") - { + sealed case class Flags(platform: String, CFLAGS: String = "", LIBS: String = "") { val changed: List[(String, String)] = List("CFLAGS" -> CFLAGS, "LIBS" -> LIBS).filter(p => p._2.nonEmpty) @@ -26,8 +24,7 @@ Some(" * " + platform + ":\n" + changed.map(p => " " + Properties.Eq(p)) .mkString("\n")) - def change(path: Path): Unit = - { + def change(path: Path): Unit = { def change_line(line: String, p: (String, String)): String = line.replaceAll(p._1 + "=.*", Properties.Eq(p)) File.change_lines(path) { _.map(line => changed.foldLeft(line)(change_line)) } @@ -55,12 +52,11 @@ verbose: Boolean = false, progress: Progress = new Progress, target_dir: Path = Path.current, - mingw: MinGW = MinGW.none): Unit = - { + mingw: MinGW = MinGW.none + ): Unit = { mingw.check - Isabelle_System.with_tmp_dir("build")(tmp_dir => - { + Isabelle_System.with_tmp_dir("build") { tmp_dir => /* component */ val Archive_Name = """^.*?([^/]+)$""".r @@ -163,7 +159,7 @@ Makarius """ + Date.Format.date(Date.now()) + "\n") - }) + } } @@ -171,14 +167,13 @@ val isabelle_tool = Isabelle_Tool("build_csdp", "build prover component from official download", Scala_Project.here, - args => - { - var target_dir = Path.current - var mingw = MinGW.none - var download_url = default_download_url - var verbose = false + { args => + var target_dir = Path.current + var mingw = MinGW.none + var download_url = default_download_url + var verbose = false - val getopts = Getopts(""" + val getopts = Getopts(""" Usage: isabelle build_csdp [OPTIONS] Options are: @@ -190,17 +185,17 @@ Build prover component from official download. """, - "D:" -> (arg => target_dir = Path.explode(arg)), - "M:" -> (arg => mingw = MinGW(Path.explode(arg))), - "U:" -> (arg => download_url = arg), - "v" -> (_ => verbose = true)) + "D:" -> (arg => target_dir = Path.explode(arg)), + "M:" -> (arg => mingw = MinGW(Path.explode(arg))), + "U:" -> (arg => download_url = arg), + "v" -> (_ => verbose = true)) - val more_args = getopts(args) - if (more_args.nonEmpty) getopts.usage() + val more_args = getopts(args) + if (more_args.nonEmpty) getopts.usage() - val progress = new Console_Progress() + val progress = new Console_Progress() - build_csdp(download_url = download_url, verbose = verbose, progress = progress, - target_dir = target_dir, mingw = mingw) - }) + build_csdp(download_url = download_url, verbose = verbose, progress = progress, + target_dir = target_dir, mingw = mingw) + }) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/build_cygwin.scala --- a/src/Pure/Admin/build_cygwin.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/build_cygwin.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Build_Cygwin -{ +object Build_Cygwin { val default_mirror: String = "https://isabelle.sketis.net/cygwin_2021-1" val packages: List[String] = @@ -16,12 +15,11 @@ def build_cygwin(progress: Progress, mirror: String = default_mirror, - more_packages: List[String] = Nil): Unit = - { + more_packages: List[String] = Nil + ): Unit = { require(Platform.is_windows, "Windows platform expected") - Isabelle_System.with_tmp_dir("cygwin")(tmp_dir => - { + Isabelle_System.with_tmp_dir("cygwin") { tmp_dir => val cygwin = tmp_dir + Path.explode("cygwin") val cygwin_etc = cygwin + Path.explode("etc") val cygwin_isabelle = Isabelle_System.make_directory(cygwin + Path.explode("isabelle")) @@ -54,7 +52,7 @@ val archive = "cygwin-" + Date.Format.alt_date(Date.now()) + ".tar.gz" Isabelle_System.gnutar("-czf " + Bash.string(archive) + " cygwin", dir = tmp_dir).check - }) + } } @@ -62,13 +60,13 @@ val isabelle_tool = Isabelle_Tool("build_cygwin", "produce pre-canned Cygwin distribution for Isabelle", - Scala_Project.here, args => - { - var mirror = default_mirror - var more_packages: List[String] = Nil + Scala_Project.here, + { args => + var mirror = default_mirror + var more_packages: List[String] = Nil - val getopts = - Getopts(""" + val getopts = + Getopts(""" Usage: isabelle build_cygwin [OPTIONS] Options are: @@ -78,12 +76,12 @@ Produce pre-canned Cygwin distribution for Isabelle: this requires Windows administrator mode. """, - "R:" -> (arg => mirror = arg), - "p:" -> (arg => more_packages ::= arg)) + "R:" -> (arg => mirror = arg), + "p:" -> (arg => more_packages ::= arg)) - val more_args = getopts(args) - if (more_args.nonEmpty) getopts.usage() + val more_args = getopts(args) + if (more_args.nonEmpty) getopts.usage() - build_cygwin(new Console_Progress(), mirror = mirror, more_packages = more_packages) - }) + build_cygwin(new Console_Progress(), mirror = mirror, more_packages = more_packages) + }) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/build_doc.scala --- a/src/Pure/Admin/build_doc.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/build_doc.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Build_Doc -{ +object Build_Doc { /* build_doc */ def build_doc( @@ -17,8 +16,8 @@ all_docs: Boolean = false, max_jobs: Int = 1, sequential: Boolean = false, - docs: List[String] = Nil): Unit = - { + docs: List[String] = Nil + ): Unit = { val store = Sessions.store(options) val sessions_structure = Sessions.load_structure(options) @@ -74,15 +73,15 @@ /* Isabelle tool wrapper */ val isabelle_tool = - Isabelle_Tool("build_doc", "build Isabelle documentation", Scala_Project.here, args => - { - var all_docs = false - var max_jobs = 1 - var sequential = false - var options = Options.init() + Isabelle_Tool("build_doc", "build Isabelle documentation", Scala_Project.here, + { args => + var all_docs = false + var max_jobs = 1 + var sequential = false + var options = Options.init() - val getopts = - Getopts(""" + val getopts = + Getopts(""" Usage: isabelle build_doc [OPTIONS] [DOCS ...] Options are: @@ -94,20 +93,20 @@ Build Isabelle documentation from documentation sessions with suitable document_variants entry. """, - "a" -> (_ => all_docs = true), - "j:" -> (arg => max_jobs = Value.Int.parse(arg)), - "o:" -> (arg => options = options + arg), - "s" -> (_ => sequential = true)) + "a" -> (_ => all_docs = true), + "j:" -> (arg => max_jobs = Value.Int.parse(arg)), + "o:" -> (arg => options = options + arg), + "s" -> (_ => sequential = true)) - val docs = getopts(args) + val docs = getopts(args) - if (!all_docs && docs.isEmpty) getopts.usage() + if (!all_docs && docs.isEmpty) getopts.usage() - val progress = new Console_Progress() + val progress = new Console_Progress() - progress.interrupt_handler { - build_doc(options, progress = progress, all_docs = all_docs, max_jobs = max_jobs, - sequential = sequential, docs = docs) - } - }) + progress.interrupt_handler { + build_doc(options, progress = progress, all_docs = all_docs, max_jobs = max_jobs, + sequential = sequential, docs = docs) + } + }) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/build_e.scala --- a/src/Pure/Admin/build_e.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/build_e.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Build_E -{ +object Build_E { /* build E prover */ val default_version = "2.6" @@ -19,10 +18,9 @@ download_url: String = default_download_url, verbose: Boolean = false, progress: Progress = new Progress, - target_dir: Path = Path.current): Unit = - { - Isabelle_System.with_tmp_dir("build")(tmp_dir => - { + target_dir: Path = Path.current + ): Unit = { + Isabelle_System.with_tmp_dir("build") { tmp_dir => /* component */ val component_name = "e-" + version @@ -51,8 +49,7 @@ progress.echo("Building E prover for " + platform_name + " ...") val build_dir = tmp_dir + Path.basic("E") - val build_options = - { + val build_options = { val result = Isabelle_System.bash("./configure --help", cwd = build_dir.file) if (result.check.out.containsSlice("--enable-ho")) " --enable-ho" else "" } @@ -103,21 +100,20 @@ Makarius """ + Date.Format.date(Date.now()) + "\n") - }) + } } /* Isabelle tool wrapper */ val isabelle_tool = Isabelle_Tool("build_e", "build prover component from source distribution", Scala_Project.here, - args => - { - var target_dir = Path.current - var version = default_version - var download_url = default_download_url - var verbose = false + { args => + var target_dir = Path.current + var version = default_version + var download_url = default_download_url + var verbose = false - val getopts = Getopts(""" + val getopts = Getopts(""" Usage: isabelle build_e [OPTIONS] Options are: @@ -129,17 +125,17 @@ Build prover component from the specified source distribution. """, - "D:" -> (arg => target_dir = Path.explode(arg)), - "U:" -> (arg => download_url = arg), - "V:" -> (arg => version = arg), - "v" -> (_ => verbose = true)) + "D:" -> (arg => target_dir = Path.explode(arg)), + "U:" -> (arg => download_url = arg), + "V:" -> (arg => version = arg), + "v" -> (_ => verbose = true)) - val more_args = getopts(args) - if (more_args.nonEmpty) getopts.usage() + val more_args = getopts(args) + if (more_args.nonEmpty) getopts.usage() - val progress = new Console_Progress() + val progress = new Console_Progress() - build_e(version = version, download_url = download_url, - verbose = verbose, progress = progress, target_dir = target_dir) - }) + build_e(version = version, download_url = download_url, + verbose = verbose, progress = progress, target_dir = target_dir) + }) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/build_fonts.scala --- a/src/Pure/Admin/build_fonts.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/build_fonts.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,12 +7,10 @@ package isabelle -object Build_Fonts -{ +object Build_Fonts { /* relevant codepoint ranges */ - object Range - { + object Range { def base_font: Seq[Int] = (0x0020 to 0x007e) ++ // ASCII (0x00a0 to 0x024f) ++ // Latin Extended-A/B @@ -137,8 +135,8 @@ plain: String = "", bold: String = "", italic: String = "", - bold_italic: String = "") - { + bold_italic: String = "" + ) { val fonts: List[String] = proper_string(plain).toList ::: proper_string(bold).toList ::: @@ -148,8 +146,7 @@ def get(index: Int): String = fonts(index % fonts.length) } - object Family - { + object Family { def isabelle_symbols: Family = Family( plain = "IsabelleSymbols.sfd", @@ -183,8 +180,7 @@ /* hinting */ // see https://www.freetype.org/ttfautohint/doc/ttfautohint.html - private def auto_hint(source: Path, target: Path): Unit = - { + private def auto_hint(source: Path, target: Path): Unit = { Isabelle_System.bash("ttfautohint -i " + File.bash_path(source) + " " + File.bash_path(target)).check } @@ -197,8 +193,7 @@ /* build fonts */ - private def find_file(dirs: List[Path], name: String): Path = - { + private def find_file(dirs: List[Path], name: String): Path = { val path = Path.explode(name) dirs.collectFirst({ case dir if (dir + path).is_file => dir + path }) match { case Some(file) => file @@ -220,8 +215,8 @@ target_prefix: String = "Isabelle", target_version: String = "", target_dir: Path = default_target_dir, - progress: Progress = new Progress): Unit = - { + progress: Progress = new Progress + ): Unit = { Isabelle_System.require_command("ttfautohint") progress.echo("Directory " + target_dir) @@ -243,8 +238,7 @@ val target_names = source_names.update(prefix = target_prefix, version = target_version) - Isabelle_System.with_tmp_file("font", "ttf")(tmp_file => - { + Isabelle_System.with_tmp_file("font", "ttf") { tmp_file => for (hinted <- hinting) { val target_file = target_dir + hinted_path(hinted) + target_names.ttf progress.echo("Font " + target_file.toString + " ...") @@ -271,7 +265,7 @@ Fontforge.close) ).check } - }) + } (target_names.ttf, index) } @@ -342,11 +336,11 @@ /* Isabelle tool wrapper */ val isabelle_tool = - Isabelle_Tool("build_fonts", "construct Isabelle fonts", Scala_Project.here, args => - { - var source_dirs: List[Path] = Nil + Isabelle_Tool("build_fonts", "construct Isabelle fonts", Scala_Project.here, + { args => + var source_dirs: List[Path] = Nil - val getopts = Getopts(""" + val getopts = Getopts(""" Usage: isabelle build_fonts [OPTIONS] Options are: @@ -354,17 +348,17 @@ Construct Isabelle fonts from DejaVu font families and Isabelle symbols. """, - "d:" -> (arg => source_dirs = source_dirs ::: List(Path.explode(arg)))) + "d:" -> (arg => source_dirs = source_dirs ::: List(Path.explode(arg)))) - val more_args = getopts(args) - if (more_args.nonEmpty) getopts.usage() + val more_args = getopts(args) + if (more_args.nonEmpty) getopts.usage() - val target_version = Date.Format.alt_date(Date.now()) - val target_dir = Path.explode("isabelle_fonts-" + target_version) + val target_version = Date.Format.alt_date(Date.now()) + val target_dir = Path.explode("isabelle_fonts-" + target_version) - val progress = new Console_Progress + val progress = new Console_Progress - build_fonts(source_dirs = source_dirs, target_dir = target_dir, - target_version = target_version, progress = progress) - }) + build_fonts(source_dirs = source_dirs, target_dir = target_dir, + target_version = target_version, progress = progress) + }) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/build_history.scala --- a/src/Pure/Admin/build_history.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/build_history.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Build_History -{ +object Build_History { /* log files */ val engine = "build_history" @@ -23,10 +22,9 @@ arch_64: Boolean, heap: Int, max_heap: Option[Int], - more_settings: List[String]): String = - { - val (ml_platform, ml_settings) = - { + more_settings: List[String] + ): String = { + val (ml_platform, ml_settings) = { val cygwin_32 = "x86-cygwin" val windows_32 = "x86-windows" val windows_64 = "x86_64-windows" @@ -123,8 +121,8 @@ more_preferences: List[String] = Nil, verbose: Boolean = false, build_tags: List[String] = Nil, - build_args: List[String] = Nil): List[(Process_Result, Path)] = - { + build_args: List[String] = Nil + ): List[(Process_Result, Path)] = { /* sanity checks */ if (File.eq(Path.ISABELLE_HOME, root)) @@ -148,8 +146,7 @@ /* checkout Isabelle + AFP repository */ - def checkout(dir: Path, version: String): String = - { + def checkout(dir: Path, version: String): String = { val hg = Mercurial.repository(dir) hg.update(rev = version, clean = true) progress.echo_if(verbose, hg.log(version, options = "-l1")) @@ -188,8 +185,7 @@ val build_group_id = build_host + ":" + build_history_date.time.ms var first_build = true - for ((threads, processes) <- multicore_list) yield - { + for ((threads, processes) <- multicore_list) yield { /* init settings */ val component_settings = @@ -285,86 +281,81 @@ build_out_progress.echo("Reading session build info ...") val session_build_info = - build_info.finished_sessions.flatMap(session_name => - { - val database = isabelle_output + store.database(session_name) + build_info.finished_sessions.flatMap { session_name => + val database = isabelle_output + store.database(session_name) + + if (database.is_file) { + using(SQLite.open_database(database)) { db => + val theory_timings = + try { + store.read_theory_timings(db, session_name).map(ps => + Protocol.Theory_Timing_Marker((Build_Log.SESSION_NAME, session_name) :: ps)) + } + catch { case ERROR(_) => Nil } - if (database.is_file) { - using(SQLite.open_database(database))(db => - { - val theory_timings = - try { - store.read_theory_timings(db, session_name).map(ps => - Protocol.Theory_Timing_Marker((Build_Log.SESSION_NAME, session_name) :: ps)) - } - catch { case ERROR(_) => Nil } + val session_sources = + store.read_build(db, session_name).map(_.sources) match { + case Some(sources) if sources.length == SHA1.digest_length => + List("Sources " + session_name + " " + sources) + case _ => Nil + } - val session_sources = - store.read_build(db, session_name).map(_.sources) match { - case Some(sources) if sources.length == SHA1.digest_length => - List("Sources " + session_name + " " + sources) - case _ => Nil - } - - theory_timings ::: session_sources - }) + theory_timings ::: session_sources } - else Nil - }) + } + else Nil + } build_out_progress.echo("Reading ML statistics ...") 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) + 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_ml_statistics(db, session_name)) - } - else if (log_gz.is_file) { - Build_Log.Log_File(log_gz).parse_session_info(ml_statistics = true).ml_statistics - } - else Nil + val properties = + if (database.is_file) { + using(SQLite.open_database(database))(db => + store.read_ml_statistics(db, session_name)) + } + else if (log_gz.is_file) { + Build_Log.Log_File(log_gz).parse_session_info(ml_statistics = true).ml_statistics + } + else Nil - val trimmed_properties = - if (ml_statistics_step <= 0) Nil - else if (ml_statistics_step == 1) properties - else { - (for { (ps, i) <- properties.iterator.zipWithIndex if i % ml_statistics_step == 0 } - yield ps).toList - } + val trimmed_properties = + if (ml_statistics_step <= 0) Nil + else if (ml_statistics_step == 1) properties + else { + (for { (ps, i) <- properties.iterator.zipWithIndex if i % ml_statistics_step == 0 } + yield ps).toList + } - trimmed_properties.map(ps => (Build_Log.SESSION_NAME -> session_name) :: ps) - }) + trimmed_properties.map(ps => (Build_Log.SESSION_NAME -> session_name) :: ps) + } build_out_progress.echo("Reading error messages ...") val session_errors = - build_info.failed_sessions.flatMap(session_name => - { - val database = isabelle_output + store.database(session_name) - val errors = - if (database.is_file) { - try { - using(SQLite.open_database(database))(db => store.read_errors(db, session_name)) - } // column "errors" could be missing - catch { case _: java.sql.SQLException => Nil } - } - else Nil - errors.map(msg => List(Build_Log.SESSION_NAME -> session_name, Markup.CONTENT -> msg)) - }) + build_info.failed_sessions.flatMap { session_name => + val database = isabelle_output + store.database(session_name) + val errors = + if (database.is_file) { + try { + using(SQLite.open_database(database))(db => store.read_errors(db, session_name)) + } // column "errors" could be missing + catch { case _: java.sql.SQLException => Nil } + } + else Nil + errors.map(msg => List(Build_Log.SESSION_NAME -> session_name, Markup.CONTENT -> msg)) + } build_out_progress.echo("Reading heap sizes ...") 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 - }) + 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 + } build_out_progress.echo("Writing log file " + log_path.xz + " ...") File.write_xz(log_path.xz, @@ -392,8 +383,7 @@ /* command line entry point */ - private object Multicore - { + private object Multicore { private val Pat1 = """^(\d+)$""".r private val Pat2 = """^(\d+)x(\d+)$""".r @@ -405,8 +395,7 @@ } } - def main(args: Array[String]): Unit = - { + def main(args: Array[String]): Unit = { Command_Line.tool { var afp_rev: Option[String] = None var multicore_base = false @@ -546,8 +535,8 @@ rev: String = "", afp_rev: Option[String] = None, options: String = "", - args: String = ""): List[(String, Bytes)] = - { + args: String = "" + ): List[(String, Bytes)] = { /* Isabelle self repository */ val self_hg = @@ -594,8 +583,7 @@ /* Admin/build_history */ - ssh.with_tmp_dir(tmp_dir => - { + ssh.with_tmp_dir { tmp_dir => val output_file = tmp_dir + Path.explode("output") val rev_options = if (rev == "") "" else " -r " + Bash.string(rev_id) @@ -619,6 +607,6 @@ ssh.rm(log) (log.file_name, bytes) } - }) + } } } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/build_jcef.scala --- a/src/Pure/Admin/build_jcef.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/build_jcef.scala Fri Apr 01 23:26:19 2022 +0200 @@ -11,8 +11,7 @@ package isabelle -object Build_JCEF -{ +object Build_JCEF { /* platform information */ sealed case class JCEF_Platform( @@ -48,8 +47,8 @@ base_url: String = default_url, version: String = default_version, target_dir: Path = Path.current, - progress: Progress = new Progress): Unit = - { + progress: Progress = new Progress + ): Unit = { /* component name */ val component = "jcef-" + version @@ -61,8 +60,7 @@ val platform_settings: List[String] = for (platform <- platforms) yield { - Isabelle_System.with_tmp_file("archive", ext = "tar.gz")(archive_file => - { + Isabelle_System.with_tmp_file("archive", ext = "tar.gz") { archive_file => val url = base_url + "/" + version + "/" + platform.archive Isabelle_System.download_file(url, archive_file, progress = progress) @@ -87,7 +85,7 @@ " " + "ISABELLE_JCEF_LIB=\"$ISABELLE_JCEF_HOME/" + platform.lib + "\"\n" + " " + platform.library + "\n" + " " + ";;" - }) + } } @@ -144,13 +142,13 @@ val isabelle_tool = Isabelle_Tool("build_jcef", "build component for Java Chromium Embedded Framework", - Scala_Project.here, args => - { - var target_dir = Path.current - var base_url = default_url - var version = default_version + Scala_Project.here, + { args => + var target_dir = Path.current + var base_url = default_url + var version = default_version - val getopts = Getopts(""" + val getopts = Getopts(""" Usage: isabelle build_jcef [OPTIONS] Options are: @@ -160,16 +158,16 @@ Build component for Java Chromium Embedded Framework. """, - "D:" -> (arg => target_dir = Path.explode(arg)), - "U:" -> (arg => base_url = arg), - "V:" -> (arg => version = arg)) + "D:" -> (arg => target_dir = Path.explode(arg)), + "U:" -> (arg => base_url = arg), + "V:" -> (arg => version = arg)) - val more_args = getopts(args) - if (more_args.nonEmpty) getopts.usage() + val more_args = getopts(args) + if (more_args.nonEmpty) getopts.usage() - val progress = new Console_Progress() + val progress = new Console_Progress() - build_jcef(base_url = base_url, version = version, target_dir = target_dir, - progress = progress) - }) + build_jcef(base_url = base_url, version = version, target_dir = target_dir, + progress = progress) + }) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/build_jdk.scala --- a/src/Pure/Admin/build_jdk.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/build_jdk.scala Fri Apr 01 23:26:19 2022 +0200 @@ -14,8 +14,7 @@ import scala.util.matching.Regex -object Build_JDK -{ +object Build_JDK { /* platform and version information */ sealed case class JDK_Platform( @@ -23,16 +22,14 @@ platform_regex: Regex, exe: String = "java", macos_home: Boolean = false, - jdk_version: String = "") - { + jdk_version: String = "" + ) { override def toString: String = platform_name def platform_path: Path = Path.explode(platform_name) - def detect(jdk_dir: Path): Option[JDK_Platform] = - { - val major_version = - { + def detect(jdk_dir: Path): Option[JDK_Platform] = { + val major_version = { val Major_Version = """.*jdk(\d+).*$""".r val jdk_name = jdk_dir.file.getName jdk_name match { @@ -112,8 +109,7 @@ /* extract archive */ - def extract_archive(dir: Path, archive: Path): JDK_Platform = - { + def extract_archive(dir: Path, archive: Path): JDK_Platform = { try { val tmp_dir = Isabelle_System.make_directory(dir + Path.explode("tmp")) @@ -152,61 +148,60 @@ def build_jdk( archives: List[Path], progress: Progress = new Progress, - target_dir: Path = Path.current): Unit = - { + target_dir: Path = Path.current + ): Unit = { if (Platform.is_windows) error("Cannot build jdk on Windows") - Isabelle_System.with_tmp_dir("jdk")(dir => - { - progress.echo("Extracting ...") - val platforms = archives.map(extract_archive(dir, _)) + Isabelle_System.with_tmp_dir("jdk") { dir => + progress.echo("Extracting ...") + val platforms = archives.map(extract_archive(dir, _)) - val jdk_version = - platforms.map(_.jdk_version).distinct match { - case List(version) => version - case Nil => error("No archives") - case versions => - error("Archives contain multiple JDK versions: " + commas_quote(versions)) - } - - templates.filterNot(p1 => platforms.exists(p2 => p1.platform_name == p2.platform_name)) - match { - case Nil => - case missing => error("Missing platforms: " + commas_quote(missing.map(_.platform_name))) + val jdk_version = + platforms.map(_.jdk_version).distinct match { + case List(version) => version + case Nil => error("No archives") + case versions => + error("Archives contain multiple JDK versions: " + commas_quote(versions)) } - val jdk_name = "jdk-" + jdk_version - val jdk_path = Path.explode(jdk_name) - val component_dir = dir + jdk_path + templates.filterNot(p1 => platforms.exists(p2 => p1.platform_name == p2.platform_name)) + match { + case Nil => + case missing => error("Missing platforms: " + commas_quote(missing.map(_.platform_name))) + } - Isabelle_System.make_directory(component_dir + Path.explode("etc")) - File.write(Components.settings(component_dir), settings) - File.write(component_dir + Path.explode("README"), readme(jdk_version)) + val jdk_name = "jdk-" + jdk_version + val jdk_path = Path.explode(jdk_name) + val component_dir = dir + jdk_path - for (platform <- platforms) { - Isabelle_System.move_file(dir + platform.platform_path, component_dir) - } + Isabelle_System.make_directory(component_dir + Path.explode("etc")) + File.write(Components.settings(component_dir), settings) + File.write(component_dir + Path.explode("README"), readme(jdk_version)) + + for (platform <- platforms) { + Isabelle_System.move_file(dir + platform.platform_path, component_dir) + } - for (file <- File.find_files(component_dir.file, include_dirs = true)) { - val path = file.toPath - val perms = Files.getPosixFilePermissions(path) - perms.add(PosixFilePermission.OWNER_READ) - perms.add(PosixFilePermission.GROUP_READ) - perms.add(PosixFilePermission.OTHERS_READ) + for (file <- File.find_files(component_dir.file, include_dirs = true)) { + val path = file.toPath + val perms = Files.getPosixFilePermissions(path) + perms.add(PosixFilePermission.OWNER_READ) + perms.add(PosixFilePermission.GROUP_READ) + perms.add(PosixFilePermission.OTHERS_READ) + perms.add(PosixFilePermission.OWNER_WRITE) + if (file.isDirectory) { perms.add(PosixFilePermission.OWNER_WRITE) - if (file.isDirectory) { - perms.add(PosixFilePermission.OWNER_WRITE) - perms.add(PosixFilePermission.OWNER_EXECUTE) - perms.add(PosixFilePermission.GROUP_EXECUTE) - perms.add(PosixFilePermission.OTHERS_EXECUTE) - } - Files.setPosixFilePermissions(path, perms) + perms.add(PosixFilePermission.OWNER_EXECUTE) + perms.add(PosixFilePermission.GROUP_EXECUTE) + perms.add(PosixFilePermission.OTHERS_EXECUTE) } + Files.setPosixFilePermissions(path, perms) + } - progress.echo("Archiving ...") - Isabelle_System.gnutar( - "-czf " + File.bash_path(target_dir + jdk_path.tar.gz) + " " + jdk_name, dir = dir).check - }) + progress.echo("Archiving ...") + Isabelle_System.gnutar( + "-czf " + File.bash_path(target_dir + jdk_path.tar.gz) + " " + jdk_name, dir = dir).check + } } @@ -214,11 +209,11 @@ val isabelle_tool = Isabelle_Tool("build_jdk", "build Isabelle jdk component from original archives", - Scala_Project.here, args => - { - var target_dir = Path.current + Scala_Project.here, + { args => + var target_dir = Path.current - val getopts = Getopts(""" + val getopts = Getopts(""" Usage: isabelle build_jdk [OPTIONS] ARCHIVES... Options are: @@ -227,14 +222,14 @@ Build jdk component from tar.gz archives, with original jdk archives for Linux, Windows, and macOS. """, - "D:" -> (arg => target_dir = Path.explode(arg))) + "D:" -> (arg => target_dir = Path.explode(arg))) - val more_args = getopts(args) - if (more_args.isEmpty) getopts.usage() + val more_args = getopts(args) + if (more_args.isEmpty) getopts.usage() - val archives = more_args.map(Path.explode) - val progress = new Console_Progress() + val archives = more_args.map(Path.explode) + val progress = new Console_Progress() - build_jdk(archives = archives, progress = progress, target_dir = target_dir) - }) + build_jdk(archives = archives, progress = progress, target_dir = target_dir) + }) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/build_jedit.scala --- a/src/Pure/Admin/build_jedit.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/build_jedit.scala Fri Apr 01 23:26:19 2022 +0200 @@ -12,12 +12,10 @@ import scala.jdk.CollectionConverters._ -object Build_JEdit -{ +object Build_JEdit { /* modes */ - object Mode - { + object Mode { val empty: Mode = new Mode("", "", Nil) val init: Mode = @@ -28,8 +26,7 @@ ("tabSize" -> "2") + ("indentSize" -> "2") - val list: List[Mode] = - { + val list: List[Mode] = { val isabelle_news: Mode = init.define("isabelle-news", "Isabelle NEWS") val isabelle: Mode = @@ -53,8 +50,7 @@ } } - final case class Mode private(name: String, description: String, rev_props: Properties.T) - { + final case class Mode private(name: String, description: String, rev_props: Properties.T) { override def toString: String = name def define(a: String, b: String): Mode = new Mode(a, b, rev_props) @@ -62,8 +58,7 @@ def + (entry: Properties.Entry): Mode = new Mode(name, description, Properties.put(rev_props, entry)) - def write(dir: Path): Unit = - { + def write(dir: Path): Unit = { require(name.nonEmpty && description.nonEmpty, "Bad Isabelle/jEdit mode content") val properties = @@ -112,8 +107,8 @@ version: String, original: Boolean = false, java_home: Path = default_java_home, - progress: Progress = new Progress): Unit = - { + progress: Progress = new Progress + ): Unit = { Isabelle_System.require_command("ant", test = "-version") Isabelle_System.require_command("patch") Isabelle_System.require_command("unzip", test = "-h") @@ -131,8 +126,7 @@ val jedit_dir = Isabelle_System.make_directory(component_dir + Path.basic(jedit)) val jedit_patched_dir = component_dir + Path.basic(jedit_patched) - def download_jedit(dir: Path, name: String, target_name: String = ""): Path = - { + def download_jedit(dir: Path, name: String, target_name: String = ""): Path = { val jedit_name = jedit + name val url = "https://sourceforge.net/projects/jedit/files/jedit/" + @@ -142,8 +136,7 @@ path } - Isabelle_System.with_tmp_dir("tmp")(tmp_dir => - { + Isabelle_System.with_tmp_dir("tmp") { tmp_dir => /* original version */ val install_path = download_jedit(tmp_dir, "install.jar") @@ -198,7 +191,7 @@ "no_build = true\n" + "requirements = env:JEDIT_JARS\n" + ("sources =" :: java_sources.sorted.map(" " + _)).mkString("", " \\\n", "\n")) - }) + } /* jars */ @@ -210,14 +203,13 @@ } for { (name, vers) <- download_plugins } { - Isabelle_System.with_tmp_file("tmp", ext = "zip")(zip_path => - { + Isabelle_System.with_tmp_file("tmp", ext = "zip") { zip_path => val url = "https://sourceforge.net/projects/jedit-plugins/files/" + name + "/" + vers + "/" + name + "-" + vers + "-bin.zip/download" Isabelle_System.download_file(url, zip_path, progress = progress) Isabelle_System.bash("unzip -x " + File.bash_path(zip_path), cwd = jars_dir.file).check - }) + } } @@ -519,14 +511,14 @@ val isabelle_tool = Isabelle_Tool("build_jedit", "build Isabelle component from the jEdit text-editor", - Scala_Project.here, args => - { - var target_dir = Path.current - var java_home = default_java_home - var original = false - var version = default_version + Scala_Project.here, + { args => + var target_dir = Path.current + var java_home = default_java_home + var original = false + var version = default_version - val getopts = Getopts(""" + val getopts = Getopts(""" Usage: isabelle build_jedit [OPTIONS] Options are: @@ -537,18 +529,18 @@ Build auxiliary jEdit component from original sources, with some patches. """, - "D:" -> (arg => target_dir = Path.explode(arg)), - "J:" -> (arg => java_home = Path.explode(arg)), - "O" -> (_ => original = true), - "V:" -> (arg => version = arg)) + "D:" -> (arg => target_dir = Path.explode(arg)), + "J:" -> (arg => java_home = Path.explode(arg)), + "O" -> (_ => original = true), + "V:" -> (arg => version = arg)) - val more_args = getopts(args) - if (more_args.nonEmpty) getopts.usage() + val more_args = getopts(args) + if (more_args.nonEmpty) getopts.usage() - val component_dir = target_dir + Path.basic("jedit-" + Date.Format.alt_date(Date.now())) - val progress = new Console_Progress() + val component_dir = target_dir + Path.basic("jedit-" + Date.Format.alt_date(Date.now())) + val progress = new Console_Progress() - build_jedit(component_dir, version, original = original, - java_home = java_home, progress = progress) - }) + build_jedit(component_dir, version, original = original, + java_home = java_home, progress = progress) + }) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/build_log.scala --- a/src/Pure/Admin/build_log.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/build_log.scala Fri Apr 01 23:26:19 2022 +0200 @@ -16,14 +16,12 @@ import scala.util.matching.Regex -object Build_Log -{ +object Build_Log { /** content **/ /* properties */ - object Prop - { + object Prop { val build_tags = SQL.Column.string("build_tags") // lines val build_args = SQL.Column.string("build_args") // lines val build_group_id = SQL.Column.string("build_group_id") @@ -43,8 +41,7 @@ /* settings */ - object Settings - { + object Settings { val ISABELLE_BUILD_OPTIONS = SQL.Column.string("ISABELLE_BUILD_OPTIONS") val ML_PLATFORM = SQL.Column.string("ML_PLATFORM") val ML_HOME = SQL.Column.string("ML_HOME") @@ -57,8 +54,7 @@ type Entry = (String, String) type T = List[Entry] - object Entry - { + object Entry { def unapply(s: String): Option[Entry] = for { (a, b) <- Properties.Eq.unapply(s) } yield (a, Library.perhaps_unquote(b)) @@ -93,12 +89,10 @@ def print_date(date: Date): String = Log_File.Date_Format(date) - object Log_File - { + object Log_File { /* log file */ - def plain_name(name: String): String = - { + def plain_name(name: String): String = { List(".log", ".log.gz", ".log.xz", ".gz", ".xz").find(name.endsWith) match { case Some(s) => Library.try_unsuffix(s, name).get case None => name @@ -111,8 +105,7 @@ def apply(name: String, text: String): Log_File = new Log_File(plain_name(name), Library.trim_split_lines(text)) - def apply(file: JFile): Log_File = - { + def apply(file: JFile): Log_File = { val name = file.getName val text = if (name.endsWith(".gz")) File.read_gzip(file) @@ -130,8 +123,8 @@ prefixes: List[String] = List(Build_History.log_prefix, Identify.log_prefix, Identify.log_prefix2, Isatest.log_prefix, AFP_Test.log_prefix, Jenkins.log_prefix), - suffixes: List[String] = List(".log", ".log.gz", ".log.xz")): Boolean = - { + suffixes: List[String] = List(".log", ".log.gz", ".log.xz") + ): Boolean = { val name = file.getName prefixes.exists(name.startsWith) && @@ -144,8 +137,7 @@ /* date format */ - val Date_Format = - { + val Date_Format = { val fmts = Date.Formatter.variants( List("EEE MMM d HH:mm:ss O yyyy", "EEE MMM d HH:mm:ss VV yyyy"), @@ -185,8 +177,7 @@ } } - class Log_File private(val name: String, val lines: List[String]) - { + class Log_File private(val name: String, val lines: List[String]) { log_file => override def toString: String = name @@ -199,8 +190,7 @@ /* date format */ - object Strict_Date - { + object Strict_Date { def unapply(s: String): Some[Date] = try { Some(Log_File.Date_Format.parse(s)) } catch { case exn: DateTimeParseException => log_file.err(exn.getMessage) } @@ -269,13 +259,11 @@ /** digested meta info: produced by Admin/build_history in log.xz file **/ - object Meta_Info - { + object Meta_Info { val empty: Meta_Info = Meta_Info(Nil, Nil) } - sealed case class Meta_Info(props: Properties.T, settings: Settings.T) - { + sealed case class Meta_Info(props: Properties.T, settings: Settings.T) { def is_empty: Boolean = props.isEmpty && settings.isEmpty def get(c: SQL.Column): Option[String] = @@ -286,8 +274,7 @@ get(c).map(Log_File.Date_Format.parse) } - object Identify - { + object Identify { val log_prefix = "isabelle_identify_" val log_prefix2 = "plain_identify_" @@ -308,8 +295,7 @@ val AFP_Version = List(new Regex("""^AFP version: (\S+)$""")) } - object Isatest - { + object Isatest { val log_prefix = "isatest-makeall-" val engine = "isatest" val Start = new Regex("""^------------------- starting test --- (.+) --- (.+)$""") @@ -317,8 +303,7 @@ val Isabelle_Version = List(new Regex("""^Isabelle version: (\S+)$""")) } - object AFP_Test - { + object AFP_Test { val log_prefix = "afp-test-devel-" val engine = "afp-test" val Start = new Regex("""^Start test(?: for .+)? at ([^,]+), (.*)$""") @@ -329,8 +314,7 @@ val Bad_Init = new Regex("""^cp:.*: Disc quota exceeded$""") } - object Jenkins - { + object Jenkins { val log_prefix = "jenkins_" val engine = "jenkins" val Host = new Regex("""^Building remotely on (\S+) \((\S+)\).*$""") @@ -348,13 +332,11 @@ val BUILD = "=== BUILD ===" } - private def parse_meta_info(log_file: Log_File): Meta_Info = - { + private def parse_meta_info(log_file: Log_File): Meta_Info = { def parse(engine: String, host: String, start: Date, - End: Regex, Isabelle_Version: List[Regex], AFP_Version: List[Regex]): Meta_Info = - { - val build_id = - { + End: Regex, Isabelle_Version: List[Regex], AFP_Version: List[Regex] + ): Meta_Info = { + val build_id = { val prefix = proper_string(host) orElse proper_string(engine) getOrElse "build" prefix + ":" + start.time.ms } @@ -426,8 +408,7 @@ val SESSION_NAME = "session_name" - object Session_Status extends Enumeration - { + object Session_Status extends Enumeration { val existing, finished, failed, cancelled = Value } @@ -442,29 +423,25 @@ status: Option[Session_Status.Value] = None, errors: List[String] = Nil, theory_timings: Map[String, Timing] = Map.empty, - ml_statistics: List[Properties.T] = Nil) - { + ml_statistics: List[Properties.T] = Nil + ) { def proper_groups: Option[String] = if (groups.isEmpty) None else Some(cat_lines(groups)) def finished: Boolean = status == Some(Session_Status.finished) def failed: Boolean = status == Some(Session_Status.failed) } - object Build_Info - { + object Build_Info { val sessions_dummy: Map[String, Session_Entry] = Map("" -> Session_Entry(theory_timings = Map("" -> Timing.zero))) } - sealed case class Build_Info(sessions: Map[String, Session_Entry]) - { + sealed case class Build_Info(sessions: Map[String, Session_Entry]) { def finished_sessions: List[String] = for ((a, b) <- sessions.toList if b.finished) yield a def failed_sessions: List[String] = for ((a, b) <- sessions.toList if b.failed) yield a } - private def parse_build_info(log_file: Log_File, parse_ml_statistics: Boolean): Build_Info = - { - object Chapter_Name - { + private def parse_build_info(log_file: Log_File, parse_ml_statistics: Boolean): Build_Info = { + object Chapter_Name { def unapply(s: String): Some[(String, String)] = space_explode('/', s) match { case List(chapter, name) => Some((chapter, name)) @@ -484,8 +461,7 @@ val Sources = new Regex("""^Sources (\S+) (\S{""" + SHA1.digest_length + """})$""") val Heap = new Regex("""^Heap (\S+) \((\d+) bytes\)$""") - object Theory_Timing - { + object Theory_Timing { def unapply(line: String): Option[(String, (String, Timing))] = Protocol.Theory_Timing_Marker.unapply(line.replace('~', '-')).map(log_file.parse_props) match { @@ -614,8 +590,8 @@ theory_timings: List[Properties.T], ml_statistics: List[Properties.T], task_statistics: List[Properties.T], - errors: List[String]) - { + errors: List[String] + ) { def error(s: String): Session_Info = copy(errors = errors ::: List(s)) } @@ -625,8 +601,8 @@ command_timings: Boolean, theory_timings: Boolean, ml_statistics: Boolean, - task_statistics: Boolean): Session_Info = - { + task_statistics: Boolean + ): Session_Info = { Session_Info( session_timing = log_file.find_props(Protocol.Session_Timing_Marker) getOrElse Nil, command_timings = @@ -660,8 +636,7 @@ /* SQL data model */ - object Data - { + object Data { def build_log_table(name: String, columns: List[SQL.Column], body: String = ""): SQL.Table = SQL.Table("isabelle_build_log_" + name, columns, body) @@ -712,8 +687,7 @@ /* AFP versions */ - val isabelle_afp_versions_table: SQL.Table = - { + val isabelle_afp_versions_table: SQL.Table = { val version1 = Prop.isabelle_version val version2 = Prop.afp_version build_log_table("isabelle_afp_versions", List(version1.make_primary_key, version2), @@ -728,8 +702,7 @@ if (afp) SQL.Column.date("afp_pull_date") else SQL.Column.date("pull_date") - def pull_date_table(afp: Boolean = false): SQL.Table = - { + def pull_date_table(afp: Boolean = false): SQL.Table = { val (name, versions) = if (afp) ("afp_pull_date", List(Prop.isabelle_version, Prop.afp_version)) else ("pull_date", List(Prop.isabelle_version)) @@ -749,8 +722,10 @@ "now() - INTERVAL '" + days.max(0) + " days'" def recent_pull_date_table( - days: Int, rev: String = "", afp_rev: Option[String] = None): SQL.Table = - { + days: Int, + rev: String = "", + afp_rev: Option[String] = None + ): SQL.Table = { val afp = afp_rev.isDefined val rev2 = afp_rev.getOrElse("") val table = pull_date_table(afp) @@ -769,17 +744,19 @@ else ""))) } - def select_recent_log_names(days: Int): SQL.Source = - { + def select_recent_log_names(days: Int): SQL.Source = { val table1 = meta_info_table val table2 = recent_pull_date_table(days) table1.select(List(log_name), distinct = true) + SQL.join_inner + table2.query_named + " ON " + Prop.isabelle_version(table1) + " = " + Prop.isabelle_version(table2) } - def select_recent_versions(days: Int, - rev: String = "", afp_rev: Option[String] = None, sql: SQL.Source = ""): SQL.Source = - { + def select_recent_versions( + days: Int, + rev: String = "", + afp_rev: Option[String] = None, + sql: SQL.Source = "" + ): SQL.Source = { val afp = afp_rev.isDefined val version = Prop.isabelle_version val table1 = recent_pull_date_table(days, rev = rev, afp_rev = afp_rev) @@ -798,8 +775,7 @@ /* universal view on main data */ - val universal_table: SQL.Table = - { + val universal_table: SQL.Table = { val afp_pull_date = pull_date(afp = true) val version1 = Prop.isabelle_version val version2 = Prop.afp_version @@ -846,8 +822,7 @@ def store(options: Options, cache: XML.Cache = XML.Cache.make()): Store = new Store(options, cache) - class Store private[Build_Log](options: Options, val cache: XML.Cache) - { + class Store private[Build_Log](options: Options, val cache: XML.Cache) { def open_database( user: String = options.string("build_log_database_user"), password: String = options.string("build_log_database_password"), @@ -856,8 +831,8 @@ port: Int = options.int("build_log_database_port"), ssh_host: String = options.string("build_log_ssh_host"), ssh_user: String = options.string("build_log_ssh_user"), - ssh_port: Int = options.int("build_log_ssh_port")): PostgreSQL.Database = - { + ssh_port: Int = options.int("build_log_ssh_port") + ): PostgreSQL.Database = { PostgreSQL.open_database( user = user, password = password, database = database, host = host, port = port, ssh = @@ -867,8 +842,7 @@ } def update_database( - db: PostgreSQL.Database, dirs: List[Path], ml_statistics: Boolean = false): Unit = - { + db: PostgreSQL.Database, dirs: List[Path], ml_statistics: Boolean = false): Unit = { val log_files = dirs.flatMap(dir => File.find_files(dir.file, pred = Log_File.is_log(_), follow_links = true)) @@ -879,14 +853,16 @@ db.create_view(Data.universal_table) } - def snapshot_database(db: PostgreSQL.Database, sqlite_database: Path, - days: Int = 100, ml_statistics: Boolean = false): Unit = - { + def snapshot_database( + db: PostgreSQL.Database, + sqlite_database: Path, + days: Int = 100, + ml_statistics: Boolean = false + ): Unit = { Isabelle_System.make_directory(sqlite_database.dir) sqlite_database.file.delete - using(SQLite.open_database(sqlite_database))(db2 => - { + using(SQLite.open_database(sqlite_database)) { db2 => db.transaction { db2.transaction { // main content @@ -912,16 +888,13 @@ } // pull_date - for (afp <- List(false, true)) - { + for (afp <- List(false, true)) { val afp_rev = if (afp) Some("") else None val table = Data.pull_date_table(afp) db2.create_table(table) - db2.using_statement(table.insert())(stmt2 => - { + db2.using_statement(table.insert()) { stmt2 => db.using_statement( - Data.recent_pull_date_table(days, afp_rev = afp_rev).query)(stmt => - { + Data.recent_pull_date_table(days, afp_rev = afp_rev).query) { stmt => val res = stmt.execute_query() while (res.next()) { for ((c, i) <- table.columns.zipWithIndex) { @@ -929,8 +902,8 @@ } stmt2.execute() } - }) - }) + } + } } // full view @@ -938,18 +911,16 @@ } } db2.rebuild - }) + } } def domain(db: SQL.Database, table: SQL.Table, column: SQL.Column): Set[String] = db.using_statement(table.select(List(column), distinct = true))(stmt => stmt.execute_query().iterator(_.string(column)).toSet) - def update_meta_info(db: SQL.Database, log_name: String, meta_info: Meta_Info): Unit = - { + def update_meta_info(db: SQL.Database, log_name: String, meta_info: Meta_Info): Unit = { val table = Data.meta_info_table - db.using_statement(db.insert_permissive(table))(stmt => - { + db.using_statement(db.insert_permissive(table)) { stmt => stmt.string(1) = log_name for ((c, i) <- table.columns.tail.zipWithIndex) { if (c.T == SQL.Type.Date) @@ -958,14 +929,12 @@ stmt.string(i + 2) = meta_info.get(c) } stmt.execute() - }) + } } - def update_sessions(db: SQL.Database, log_name: String, build_info: Build_Info): Unit = - { + def update_sessions(db: SQL.Database, log_name: String, build_info: Build_Info): Unit = { val table = Data.sessions_table - db.using_statement(db.insert_permissive(table))(stmt => - { + db.using_statement(db.insert_permissive(table)) { stmt => val sessions = if (build_info.sessions.isEmpty) Build_Info.sessions_dummy else build_info.sessions @@ -989,14 +958,12 @@ stmt.string(17) = session.sources stmt.execute() } - }) + } } - def update_theories(db: SQL.Database, log_name: String, build_info: Build_Info): Unit = - { + def update_theories(db: SQL.Database, log_name: String, build_info: Build_Info): Unit = { val table = Data.theories_table - db.using_statement(db.insert_permissive(table))(stmt => - { + db.using_statement(db.insert_permissive(table)) { stmt => val sessions = if (build_info.sessions.forall({ case (_, session) => session.theory_timings.isEmpty })) Build_Info.sessions_dummy @@ -1013,14 +980,12 @@ stmt.long(6) = timing.gc.ms stmt.execute() } - }) + } } - def update_ml_statistics(db: SQL.Database, log_name: String, build_info: Build_Info): Unit = - { + def update_ml_statistics(db: SQL.Database, log_name: String, build_info: Build_Info): Unit = { val table = Data.ml_statistics_table - db.using_statement(db.insert_permissive(table))(stmt => - { + db.using_statement(db.insert_permissive(table)) { stmt => val ml_stats: List[(String, Option[Bytes])] = Par_List.map[(String, Session_Entry), (String, Option[Bytes])]( { case (a, b) => (a, Properties.compress(b.ml_statistics, cache = cache.xz).proper) }, @@ -1032,21 +997,18 @@ stmt.bytes(3) = ml_statistics stmt.execute() } - }) + } } - def write_info(db: SQL.Database, files: List[JFile], ml_statistics: Boolean = false): Unit = - { - abstract class Table_Status(table: SQL.Table) - { + def write_info(db: SQL.Database, files: List[JFile], ml_statistics: Boolean = false): Unit = { + abstract class Table_Status(table: SQL.Table) { db.create_table(table) private var known: Set[String] = domain(db, table, Data.log_name) def required(file: JFile): Boolean = !known(Log_File.plain_name(file.getName)) def update_db(db: SQL.Database, log_file: Log_File): Unit - def update(log_file: Log_File): Unit = - { + def update(log_file: Log_File): Unit = { if (!known(log_file.name)) { update_db(db, log_file) known += log_file.name @@ -1077,19 +1039,16 @@ for (file_group <- files.filter(file => status.exists(_.required(file))). - grouped(options.int("build_log_transaction_size") max 1)) - { + grouped(options.int("build_log_transaction_size") max 1)) { val log_files = Par_List.map[JFile, Log_File](Log_File.apply, file_group) db.transaction { log_files.foreach(log_file => status.foreach(_.update(log_file))) } } } - def read_meta_info(db: SQL.Database, log_name: String): Option[Meta_Info] = - { + def read_meta_info(db: SQL.Database, log_name: String): Option[Meta_Info] = { val table = Data.meta_info_table val columns = table.columns.tail - db.using_statement(table.select(columns, Data.log_name.where_equal(log_name)))(stmt => - { + db.using_statement(table.select(columns, Data.log_name.where_equal(log_name))) { stmt => val res = stmt.execute_query() if (!res.next()) None else { @@ -1104,15 +1063,14 @@ val settings = for ((x, Some(y)) <- results.drop(n)) yield (x, y) Some(Meta_Info(props, settings)) } - }) + } } def read_build_info( db: SQL.Database, log_name: String, session_names: List[String] = Nil, - ml_statistics: Boolean = false): Build_Info = - { + ml_statistics: Boolean = false): Build_Info = { val table1 = Data.sessions_table val table2 = Data.ml_statistics_table @@ -1136,10 +1094,8 @@ else (columns1, table1.ident) val sessions = - db.using_statement(SQL.select(columns) + from + " " + where)(stmt => - { - stmt.execute_query().iterator(res => - { + db.using_statement(SQL.select(columns) + from + " " + where) { stmt => + stmt.execute_query().iterator({ res => val session_name = res.string(Data.session_name) val session_entry = Session_Entry( @@ -1160,7 +1116,7 @@ else Nil) session_name -> session_entry }).toMap - }) + } Build_Info(sessions) } } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/build_minisat.scala --- a/src/Pure/Admin/build_minisat.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/build_minisat.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Build_Minisat -{ +object Build_Minisat { val default_download_url = "https://github.com/stp/minisat/archive/releases/2.2.1.tar.gz" def make_component_name(version: String): String = "minisat-" + version @@ -21,10 +20,9 @@ component_name: String = "", verbose: Boolean = false, progress: Progress = new Progress, - target_dir: Path = Path.current): Unit = - { - Isabelle_System.with_tmp_dir("build")(tmp_dir => - { + target_dir: Path = Path.current + ): Unit = { + Isabelle_System.with_tmp_dir("build") { tmp_dir => /* component */ val Archive_Name = """^.*?([^/]+)$""".r @@ -115,7 +113,7 @@ Makarius """ + Date.Format.date(Date.now()) + "\n") - }) + } } @@ -123,14 +121,13 @@ val isabelle_tool = Isabelle_Tool("build_minisat", "build prover component from sources", Scala_Project.here, - args => - { - var target_dir = Path.current - var download_url = default_download_url - var component_name = "" - var verbose = false + { args => + var target_dir = Path.current + var download_url = default_download_url + var component_name = "" + var verbose = false - val getopts = Getopts(""" + val getopts = Getopts(""" Usage: isabelle build_minisat [OPTIONS] Options are: @@ -142,17 +139,17 @@ Build prover component from official download. """, - "D:" -> (arg => target_dir = Path.explode(arg)), - "U:" -> (arg => download_url = arg), - "n:" -> (arg => component_name = arg), - "v" -> (_ => verbose = true)) + "D:" -> (arg => target_dir = Path.explode(arg)), + "U:" -> (arg => download_url = arg), + "n:" -> (arg => component_name = arg), + "v" -> (_ => verbose = true)) - val more_args = getopts(args) - if (more_args.nonEmpty) getopts.usage() + val more_args = getopts(args) + if (more_args.nonEmpty) getopts.usage() - val progress = new Console_Progress() + val progress = new Console_Progress() - build_minisat(download_url = download_url, component_name = component_name, - verbose = verbose, progress = progress, target_dir = target_dir) - }) + build_minisat(download_url = download_url, component_name = component_name, + verbose = verbose, progress = progress, target_dir = target_dir) + }) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/build_pdfjs.scala --- a/src/Pure/Admin/build_pdfjs.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/build_pdfjs.scala Fri Apr 01 23:26:19 2022 +0200 @@ -13,8 +13,7 @@ package isabelle -object Build_PDFjs -{ +object Build_PDFjs { /* build pdfjs component */ val default_url = "https://github.com/mozilla/pdf.js/releases/download" @@ -24,8 +23,8 @@ base_url: String = default_url, version: String = default_version, target_dir: Path = Path.current, - progress: Progress = new Progress): Unit = - { + progress: Progress = new Progress + ): Unit = { Isabelle_System.require_command("unzip", test = "-h") @@ -39,13 +38,12 @@ /* download */ val download_url = base_url + "/v" + version - Isabelle_System.with_tmp_file("archive", ext = "zip")(archive_file => - { + Isabelle_System.with_tmp_file("archive", ext = "zip") { archive_file => Isabelle_System.download_file(download_url + "/pdfjs-" + version + "-dist.zip", archive_file, progress = progress) Isabelle_System.bash("unzip -x " + File.bash_path(archive_file), cwd = component_dir.file).check - }) + } /* settings */ @@ -74,13 +72,13 @@ val isabelle_tool = Isabelle_Tool("build_pdfjs", "build component for Mozilla PDF.js", - Scala_Project.here, args => - { - var target_dir = Path.current - var base_url = default_url - var version = default_version + Scala_Project.here, + { args => + var target_dir = Path.current + var base_url = default_url + var version = default_version - val getopts = Getopts(""" + val getopts = Getopts(""" Usage: isabelle build_pdfjs [OPTIONS] Options are: @@ -90,16 +88,16 @@ Build component for PDF.js. """, - "D:" -> (arg => target_dir = Path.explode(arg)), - "U:" -> (arg => base_url = arg), - "V:" -> (arg => version = arg)) + "D:" -> (arg => target_dir = Path.explode(arg)), + "U:" -> (arg => base_url = arg), + "V:" -> (arg => version = arg)) - val more_args = getopts(args) - if (more_args.nonEmpty) getopts.usage() + val more_args = getopts(args) + if (more_args.nonEmpty) getopts.usage() - val progress = new Console_Progress() + val progress = new Console_Progress() - build_pdfjs(base_url = base_url, version = version, target_dir = target_dir, - progress = progress) - }) + build_pdfjs(base_url = base_url, version = version, target_dir = target_dir, + progress = progress) + }) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/build_polyml.scala --- a/src/Pure/Admin/build_polyml.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/build_polyml.scala Fri Apr 01 23:26:19 2022 +0200 @@ -10,8 +10,7 @@ import scala.util.matching.Regex -object Build_PolyML -{ +object Build_PolyML { /** platform-specific build **/ sealed case class Platform_Info( @@ -42,8 +41,8 @@ progress: Progress = new Progress, arch_64: Boolean = false, options: List[String] = Nil, - mingw: MinGW = MinGW.none): Unit = - { + mingw: MinGW = MinGW.none + ): Unit = { if (!((root + Path.explode("configure")).is_file && (root + Path.explode("PolyML")).is_dir)) error("Bad Poly/ML root directory: " + root) @@ -65,8 +64,8 @@ def bash( cwd: Path, script: String, redirect: Boolean = false, - echo: Boolean = false): Process_Result = - { + echo: Boolean = false + ): Process_Result = { val script1 = if (platform.is_arm && platform.is_macos) { "arch -arch arm64 bash -c " + Bash.string(script) @@ -146,8 +145,7 @@ /** skeleton for component **/ - private def extract_sources(source_archive: Path, component_dir: Path): Unit = - { + private def extract_sources(source_archive: Path, component_dir: Path): Unit = { if (source_archive.get_ext == "zip") { Isabelle_System.bash( "unzip -x " + File.bash_path(source_archive.absolute), cwd = component_dir.file).check @@ -181,8 +179,8 @@ def build_polyml_component( source_archive: Path, component_dir: Path, - sha1_root: Option[Path] = None): Unit = - { + sha1_root: Option[Path] = None + ): Unit = { Isabelle_System.new_directory(component_dir) extract_sources(source_archive, component_dir) @@ -203,13 +201,13 @@ /** Isabelle tool wrappers **/ val isabelle_tool1 = - Isabelle_Tool("build_polyml", "build Poly/ML from sources", Scala_Project.here, args => - { - var mingw = MinGW.none - var arch_64 = false - var sha1_root: Option[Path] = None + Isabelle_Tool("build_polyml", "build Poly/ML from sources", Scala_Project.here, + { args => + var mingw = MinGW.none + var arch_64 = false + var sha1_root: Option[Path] = None - val getopts = Getopts(""" + val getopts = Getopts(""" Usage: isabelle build_polyml [OPTIONS] ROOT [CONFIGURE_OPTIONS] Options are: @@ -221,32 +219,32 @@ Build Poly/ML in the ROOT directory of its sources, with additional CONFIGURE_OPTIONS (e.g. --without-gmp). """, - "M:" -> (arg => mingw = MinGW(Path.explode(arg))), - "m:" -> - { - case "32" => arch_64 = false - case "64" => arch_64 = true - case bad => error("Bad processor architecture: " + quote(bad)) - }, - "s:" -> (arg => sha1_root = Some(Path.explode(arg)))) + "M:" -> (arg => mingw = MinGW(Path.explode(arg))), + "m:" -> + { + case "32" => arch_64 = false + case "64" => arch_64 = true + case bad => error("Bad processor architecture: " + quote(bad)) + }, + "s:" -> (arg => sha1_root = Some(Path.explode(arg)))) - val more_args = getopts(args) - val (root, options) = - more_args match { - case root :: options => (Path.explode(root), options) - case Nil => getopts.usage() - } - build_polyml(root, sha1_root = sha1_root, progress = new Console_Progress, - arch_64 = arch_64, options = options, mingw = mingw) - }) + val more_args = getopts(args) + val (root, options) = + more_args match { + case root :: options => (Path.explode(root), options) + case Nil => getopts.usage() + } + build_polyml(root, sha1_root = sha1_root, progress = new Console_Progress, + arch_64 = arch_64, options = options, mingw = mingw) + }) val isabelle_tool2 = Isabelle_Tool("build_polyml_component", "make skeleton for Poly/ML component", - Scala_Project.here, args => - { - var sha1_root: Option[Path] = None - - val getopts = Getopts(""" + Scala_Project.here, + { args => + var sha1_root: Option[Path] = None + + val getopts = Getopts(""" Usage: isabelle build_polyml_component [OPTIONS] SOURCE_ARCHIVE COMPONENT_DIR Options are: @@ -255,15 +253,15 @@ Make skeleton for Poly/ML component, based on official source archive (zip or tar.gz). """, - "s:" -> (arg => sha1_root = Some(Path.explode(arg)))) + "s:" -> (arg => sha1_root = Some(Path.explode(arg)))) - val more_args = getopts(args) + val more_args = getopts(args) - val (source_archive, component_dir) = - more_args match { - case List(archive, dir) => (Path.explode(archive), Path.explode(dir)) - case _ => getopts.usage() - } - build_polyml_component(source_archive, component_dir, sha1_root = sha1_root) - }) + val (source_archive, component_dir) = + more_args match { + case List(archive, dir) => (Path.explode(archive), Path.explode(dir)) + case _ => getopts.usage() + } + build_polyml_component(source_archive, component_dir, sha1_root = sha1_root) + }) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/build_release.scala --- a/src/Pure/Admin/build_release.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/build_release.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Build_Release -{ +object Build_Release { /** release context **/ private def execute(dir: Path, script: String): Unit = @@ -20,14 +19,13 @@ private def bash_java_opens(args: String*): String = Bash.strings(args.toList.flatMap(arg => List("--add-opens", arg + "=ALL-UNNAMED"))) - object Release_Context - { + object Release_Context { def apply( target_dir: Path, release_name: String = "", components_base: Path = Components.default_components_base, - progress: Progress = new Progress): Release_Context = - { + progress: Progress = new Progress + ): Release_Context = { val date = Date.now() val dist_name = proper_string(release_name) getOrElse ("Isabelle_" + Date.Format.date(date)) val dist_dir = (target_dir + Path.explode("dist-" + dist_name)).absolute @@ -40,8 +38,8 @@ val dist_name: String, val dist_dir: Path, val components_base: Path, - val progress: Progress) - { + val progress: Progress + ) { override def toString: String = dist_name val isabelle: Path = Path.explode(dist_name) @@ -54,8 +52,7 @@ isabelle_identifier = dist_name + "-build", progress = progress) - def make_announce(id: String): Unit = - { + def make_announce(id: String): Unit = { if (release_name.isEmpty) { File.write(isabelle_dir + Path.explode("ANNOUNCE"), """ @@ -67,8 +64,7 @@ } } - def make_contrib(): Unit = - { + def make_contrib(): Unit = { Isabelle_System.make_directory(Components.contrib(isabelle_dir)) File.write(Components.contrib(isabelle_dir, name = "README"), """This directory contains add-on components that contribute to the main @@ -90,8 +86,8 @@ sealed case class Bundle_Info( platform: Platform.Family.Value, platform_description: String, - name: String) - { + name: String + ) { def path: Path = Path.explode(name) } @@ -104,13 +100,10 @@ val ISABELLE_TAGS: Path = Path.explode("etc/ISABELLE_TAGS") val ISABELLE_IDENTIFIER: Path = Path.explode("etc/ISABELLE_IDENTIFIER") - object Release_Archive - { - def make(bytes: Bytes, rename: String = ""): Release_Archive = - { + object Release_Archive { + def make(bytes: Bytes, rename: String = ""): Release_Archive = { Isabelle_System.with_tmp_dir("tmp")(dir => - Isabelle_System.with_tmp_file("archive", ext = "tar.gz")(archive_path => - { + Isabelle_System.with_tmp_file("archive", ext = "tar.gz") { archive_path => val isabelle_dir = Isabelle_System.make_directory(dir + ISABELLE) Bytes.write(archive_path, bytes) @@ -129,15 +122,18 @@ (Bytes.read(archive_path), rename) } new Release_Archive(bytes1, id, tags, identifier1) - }) + } ) } def read(path: Path, rename: String = ""): Release_Archive = make(Bytes.read(path), rename = rename) - def get(url: String, rename: String = "", progress: Progress = new Progress): Release_Archive = - { + def get( + url: String, + rename: String = "", + progress: Progress = new Progress + ): Release_Archive = { val bytes = if (Path.is_wellformed(url)) Bytes.read(Path.explode(url)) else Isabelle_System.download(url, progress = progress).bytes @@ -146,8 +142,7 @@ } case class Release_Archive private[Build_Release]( - bytes: Bytes, id: String, tags: String, identifier: String) - { + bytes: Bytes, id: String, tags: String, identifier: String) { override def toString: String = identifier } @@ -157,8 +152,7 @@ /* bundled components */ - class Bundled(platform: Option[Platform.Family.Value] = None) - { + class Bundled(platform: Option[Platform.Family.Value] = None) { def detect(s: String): Boolean = s.startsWith("#bundled") && !s.startsWith("#bundled ") @@ -176,8 +170,7 @@ } } - def record_bundled_components(dir: Path): Unit = - { + def record_bundled_components(dir: Path): Unit = { val catalogs = List("main", "bundled").map((_, new Bundled())) ::: Platform.Family.list.flatMap(platform => @@ -195,8 +188,7 @@ } yield bundled(line)).toList)) } - def get_bundled_components(dir: Path, platform: Platform.Family.Value): (List[String], String) = - { + def get_bundled_components(dir: Path, platform: Platform.Family.Value): (List[String], String) = { val Bundled = new Bundled(platform = Some(platform)) val components = for { Bundled(name) <- Components.read_components(dir) } yield name @@ -206,8 +198,7 @@ } def activate_components( - dir: Path, platform: Platform.Family.Value, more_names: List[String]): Unit = - { + dir: Path, platform: Platform.Family.Value, more_names: List[String]): Unit = { def contrib_name(name: String): String = Components.contrib(name = name).implode @@ -231,8 +222,8 @@ options: Options, platform: Platform.Family.Value, build_sessions: List[String], - local_dir: Path): Unit = - { + local_dir: Path + ): Unit = { val server_option = "build_host_" + platform.toString val ssh = options.string(server_option) match { @@ -244,11 +235,9 @@ case s => error("Malformed option " + server_option + ": " + quote(s)) } try { - Isabelle_System.with_tmp_file("tmp", ext = "tar")(local_tmp_tar => - { + Isabelle_System.with_tmp_file("tmp", ext = "tar") { local_tmp_tar => execute_tar(local_dir, "-cf " + File.bash_path(local_tmp_tar) + " .") - ssh.with_tmp_dir(remote_dir => - { + ssh.with_tmp_dir { remote_dir => val remote_tmp_tar = remote_dir + Path.basic("tmp.tar") ssh.write_file(remote_tmp_tar, local_tmp_tar) val remote_commands = @@ -259,9 +248,9 @@ "tar -cf tmp.tar heaps") ssh.execute(remote_commands.mkString(" && "), settings = false).check ssh.read_file(remote_tmp_tar, local_tmp_tar) - }) + } execute_tar(local_dir, "-xf " + File.bash_path(local_tmp_tar)) - }) + } } finally { ssh.close() } } @@ -269,8 +258,7 @@ /* Isabelle application */ - def make_isabelle_options(path: Path, options: List[String], line_ending: String = "\n"): Unit = - { + def make_isabelle_options(path: Path, options: List[String], line_ending: String = "\n"): Unit = { val title = "# Java runtime options" File.write(path, (title :: options).map(_ + line_ending).mkString) } @@ -281,8 +269,7 @@ isabelle_name: String, jdk_component: String, classpath: List[Path], - dock_icon: Boolean = false): Unit = - { + dock_icon: Boolean = false): Unit = { val script = """#!/usr/bin/env bash # # Author: Makarius @@ -330,8 +317,7 @@ } - def make_isabelle_plist(path: Path, isabelle_name: String, isabelle_rev: String): Unit = - { + def make_isabelle_plist(path: Path, isabelle_name: String, isabelle_rev: String): Unit = { File.write(path, """ @@ -394,8 +380,8 @@ def use_release_archive( context: Release_Context, archive: Release_Archive, - id: String = ""): Unit = - { + id: String = "" + ): Unit = { if (id.nonEmpty && id != archive.id) { error("Mismatch of release identification " + id + " vs. archive " + archive.id) } @@ -408,8 +394,8 @@ def build_release_archive( context: Release_Context, version: String, - parallel_jobs: Int = 1): Unit = - { + parallel_jobs: Int = 1 + ): Unit = { val progress = context.progress val hg = Mercurial.repository(Path.ISABELLE_HOME) @@ -497,8 +483,8 @@ website: Option[Path] = None, build_sessions: List[String] = Nil, build_library: Boolean = false, - parallel_jobs: Int = 1): Unit = - { + parallel_jobs: Int = 1 + ): Unit = { val progress = context.progress @@ -510,15 +496,14 @@ Isabelle_System.rm_tree(context.dist_dir + path) } - Isabelle_System.with_tmp_file("archive", ext = "tar.gz")(archive_path => - { + Isabelle_System.with_tmp_file("archive", ext = "tar.gz") { archive_path => Bytes.write(archive_path, archive.bytes) val extract = List("README", "NEWS", "ANNOUNCE", "COPYRIGHT", "CONTRIBUTORS", "doc"). map(name => context.dist_name + "/" + name) execute_tar(context.dist_dir, "-xzf " + File.bash_path(archive_path) + " " + Bash.strings(extract)) - }) + } Isabelle_System.symlink(Path.explode(context.dist_name), context.dist_dir + ISABELLE) @@ -533,8 +518,7 @@ progress.echo("\nApplication bundle for " + platform) - Isabelle_System.with_tmp_dir("build_release")(tmp_dir => - { + Isabelle_System.with_tmp_dir("build_release") { tmp_dir => // release archive execute_tar(tmp_dir, "-xzf " + File.bash_path(context.isabelle_archive)) @@ -578,19 +562,17 @@ if (opt.startsWith(s)) s + isabelle_name else opt }) ::: List("-Disabelle.jedit_server=" + isabelle_name) - val classpath: List[Path] = - { + val classpath: List[Path] = { val base = isabelle_target.absolute val classpath1 = Path.split(other_isabelle.getenv("ISABELLE_CLASSPATH")) val classpath2 = Path.split(other_isabelle.getenv("ISABELLE_SETUP_CLASSPATH")) - (classpath1 ::: classpath2).map(path => - { + (classpath1 ::: classpath2).map { path => val abs_path = path.absolute File.relative_path(base, abs_path) match { case Some(rel_path) => rel_path case None => error("Bad classpath element: " + abs_path) } - }) + } } val jedit_options = Path.explode("src/Tools/jEdit/etc/options") @@ -784,7 +766,7 @@ Bytes.read(sfx_exe) + Bytes(sfx_txt) + Bytes.read(exe_archive)) File.set_executable(context.dist_dir + isabelle_exe, true) } - }) + } progress.echo("DONE") } @@ -833,27 +815,26 @@ progress.echo_warning("Library archive already exists: " + context.isabelle_library_archive) } else { - Isabelle_System.with_tmp_dir("build_release")(tmp_dir => - { - val bundle = - context.dist_dir + Path.explode(context.dist_name + "_" + Platform.family + ".tar.gz") - execute_tar(tmp_dir, "-xzf " + File.bash_path(bundle)) + Isabelle_System.with_tmp_dir("build_release") { tmp_dir => + val bundle = + context.dist_dir + Path.explode(context.dist_name + "_" + Platform.family + ".tar.gz") + execute_tar(tmp_dir, "-xzf " + File.bash_path(bundle)) - val other_isabelle = context.other_isabelle(tmp_dir) + val other_isabelle = context.other_isabelle(tmp_dir) - Isabelle_System.make_directory(other_isabelle.etc) - File.write(other_isabelle.etc_settings, "ML_OPTIONS=\"--minheap 1000 --maxheap 4000\"\n") + Isabelle_System.make_directory(other_isabelle.etc) + File.write(other_isabelle.etc_settings, "ML_OPTIONS=\"--minheap 1000 --maxheap 4000\"\n") - other_isabelle.bash("bin/isabelle build -f -j " + parallel_jobs + - " -o browser_info -o document=pdf -o document_variants=document:outline=/proof,/ML" + - " -o system_heaps -c -a -d '~~/src/Benchmarks'", echo = true).check - other_isabelle.isabelle_home_user.file.delete + other_isabelle.bash("bin/isabelle build -f -j " + parallel_jobs + + " -o browser_info -o document=pdf -o document_variants=document:outline=/proof,/ML" + + " -o system_heaps -c -a -d '~~/src/Benchmarks'", echo = true).check + other_isabelle.isabelle_home_user.file.delete - execute(tmp_dir, "chmod -R a+r " + Bash.string(context.dist_name)) - execute(tmp_dir, "chmod -R g=o " + Bash.string(context.dist_name)) - execute_tar(tmp_dir, "-czf " + File.bash_path(context.isabelle_library_archive) + - " " + Bash.string(context.dist_name + "/browser_info")) - }) + execute(tmp_dir, "chmod -R a+r " + Bash.string(context.dist_name)) + execute(tmp_dir, "chmod -R g=o " + Bash.string(context.dist_name)) + execute_tar(tmp_dir, "-czf " + File.bash_path(context.isabelle_library_archive) + + " " + Bash.string(context.dist_name + "/browser_info")) + } } } } @@ -862,8 +843,7 @@ /** command line entry point **/ - def main(args: Array[String]): Unit = - { + def main(args: Array[String]): Unit = { Command_Line.tool { var afp_rev = "" var components_base: Path = Components.default_components_base diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/build_scala.scala --- a/src/Pure/Admin/build_scala.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/build_scala.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Build_Scala -{ +object Build_Scala { /* downloads */ sealed case class Download( @@ -16,8 +15,8 @@ version: String, url: String, physical_url: String = "", - base_version: String = "3") - { + base_version: String = "3" + ) { def make_url(template: String): String = template.replace("{V}", version).replace("{B}", base_version) @@ -30,13 +29,12 @@ Isabelle_System.download_file(proper_url, path, progress = progress) def get_unpacked(dir: Path, strip: Int = 0, progress: Progress = new Progress): Unit = - Isabelle_System.with_tmp_file("archive")(archive_path => - { + Isabelle_System.with_tmp_file("archive"){ archive_path => get(archive_path, progress = progress) progress.echo("Unpacking " + artifact) - Isabelle_System.gnutar("-xzf " + File.bash_path(archive_path), - dir = dir, strip = strip).check - }) + Isabelle_System.gnutar( + "-xzf " + File.bash_path(archive_path), dir = dir, strip = strip).check + } def print: String = " * " + name + " " + version + @@ -68,8 +66,8 @@ def build_scala( target_dir: Path = Path.current, - progress: Progress = new Progress): Unit = - { + progress: Progress = new Progress + ): Unit = { /* component */ val component_name = main_download.name + "-" + main_download.version @@ -91,8 +89,7 @@ /* classpath */ - val classpath: List[String] = - { + val classpath: List[String] = { def no_function(name: String): String = "function " + name + "() {\n:\n}" val script = cat_lines(List( @@ -142,11 +139,11 @@ val isabelle_tool = Isabelle_Tool("build_scala", "build Isabelle Scala component from official downloads", - Scala_Project.here, args => - { - var target_dir = Path.current + Scala_Project.here, + { args => + var target_dir = Path.current - val getopts = Getopts(""" + val getopts = Getopts(""" Usage: isabelle build_scala [OPTIONS] Options are: @@ -154,13 +151,13 @@ Build Isabelle Scala component from official downloads. """, - "D:" -> (arg => target_dir = Path.explode(arg))) + "D:" -> (arg => target_dir = Path.explode(arg))) - val more_args = getopts(args) - if (more_args.nonEmpty) getopts.usage() + val more_args = getopts(args) + if (more_args.nonEmpty) getopts.usage() - val progress = new Console_Progress() + val progress = new Console_Progress() - build_scala(target_dir = target_dir, progress = progress) - }) + build_scala(target_dir = target_dir, progress = progress) + }) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/build_spass.scala --- a/src/Pure/Admin/build_spass.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/build_spass.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Build_SPASS -{ +object Build_SPASS { /* build SPASS */ val default_download_url = "https://www.cs.vu.nl/~jbe248/spass-3.8ds-src.tar.gz" @@ -18,10 +17,8 @@ download_url: String = default_download_url, verbose: Boolean = false, progress: Progress = new Progress, - target_dir: Path = Path.current): Unit = - { - Isabelle_System.with_tmp_dir("build")(tmp_dir => - { + target_dir: Path = Path.current): Unit = { + Isabelle_System.with_tmp_dir("build") { tmp_dir => Isabelle_System.require_command("bison") Isabelle_System.require_command("flex") @@ -143,20 +140,20 @@ Makarius """ + Date.Format.date(Date.now()) + "\n") - }) + } } /* Isabelle tool wrapper */ val isabelle_tool = Isabelle_Tool("build_spass", "build prover component from source distribution", - Scala_Project.here, args => - { - var target_dir = Path.current - var download_url = default_download_url - var verbose = false + Scala_Project.here, + { args => + var target_dir = Path.current + var download_url = default_download_url + var verbose = false - val getopts = Getopts(""" + val getopts = Getopts(""" Usage: isabelle build_spass [OPTIONS] Options are: @@ -167,16 +164,16 @@ Build prover component from the specified source distribution. """, - "D:" -> (arg => target_dir = Path.explode(arg)), - "U:" -> (arg => download_url = arg), - "v" -> (_ => verbose = true)) + "D:" -> (arg => target_dir = Path.explode(arg)), + "U:" -> (arg => download_url = arg), + "v" -> (_ => verbose = true)) - val more_args = getopts(args) - if (more_args.nonEmpty) getopts.usage() + val more_args = getopts(args) + if (more_args.nonEmpty) getopts.usage() - val progress = new Console_Progress() + val progress = new Console_Progress() - build_spass(download_url = download_url, verbose = verbose, progress = progress, - target_dir = target_dir) - }) + build_spass(download_url = download_url, verbose = verbose, progress = progress, + target_dir = target_dir) + }) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/build_sqlite.scala --- a/src/Pure/Admin/build_sqlite.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/build_sqlite.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,15 +7,14 @@ package isabelle -object Build_SQLite -{ +object Build_SQLite { /* build sqlite */ def build_sqlite( download_url: String, progress: Progress = new Progress, - target_dir: Path = Path.current): Unit = - { + target_dir: Path = Path.current + ): Unit = { val Download_Name = """^.*/([^/]+)\.jar""".r val download_name = download_url match { @@ -55,8 +54,7 @@ val jar = component_dir + Path.basic(download_name).ext("jar") Isabelle_System.download_file(download_url, jar, progress = progress) - Isabelle_System.with_tmp_dir("sqlite")(jar_dir => - { + Isabelle_System.with_tmp_dir("sqlite") { jar_dir => progress.echo("Unpacking " + jar) Isabelle_System.bash("isabelle_jdk jar xf " + File.bash_path(jar.absolute), cwd = jar_dir.file).check @@ -77,7 +75,7 @@ } File.set_executable(component_dir + Path.explode("x86_64-windows/sqlitejdbc.dll"), true) - }) + } } @@ -85,11 +83,11 @@ val isabelle_tool = Isabelle_Tool("build_sqlite", "build Isabelle sqlite-jdbc component from official download", - Scala_Project.here, args => - { - var target_dir = Path.current + Scala_Project.here, + { args => + var target_dir = Path.current - val getopts = Getopts(""" + val getopts = Getopts(""" Usage: isabelle build_sqlite [OPTIONS] DOWNLOAD Options are: @@ -99,17 +97,17 @@ https://github.com/xerial/sqlite-jdbc and https://oss.sonatype.org/content/repositories/releases/org/xerial/sqlite-jdbc """, - "D:" -> (arg => target_dir = Path.explode(arg))) + "D:" -> (arg => target_dir = Path.explode(arg))) - val more_args = getopts(args) - val download_url = - more_args match { - case List(download_url) => download_url - case _ => getopts.usage() - } + val more_args = getopts(args) + val download_url = + more_args match { + case List(download_url) => download_url + case _ => getopts.usage() + } - val progress = new Console_Progress() + val progress = new Console_Progress() - build_sqlite(download_url, progress = progress, target_dir = target_dir) - }) + build_sqlite(download_url, progress = progress, target_dir = target_dir) + }) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/build_status.scala --- a/src/Pure/Admin/build_status.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/build_status.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Build_Status -{ +object Build_Status { /* defaults */ val default_target_dir = Path.explode("build_status") @@ -22,15 +21,22 @@ /* data profiles */ sealed case class Profile( - description: String, history: Int = 0, afp: Boolean = false, bulky: Boolean = false, sql: String) - { + description: String, + history: Int = 0, + afp: Boolean = false, + bulky: Boolean = false, + sql: String + ) { def days(options: Options): Int = options.int("build_log_history") max history def stretch(options: Options): Double = (days(options) max default_history min (default_history * 5)).toDouble / default_history - def select(options: Options, columns: List[SQL.Column], only_sessions: Set[String]): SQL.Source = - { + def select( + options: Options, + columns: List[SQL.Column], + only_sessions: Set[String] + ): SQL.Source = { Build_Log.Data.universal_table.select(columns, distinct = true, sql = "WHERE " + Build_Log.Data.pull_date(afp) + " > " + Build_Log.Data.recent_time(days(options)) + @@ -55,8 +61,8 @@ verbose: Boolean = false, target_dir: Path = default_target_dir, ml_statistics: Boolean = false, - image_size: (Int, Int) = default_image_size): Unit = - { + image_size: (Int, Int) = default_image_size + ): Unit = { val ml_statistics_domain = Iterator(ML_Statistics.heap_fields, ML_Statistics.program_fields, ML_Statistics.tasks_fields, ML_Statistics.workers_fields).flatMap(_._2).toSet @@ -74,15 +80,21 @@ sealed case class Data(date: Date, entries: List[Data_Entry]) sealed case class Data_Entry( - name: String, hosts: List[String], stretch: Double, sessions: List[Session]) - { + name: String, + hosts: List[String], + stretch: Double, + sessions: List[Session] + ) { def failed_sessions: List[Session] = sessions.filter(_.head.failed).sortBy(_.name) } sealed case class Session( - name: String, threads: Int, entries: List[Entry], - ml_statistics: ML_Statistics, ml_statistics_date: Long) - { + name: String, + threads: Int, + entries: List[Entry], + ml_statistics: ML_Statistics, + ml_statistics_date: Long + ) { require(entries.nonEmpty, "no entries") lazy val sorted_entries: List[Entry] = entries.sortBy(entry => - entry.date) @@ -101,8 +113,7 @@ entry.average_heap > 0 || entry.stored_heap > 0) - def make_csv: CSV.File = - { + def make_csv: CSV.File = { val header = List("session_name", "chapter", @@ -167,15 +178,14 @@ average_heap: Long, stored_heap: Long, status: Build_Log.Session_Status.Value, - errors: List[String]) - { + errors: List[String] + ) { val date: Long = (afp_pull_date getOrElse pull_date).unix_epoch def finished: Boolean = status == Build_Log.Session_Status.finished def failed: Boolean = status == Build_Log.Session_Status.failed - def present_errors(name: String): XML.Body = - { + def present_errors(name: String): XML.Body = { if (errors.isEmpty) HTML.text(name + print_version(isabelle_version, afp_version, chapter)) else { @@ -185,14 +195,15 @@ } } - sealed case class Image(name: String, width: Int, height: Int) - { + sealed case class Image(name: String, width: Int, height: Int) { def path: Path = Path.basic(name) } def print_version( - isabelle_version: String, afp_version: String = "", chapter: String = AFP.chapter): String = - { + isabelle_version: String, + afp_version: String = "", + chapter: String = AFP.chapter + ): String = { val body = proper_string(isabelle_version).map("Isabelle/" + _).toList ::: (if (chapter == AFP.chapter) proper_string(afp_version).map("AFP/" + _) else None).toList @@ -205,8 +216,8 @@ only_sessions: Set[String] = Set.empty, ml_statistics: Boolean = false, ml_statistics_domain: String => Boolean = _ => true, - verbose: Boolean = false): Data = - { + verbose: Boolean = false + ): Data = { val date = Date.now() var data_hosts = Map.empty[String, Set[String]] var data_stretch = Map.empty[String, Double] @@ -217,8 +228,7 @@ val store = Build_Log.store(options) - using(store.open_database())(db => - { + using(store.open_database()) { db => for (profile <- profiles.sortBy(_.description)) { progress.echo("input " + quote(profile.description)) @@ -252,15 +262,13 @@ val sql = profile.select(options, columns, only_sessions) progress.echo_if(verbose, sql) - db.using_statement(sql)(stmt => - { + db.using_statement(sql) { stmt => val res = stmt.execute_query() while (res.next()) { val session_name = res.string(Build_Log.Data.session_name) val chapter = res.string(Build_Log.Data.chapter) val groups = split_lines(res.string(Build_Log.Data.groups)) - val threads = - { + val threads = { val threads1 = res.string(Build_Log.Settings.ISABELLE_BUILD_OPTIONS) match { case Threads_Option(Value.Int(i)) => i @@ -341,9 +349,9 @@ data_entries += (data_name -> (sessions + (session_name -> session))) } } - }) + } } - }) + } val sorted_entries = (for { @@ -365,8 +373,8 @@ def present_data(data: Data, progress: Progress = new Progress, target_dir: Path = default_target_dir, - image_size: (Int, Int) = default_image_size): Unit = - { + image_size: (Int, Int) = default_image_size + ): Unit = { def clean_name(name: String): String = name.flatMap(c => if (c == ' ' || c == '/') "_" else if (c == ',') "" else c.toString) @@ -440,8 +448,7 @@ } max 0.1) * 1.1 val timing_range = "[0:" + max_time + "]" - def gnuplot(plot_name: String, plots: List[String], range: String): Image = - { + def gnuplot(plot_name: String, plots: List[String], range: String): Image = { val image = Image(plot_name, image_width_stretch, image_height) File.write(gnuplot_file, """ @@ -463,8 +470,7 @@ image } - val timing_plots = - { + val timing_plots = { val plots1 = List( """ using 1:2 smooth sbezier title "elapsed time (smooth)" """, @@ -492,8 +498,7 @@ """ using 1:12 smooth sbezier title "heap stored (smooth)" """, """ using 1:12 smooth csplines title "heap stored" """) - def jfreechart(plot_name: String, fields: ML_Statistics.Fields): Image = - { + def jfreechart(plot_name: String, fields: ML_Statistics.Fields): Image = { val image = Image(plot_name, image_width, image_height) val chart = session.ml_statistics.chart( @@ -577,16 +582,16 @@ val isabelle_tool = Isabelle_Tool("build_status", "present recent build status information from database", - Scala_Project.here, args => - { - var target_dir = default_target_dir - var ml_statistics = false - var only_sessions = Set.empty[String] - var options = Options.init() - var image_size = default_image_size - var verbose = false + Scala_Project.here, + { args => + var target_dir = default_target_dir + var ml_statistics = false + var only_sessions = Set.empty[String] + var options = Options.init() + var image_size = default_image_size + var verbose = false - val getopts = Getopts(""" + val getopts = Getopts(""" Usage: isabelle build_status [OPTIONS] Options are: @@ -602,24 +607,24 @@ via system options build_log_database_host, build_log_database_user, build_log_history etc. """, - "D:" -> (arg => target_dir = Path.explode(arg)), - "M" -> (_ => ml_statistics = true), - "S:" -> (arg => only_sessions = space_explode(',', arg).toSet), - "l:" -> (arg => options = options + ("build_log_history=" + arg)), - "o:" -> (arg => options = options + arg), - "s:" -> (arg => - space_explode('x', arg).map(Value.Int.parse) match { - case List(w, h) if w > 0 && h > 0 => image_size = (w, h) - case _ => error("Error bad PNG image size: " + quote(arg)) - }), - "v" -> (_ => verbose = true)) + "D:" -> (arg => target_dir = Path.explode(arg)), + "M" -> (_ => ml_statistics = true), + "S:" -> (arg => only_sessions = space_explode(',', arg).toSet), + "l:" -> (arg => options = options + ("build_log_history=" + arg)), + "o:" -> (arg => options = options + arg), + "s:" -> (arg => + space_explode('x', arg).map(Value.Int.parse) match { + case List(w, h) if w > 0 && h > 0 => image_size = (w, h) + case _ => error("Error bad PNG image size: " + quote(arg)) + }), + "v" -> (_ => verbose = true)) - val more_args = getopts(args) - if (more_args.nonEmpty) getopts.usage() + val more_args = getopts(args) + if (more_args.nonEmpty) getopts.usage() - val progress = new Console_Progress + val progress = new Console_Progress - build_status(options, progress = progress, only_sessions = only_sessions, verbose = verbose, - target_dir = target_dir, ml_statistics = ml_statistics, image_size = image_size) - }) + build_status(options, progress = progress, only_sessions = only_sessions, verbose = verbose, + target_dir = target_dir, ml_statistics = ml_statistics, image_size = image_size) + }) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/build_vampire.scala --- a/src/Pure/Admin/build_vampire.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/build_vampire.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Build_Vampire -{ +object Build_Vampire { val default_download_url = "https://github.com/vprover/vampire/archive/refs/tags/v4.6.tar.gz" val default_jobs = 1 @@ -24,12 +23,11 @@ component_name: String = "", verbose: Boolean = false, progress: Progress = new Progress, - target_dir: Path = Path.current): Unit = - { + target_dir: Path = Path.current + ): Unit = { Isabelle_System.require_command("cmake") - Isabelle_System.with_tmp_dir("build")(tmp_dir => - { + Isabelle_System.with_tmp_dir("build") { tmp_dir => /* component */ val Archive_Name = """^.*?([^/]+)$""".r @@ -120,7 +118,7 @@ Makarius """ + Date.Format.date(Date.now()) + "\n") - }) + } } @@ -128,8 +126,8 @@ val isabelle_tool = Isabelle_Tool("build_vampire", "build prover component from official download", - Scala_Project.here, args => - { + Scala_Project.here, + { args => var target_dir = Path.current var download_url = default_download_url var jobs = default_jobs diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/build_verit.scala --- a/src/Pure/Admin/build_verit.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/build_verit.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Build_VeriT -{ +object Build_VeriT { val default_download_url = "https://verit.loria.fr/rmx/2021.06.2/verit-2021.06.2-rmx.tar.gz" @@ -19,12 +18,11 @@ verbose: Boolean = false, progress: Progress = new Progress, target_dir: Path = Path.current, - mingw: MinGW = MinGW.none): Unit = - { + mingw: MinGW = MinGW.none + ): Unit = { mingw.check - Isabelle_System.with_tmp_dir("build")(tmp_dir => - { + Isabelle_System.with_tmp_dir("build") { tmp_dir => /* component */ val Archive_Name = """^.*?([^/]+)$""".r @@ -116,7 +114,7 @@ Makarius """ + Date.Format.date(Date.now()) + "\n") - }) + } } @@ -124,14 +122,14 @@ val isabelle_tool = Isabelle_Tool("build_verit", "build prover component from official download", - Scala_Project.here, args => - { - var target_dir = Path.current - var mingw = MinGW.none - var download_url = default_download_url - var verbose = false + Scala_Project.here, + { args => + var target_dir = Path.current + var mingw = MinGW.none + var download_url = default_download_url + var verbose = false - val getopts = Getopts(""" + val getopts = Getopts(""" Usage: isabelle build_verit [OPTIONS] Options are: @@ -143,17 +141,17 @@ Build prover component from official download. """, - "D:" -> (arg => target_dir = Path.explode(arg)), - "M:" -> (arg => mingw = MinGW(Path.explode(arg))), - "U:" -> (arg => download_url = arg), - "v" -> (_ => verbose = true)) + "D:" -> (arg => target_dir = Path.explode(arg)), + "M:" -> (arg => mingw = MinGW(Path.explode(arg))), + "U:" -> (arg => download_url = arg), + "v" -> (_ => verbose = true)) - val more_args = getopts(args) - if (more_args.nonEmpty) getopts.usage() + val more_args = getopts(args) + if (more_args.nonEmpty) getopts.usage() - val progress = new Console_Progress() + val progress = new Console_Progress() - build_verit(download_url = download_url, verbose = verbose, progress = progress, - target_dir = target_dir, mingw = mingw) - }) + build_verit(download_url = download_url, verbose = verbose, progress = progress, + target_dir = target_dir, mingw = mingw) + }) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/build_zipperposition.scala --- a/src/Pure/Admin/build_zipperposition.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/build_zipperposition.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Build_Zipperposition -{ +object Build_Zipperposition { val default_version = "2.1" @@ -18,10 +17,9 @@ version: String = default_version, verbose: Boolean = false, progress: Progress = new Progress, - target_dir: Path = Path.current): Unit = - { - Isabelle_System.with_tmp_dir("build")(build_dir => - { + target_dir: Path = Path.current + ): Unit = { + Isabelle_System.with_tmp_dir("build") { build_dir => if (Platform.is_linux) Isabelle_System.require_command("patchelf") @@ -85,7 +83,7 @@ Makarius """ + Date.Format.date(Date.now()) + "\n") - }) + } } @@ -93,13 +91,13 @@ val isabelle_tool = Isabelle_Tool("build_zipperposition", "build prover component from OPAM repository", - Scala_Project.here, args => - { - var target_dir = Path.current - var version = default_version - var verbose = false + Scala_Project.here, + { args => + var target_dir = Path.current + var version = default_version + var verbose = false - val getopts = Getopts(""" + val getopts = Getopts(""" Usage: isabelle build_zipperposition [OPTIONS] Options are: @@ -109,16 +107,16 @@ Build prover component from OPAM repository. """, - "D:" -> (arg => target_dir = Path.explode(arg)), - "V:" -> (arg => version = arg), - "v" -> (_ => verbose = true)) + "D:" -> (arg => target_dir = Path.explode(arg)), + "V:" -> (arg => version = arg), + "v" -> (_ => verbose = true)) - val more_args = getopts(args) - if (more_args.nonEmpty) getopts.usage() + val more_args = getopts(args) + if (more_args.nonEmpty) getopts.usage() - val progress = new Console_Progress() + val progress = new Console_Progress() - build_zipperposition(version = version, verbose = verbose, progress = progress, - target_dir = target_dir) - }) + build_zipperposition(version = version, verbose = verbose, progress = progress, + target_dir = target_dir) + }) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/check_sources.scala --- a/src/Pure/Admin/check_sources.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/check_sources.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,10 +7,8 @@ package isabelle -object Check_Sources -{ - def check_file(path: Path): Unit = - { +object Check_Sources { + def check_file(path: Path): Unit = { val file_name = path.implode val file_pos = path.position def line_pos(i: Int) = Position.Line_File(i + 1, file_name) @@ -25,13 +23,11 @@ Output.warning("Bad UTF8 encoding" + Position.here(file_pos)) } - for { (line, i) <- split_lines(content).iterator.zipWithIndex } - { + for { (line, i) <- split_lines(content).iterator.zipWithIndex } { try { Symbol.decode_strict(line) - for { c <- Codepoint.iterator(line); if c > 128 && !Character.isAlphabetic(c) } - { + for { c <- Codepoint.iterator(line); if c > 128 && !Character.isAlphabetic(c) } { Output.warning("Suspicious Unicode character " + quote(Codepoint.string(c)) + Position.here(line_pos(i))) } @@ -49,8 +45,7 @@ Output.warning("Bidirectional Unicode text" + Position.here(file_pos)) } - def check_hg(root: Path): Unit = - { + def check_hg(root: Path): Unit = { Output.writeln("Checking " + root + " ...") val hg = Mercurial.repository(root) for { @@ -64,17 +59,17 @@ val isabelle_tool = Isabelle_Tool("check_sources", "some sanity checks for Isabelle sources", - Scala_Project.here, args => - { - val getopts = Getopts(""" + Scala_Project.here, + { args => + val getopts = Getopts(""" Usage: isabelle check_sources [ROOT_DIRS...] Check .thy, .ML, ROOT against known files of Mercurial ROOT_DIRS. """) - val specs = getopts(args) - if (specs.isEmpty) getopts.usage() + val specs = getopts(args) + if (specs.isEmpty) getopts.usage() - for (root <- specs) check_hg(Path.explode(root)) - }) + for (root <- specs) check_hg(Path.explode(root)) + }) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/ci_profile.scala --- a/src/Pure/Admin/ci_profile.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/ci_profile.scala Fri Apr 01 23:26:19 2022 +0200 @@ -12,17 +12,14 @@ import java.util.{Properties => JProperties, Map => JMap} -abstract class CI_Profile extends Isabelle_Tool.Body -{ +abstract class CI_Profile extends Isabelle_Tool.Body { case class Result(rc: Int) - case object Result - { + 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) = - { + 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 { @@ -41,13 +38,11 @@ (results, end_time - start_time) } - private def load_properties(): JProperties = - { + private def load_properties(): JProperties = { val props = new JProperties() val file_name = Isabelle_System.getenv("ISABELLE_CI_PROPERTIES") - if (file_name != "") - { + if (file_name != "") { val file = Path.explode(file_name).file if (file.exists()) props.load(new java.io.FileReader(file)) @@ -57,8 +52,7 @@ props } - private def compute_timing(results: Build.Results, group: Option[String]): Timing = - { + 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 @@ -66,8 +60,7 @@ timings.foldLeft(Timing.zero)(_ + _) } - private def with_documents(options: Options): Options = - { + private def with_documents(options: Options): Options = { if (documents) options .bool.update("browser_info", true) @@ -90,8 +83,7 @@ final val start_time = Instant.now().atZone(ZoneId.systemDefault).format(DateTimeFormatter.RFC_1123_DATE_TIME) - override final def apply(args: List[String]): Unit = - { + override final def apply(args: List[String]): Unit = { print_section("CONFIGURATION") println(Build_Log.Settings.show()) val props = load_properties() diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/isabelle_cronjob.scala --- a/src/Pure/Admin/isabelle_cronjob.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/isabelle_cronjob.scala Fri Apr 01 23:26:19 2022 +0200 @@ -12,8 +12,7 @@ import scala.annotation.tailrec -object Isabelle_Cronjob -{ +object Isabelle_Cronjob { /* global resources: owned by main cronjob */ val backup = "lxbroy10:cronjob" @@ -43,8 +42,8 @@ def get_afp_rev(): String = Mercurial.repository(afp_repos).id() val init: Logger_Task = - Logger_Task("init", logger => - { + Logger_Task("init", + { logger => Isabelle_Devel.make_index() Mercurial.setup_repository(Isabelle_System.isabelle_repository.root, isabelle_repos) @@ -62,8 +61,8 @@ }) val exit: Logger_Task = - Logger_Task("exit", logger => - { + Logger_Task("exit", + { logger => Isabelle_System.bash( "rsync -a " + File.bash_path(main_dir) + "/log/." + " " + Bash.string(backup) + "/log/.") .check @@ -73,8 +72,8 @@ /* Mailman archives */ val mailman_archives: Logger_Task = - Logger_Task("mailman_archives", logger => - { + Logger_Task("mailman_archives", + { logger => Mailman.Isabelle_Users.download(mailman_archives_dir) Mailman.Isabelle_Dev.download(mailman_archives_dir) }) @@ -83,10 +82,8 @@ /* build release */ val build_release: Logger_Task = - Logger_Task("build_release", logger => - { - Isabelle_Devel.release_snapshot(logger.options, get_rev(), get_afp_rev()) - }) + Logger_Task("build_release", + { logger => Isabelle_Devel.release_snapshot(logger.options, get_rev(), get_afp_rev()) }) /* remote build_history */ @@ -95,8 +92,8 @@ known: Boolean, isabelle_version: String, afp_version: Option[String], - pull_date: Date) - { + pull_date: Date + ) { def unknown: Boolean = !known def versions: (String, Option[String]) = (isabelle_version, afp_version) @@ -105,17 +102,20 @@ (afp_rev.isEmpty || afp_rev.get != "" && afp_version == afp_rev) } - def recent_items(db: SQL.Database, - days: Int, rev: String, afp_rev: Option[String], sql: SQL.Source): List[Item] = - { + def recent_items( + db: SQL.Database, + days: Int, + rev: String, + afp_rev: Option[String], + sql: SQL.Source + ): List[Item] = { val afp = afp_rev.isDefined val select = Build_Log.Data.select_recent_versions( days = days, rev = rev, afp_rev = afp_rev, sql = "WHERE " + sql) db.using_statement(select)(stmt => - stmt.execute_query().iterator(res => - { + stmt.execute_query().iterator({ res => val known = res.bool(Build_Log.Data.known) val isabelle_version = res.string(Build_Log.Prop.isabelle_version) val afp_version = if (afp) proper_string(res.string(Build_Log.Prop.afp_version)) else None @@ -124,8 +124,7 @@ }).toList) } - def unknown_runs(items: List[Item]): List[List[Item]] = - { + 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 } @@ -150,8 +149,8 @@ bulky: Boolean = false, more_hosts: List[String] = Nil, detect: SQL.Source = "", - active: Boolean = true) - { + active: Boolean = true + ) { def ssh_session(context: SSH.Context): SSH.Session = context.open_session(host = host, user = user, port = port, actual_host = actual_host, proxy_host = proxy_host, proxy_user = proxy_user, proxy_port = proxy_port, @@ -168,15 +167,13 @@ def pick( options: Options, rev: String = "", - filter: Item => Boolean = _ => true): Option[(String, Option[String])] = - { + filter: Item => Boolean = _ => true + ): Option[(String, Option[String])] = { val afp_rev = if (afp) Some(get_afp_rev()) else None val store = Build_Log.store(options) - using(store.open_database())(db => - { - def pick_days(days: Int, gap: Int): Option[(String, Option[String])] = - { + using(store.open_database()) { db => + def pick_days(days: Int, gap: Int): Option[(String, Option[String])] = { val items = recent_items(db, days, rev, afp_rev, sql).filter(filter) def runs = unknown_runs(items).filter(run => run.length >= gap) @@ -195,7 +192,7 @@ pick_days(options.int("build_log_history") max history, 2) orElse pick_days(200, 5) orElse pick_days(2000, 1) - }) + } } def build_history_options: String = @@ -310,8 +307,7 @@ } } - val remote_builds1: List[List[Remote_Build]] = - { + val remote_builds1: List[List[Remote_Build]] = { List( List(Remote_Build("Linux A", "augsburg1", options = "-m32 -B -M4" + @@ -400,34 +396,36 @@ bulky = true, detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("AFP")))) - def remote_build_history(rev: String, afp_rev: Option[String], i: Int, r: Remote_Build) - : Logger_Task = - { + def remote_build_history( + rev: String, + afp_rev: Option[String], + i: Int, + r: Remote_Build + ) : Logger_Task = { val task_name = "build_history-" + r.host - Logger_Task(task_name, logger => - { - using(r.ssh_session(logger.ssh_context))(ssh => - { - val results = - Build_History.remote_build_history(ssh, - isabelle_repos, - isabelle_repos.ext(r.host), - isabelle_identifier = "cronjob_build_history", - self_update = r.self_update, - rev = rev, - afp_rev = afp_rev, - options = - " -N " + Bash.string(task_name) + (if (i < 0) "" else "_" + (i + 1).toString) + - " -R " + Bash.string(Components.default_component_repository) + - " -C '$USER_HOME/.isabelle/contrib' -f " + - r.build_history_options, - args = "-o timeout=10800 " + r.args) + Logger_Task(task_name, + { logger => + using(r.ssh_session(logger.ssh_context)) { ssh => + val results = + Build_History.remote_build_history(ssh, + isabelle_repos, + isabelle_repos.ext(r.host), + isabelle_identifier = "cronjob_build_history", + self_update = r.self_update, + rev = rev, + afp_rev = afp_rev, + options = + " -N " + Bash.string(task_name) + (if (i < 0) "" else "_" + (i + 1).toString) + + " -R " + Bash.string(Components.default_component_repository) + + " -C '$USER_HOME/.isabelle/contrib' -f " + + r.build_history_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) - } - }) + for ((log_name, bytes) <- results) { + logger.log(Date.now(), log_name) + Bytes.write(logger.log_dir + Path.explode(log_name), bytes) + } + } }) } @@ -438,20 +436,19 @@ /** task logging **/ - object Log_Service - { + object Log_Service { def apply(options: Options, progress: Progress = new Progress): Log_Service = new Log_Service(SSH.init_context(options), progress) } - class Log_Service private(val ssh_context: SSH.Context, progress: Progress) - { + class Log_Service private(val ssh_context: SSH.Context, progress: Progress) { current_log.file.delete private val thread: Consumer_Thread[String] = Consumer_Thread.fork("cronjob: logger", daemon = true)( - consume = (text: String) => - { // critical + consume = + { (text: String) => + // critical File.append(current_log, text + "\n") File.append(cumulative_log, text + "\n") progress.echo(text) @@ -470,8 +467,7 @@ 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): Unit = - { + def run_task(start_date: Date, task: Logger_Task): Unit = { val logger = start_logger(start_date, task.name) val res = Exn.capture { task.body(logger) } val end_date = Date.now() @@ -492,15 +488,16 @@ } class Logger private[Isabelle_Cronjob]( - val log_service: Log_Service, val start_date: Date, val task_name: String) - { + 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]): Unit = - { + def log_end(end_date: Date, err: Option[String]): Unit = { val elapsed_time = end_date.time - start_date.time val msg = (if (err.isEmpty) "finished" else "ERROR " + err.get) + @@ -513,8 +510,7 @@ log(start_date, "started") } - class Task private[Isabelle_Cronjob](name: String, body: => Unit) - { + 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 } @@ -523,8 +519,7 @@ /** cronjob **/ - def cronjob(progress: Progress, exclude_task: Set[String]): Unit = - { + def cronjob(progress: Progress, exclude_task: Set[String]): Unit = { /* soft lock */ val still_running = @@ -553,22 +548,22 @@ 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]): Unit = - { - running.partition(_.is_finished) match { - case (Nil, Nil) => - case (Nil, _ :: _) => Time.seconds(0.5).sleep(); join(running) - case (_ :: _, remaining) => join(remaining) + def PAR(tasks: List[Logger_Task]): Logger_Task = + Logger_Task(body = + { _ => + @tailrec def join(running: List[Task]): Unit = { + running.partition(_.is_finished) match { + case (Nil, Nil) => + case (Nil, _ :: _) => Time.seconds(0.5).sleep(); 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) - }) + 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) + }) /* repository structure */ @@ -576,8 +571,7 @@ val hg = Mercurial.repository(isabelle_repos) val hg_graph = hg.graph() - def history_base_filter(r: Remote_Build): Item => Boolean = - { + def history_base_filter(r: Remote_Build): Item => Boolean = { val base_rev = hg.id(r.history_base) val nodes = hg_graph.all_succs(List(base_rev)).toSet (item: Item) => nodes(item.isabelle_version) @@ -621,8 +615,7 @@ /** command line entry point **/ - def main(args: Array[String]): Unit = - { + def main(args: Array[String]): Unit = { Command_Line.tool { var force = false var verbose = false diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/isabelle_devel.scala --- a/src/Pure/Admin/isabelle_devel.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/isabelle_devel.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Isabelle_Devel -{ +object Isabelle_Devel { val RELEASE_SNAPSHOT = "release_snapshot" val BUILD_LOG_DB = "build_log.db" val BUILD_STATUS = "build_status" @@ -20,8 +19,7 @@ /* index */ - def make_index(): Unit = - { + def make_index(): Unit = { val redirect = "https://isabelle-dev.sketis.net/home/menu/view/20" HTML.write_document(root, "index.html", @@ -34,41 +32,36 @@ /* release snapshot */ - def release_snapshot(options: Options, rev: String, afp_rev: String): Unit = - { - Isabelle_System.with_tmp_dir("isadist")(target_dir => - { - Isabelle_System.update_directory(root + Path.explode(RELEASE_SNAPSHOT), - website_dir => - { + def release_snapshot(options: Options, rev: String, afp_rev: String): Unit = { + Isabelle_System.with_tmp_dir("isadist") { target_dir => + Isabelle_System.update_directory(root + Path.explode(RELEASE_SNAPSHOT), + { website_dir => val context = Build_Release.Release_Context(target_dir) Build_Release.build_release_archive(context, rev) Build_Release.build_release(options, context, afp_rev = afp_rev, build_sessions = List(Isabelle_System.getenv("ISABELLE_LOGIC")), website = Some(website_dir)) - }) - }) + } + ) + } } /* maintain build_log database */ - def build_log_database(options: Options, log_dirs: List[Path]): Unit = - { + def build_log_database(options: Options, log_dirs: List[Path]): Unit = { val store = Build_Log.store(options) - using(store.open_database())(db => - { + using(store.open_database()) { db => store.update_database(db, log_dirs) store.update_database(db, log_dirs, ml_statistics = true) store.snapshot_database(db, root + Path.explode(BUILD_LOG_DB)) - }) + } } /* present build status */ - def build_status(options: Options): Unit = - { + def build_status(options: Options): Unit = { Isabelle_System.update_directory(root + Path.explode(BUILD_STATUS), dir => Build_Status.build_status(options, target_dir = dir, ml_statistics = true)) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/jenkins.scala --- a/src/Pure/Admin/jenkins.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/jenkins.scala Fri Apr 01 23:26:19 2022 +0200 @@ -12,15 +12,13 @@ import scala.util.matching.Regex -object Jenkins -{ +object Jenkins { /* server API */ def root(): String = Isabelle_System.getenv_strict("ISABELLE_JENKINS_ROOT") - def invoke(url: String, args: String*): Any = - { + def invoke(url: String, args: String*): Any = { val req = url + "/api/json?" + args.mkString("&") val result = Url.read(req) try { JSON.parse(result) } @@ -40,8 +38,11 @@ def download_logs( - options: Options, job_names: List[String], dir: Path, progress: Progress = new Progress): Unit = - { + options: Options, + job_names: List[String], + dir: Path, + progress: Progress = new Progress + ): Unit = { val store = Sessions.store(options) val infos = job_names.flatMap(build_job_infos) Par_List.map((info: Job_Info) => info.download_log(store, dir, progress), infos) @@ -68,15 +69,14 @@ job_name: String, timestamp: Long, main_log: URL, - session_logs: List[(String, String, URL)]) - { + session_logs: List[(String, String, URL)] + ) { val date: Date = Date(Time.ms(timestamp), Date.timezone_berlin) def log_filename: Path = Build_Log.log_filename(Build_Log.Jenkins.engine, date, List(job_name)) - def read_ml_statistics(store: Sessions.Store, session_name: String): List[Properties.T] = - { + def read_ml_statistics(store: Sessions.Store, session_name: String): List[Properties.T] = { def get_log(ext: String): Option[URL] = session_logs.collectFirst({ case (a, b, url) if a == session_name && b == ext => url }) @@ -96,8 +96,7 @@ } } - def download_log(store: Sessions.Store, dir: Path, progress: Progress = new Progress): Unit = - { + def download_log(store: Sessions.Store, dir: Path, progress: Progress = new Progress): Unit = { val log_dir = dir + Build_Log.log_subdir(date) val log_path = log_dir + log_filename.xz @@ -118,8 +117,7 @@ } } - def build_job_infos(job_name: String): List[Job_Info] = - { + def build_job_infos(job_name: String): List[Job_Info] = { val Session_Log = new Regex("""^.*/log/([^/]+)\.(db|gz)$""") val infos = diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Admin/other_isabelle.scala --- a/src/Pure/Admin/other_isabelle.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Admin/other_isabelle.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Other_Isabelle -{ +object Other_Isabelle { def apply(isabelle_home: Path, isabelle_identifier: String = "", user_home: Path = Path.USER_HOME, @@ -20,8 +19,8 @@ val isabelle_home: Path, val isabelle_identifier: String, user_home: Path, - progress: Progress) -{ + progress: Progress +) { other_isabelle => override def toString: String = isabelle_home.toString @@ -66,8 +65,7 @@ /* NEWS */ - def make_news(): Unit = - { + def make_news(): Unit = { val doc_dir = isabelle_home + Path.explode("doc") val fonts_dir = Isabelle_System.make_directory(doc_dir + Path.explode("fonts")) @@ -88,8 +86,8 @@ component_repository: String = Components.default_component_repository, components_base: Path = Components.default_components_base, catalogs: List[String] = Nil, - components: List[String] = Nil): List[String] = - { + components: List[String] = Nil + ): List[String] = { val dir = Components.admin(isabelle_home) ("ISABELLE_COMPONENT_REPOSITORY=" + Bash.string(component_repository)) :: @@ -110,8 +108,7 @@ } else false - def init_settings(settings: List[String]): Unit = - { + def init_settings(settings: List[String]): Unit = { if (!clean_settings()) error("Cannot proceed with existing user settings file: " + etc_settings) @@ -125,8 +122,7 @@ /* cleanup */ - def cleanup(): Unit = - { + def cleanup(): Unit = { clean_settings() etc.file.delete isabelle_home_user.file.delete diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Concurrent/consumer_thread.scala --- a/src/Pure/Concurrent/consumer_thread.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Concurrent/consumer_thread.scala Fri Apr 01 23:26:19 2022 +0200 @@ -11,8 +11,7 @@ import scala.annotation.tailrec -object Consumer_Thread -{ +object Consumer_Thread { def fork_bulk[A](name: String = "", daemon: Boolean = false)( bulk: A => Boolean, consume: List[A] => (List[Exn.Result[Unit]], Boolean), @@ -21,10 +20,9 @@ def fork[A](name: String = "", daemon: Boolean = false)( consume: A => Boolean, - finish: () => Unit = () => ()): Consumer_Thread[A] = - { - def consume_single(args: List[A]): (List[Exn.Result[Unit]], Boolean) = - { + finish: () => Unit = () => () + ): Consumer_Thread[A] = { + def consume_single(args: List[A]): (List[Exn.Result[Unit]], Boolean) = { assert(args.length == 1) Exn.capture { consume(args.head) } match { case Exn.Res(continue) => (List(Exn.Res(())), continue) @@ -40,8 +38,8 @@ name: String, daemon: Boolean, bulk: A => Boolean, consume: List[A] => (List[Exn.Result[Unit]], Boolean), - finish: () => Unit) -{ + finish: () => Unit +) { /* thread */ private var active = true @@ -61,21 +59,18 @@ /* requests */ - private class Request(val arg: A, acknowledge: Boolean = false) - { + private class Request(val arg: A, acknowledge: Boolean = false) { val ack: Option[Synchronized[Option[Exn.Result[Unit]]]] = if (acknowledge) Some(Synchronized(None)) else None - def await(): Unit = - { + def await(): Unit = { for (a <- ack) { Exn.release(a.guarded_access({ case None => None case res => Some((res.get, res)) })) } } } - private def request(req: Request): Unit = - { + private def request(req: Request): Unit = { synchronized { if (is_active()) mailbox.send(Some(req)) else error("Consumer thread not active: " + quote(thread.getName)) @@ -95,8 +90,9 @@ val (results, continue) = consume(reqs.map(_.arg)) - for { (Some(req), Some(res)) <- reqs.map(Some(_)).zipAll(results.map(Some(_)), None, None) } - { + for { + (Some(req), Some(res)) <- reqs.map(Some(_)).zipAll(results.map(Some(_)), None, None) + } { (req.ack, res) match { case (Some(a), _) => a.change(_ => Some(res)) case (None, Exn.Res(_)) => @@ -116,8 +112,7 @@ def send(arg: A): Unit = request(new Request(arg)) def send_wait(arg: A): Unit = request(new Request(arg, acknowledge = true)) - def shutdown(): Unit = - { + def shutdown(): Unit = { synchronized { if (is_active()) { active = false; mailbox.send(None) } } thread.join() } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Concurrent/counter.scala --- a/src/Pure/Concurrent/counter.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Concurrent/counter.scala Fri Apr 01 23:26:19 2022 +0200 @@ -9,14 +9,12 @@ package isabelle -object Counter -{ +object Counter { type ID = Long def make(): Counter = new Counter } -final class Counter private -{ +final class Counter private { private var count: Counter.ID = 0 def apply(): Counter.ID = synchronized { diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Concurrent/delay.scala --- a/src/Pure/Concurrent/delay.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Concurrent/delay.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Delay -{ +object Delay { // delayed event after first invocation def first(delay: => Time, log: Logger = No_Logger, gui: Boolean = false)(event: => Unit): Delay = new Delay(true, delay, log, if (gui) GUI_Thread.later { event } else event ) @@ -18,12 +17,10 @@ new Delay(false, delay, log, if (gui) GUI_Thread.later { event } else event ) } -final class Delay private(first: Boolean, delay: => Time, log: Logger, event: => Unit) -{ +final class Delay private(first: Boolean, delay: => Time, log: Logger, event: => Unit) { private var running: Option[Event_Timer.Request] = None - private def run: Unit = - { + private def run: Unit = { val do_run = synchronized { if (running.isDefined) { running = None; true } else false } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Concurrent/event_timer.scala --- a/src/Pure/Concurrent/event_timer.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Concurrent/event_timer.scala Fri Apr 01 23:26:19 2022 +0200 @@ -13,17 +13,18 @@ import java.util.{Timer, TimerTask, Date => JDate} -object Event_Timer -{ +object Event_Timer { private lazy val event_timer = new Timer("event_timer", true) - final class Request private[Event_Timer](val time: Time, val repeat: Option[Time], task: TimerTask) - { + final class Request private[Event_Timer]( + val time: Time, + val repeat: Option[Time], + task: TimerTask + ) { def cancel(): Boolean = task.cancel() } - def request(time: Time, repeat: Option[Time] = None)(event: => Unit): Request = - { + def request(time: Time, repeat: Option[Time] = None)(event: => Unit): Request = { val task = new TimerTask { def run: Unit = event } repeat match { case None => event_timer.schedule(task, new JDate(time.ms)) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Concurrent/future.scala --- a/src/Pure/Concurrent/future.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Concurrent/future.scala Fri Apr 01 23:26:19 2022 +0200 @@ -12,8 +12,7 @@ /* futures and promises */ -object Future -{ +object Future { def value[A](x: A): Future[A] = new Value_Future(x) def fork[A](body: => A): Future[A] = new Task_Future[A](body) def promise[A]: Promise[A] = new Promise_Future[A] @@ -24,14 +23,14 @@ pri: Int = Thread.NORM_PRIORITY, daemon: Boolean = false, inherit_locals: Boolean = false, - uninterruptible: Boolean = false)(body: => A): Future[A] = - { - new Thread_Future[A](name, group, pri, daemon, inherit_locals, uninterruptible, body) - } + uninterruptible: Boolean = false)( + body: => A + ): Future[A] = { + new Thread_Future[A](name, group, pri, daemon, inherit_locals, uninterruptible, body) + } } -trait Future[A] -{ +trait Future[A] { def peek: Option[Exn.Result[A]] def is_finished: Boolean = peek.isDefined def get_finished: A = { require(is_finished, "future not finished"); Exn.release(peek.get) } @@ -48,8 +47,7 @@ } } -trait Promise[A] extends Future[A] -{ +trait Promise[A] extends Future[A] { def fulfill_result(res: Exn.Result[A]): Unit def fulfill(x: A): Unit } @@ -57,8 +55,7 @@ /* value future */ -private class Value_Future[A](x: A) extends Future[A] -{ +private class Value_Future[A](x: A) extends Future[A] { val peek: Option[Exn.Result[A]] = Some(Exn.Res(x)) def join_result: Exn.Result[A] = peek.get def cancel(): Unit = {} @@ -67,8 +64,7 @@ /* task future via thread pool */ -private class Task_Future[A](body: => A) extends Future[A] -{ +private class Task_Future[A](body: => A) extends Future[A] { private sealed abstract class Status private case object Ready extends Status private case class Running(thread: Thread) extends Status @@ -83,8 +79,7 @@ case _ => None } - private def try_run(): Unit = - { + private def try_run(): Unit = { val do_run = status.change_result { case Ready => (true, Running(Thread.currentThread)) @@ -98,8 +93,7 @@ } private val task = Isabelle_Thread.pool.submit(new Callable[Unit] { def call = try_run() }) - def join_result: Exn.Result[A] = - { + def join_result: Exn.Result[A] = { try_run() status.guarded_access { case st @ Finished(result) => Some((result, st)) @@ -107,8 +101,7 @@ } } - def cancel(): Unit = - { + def cancel(): Unit = { status.change { case Ready => task.cancel(false); Finished(Exn.Exn(Exn.Interrupt())) case st @ Running(thread) => thread.interrupt(); st @@ -120,8 +113,7 @@ /* promise future */ -private class Promise_Future[A] extends Promise[A] -{ +private class Promise_Future[A] extends Promise[A] { private val state = Synchronized[Option[Exn.Result[A]]](None) def peek: Option[Exn.Result[A]] = state.value @@ -147,13 +139,12 @@ daemon: Boolean, inherit_locals: Boolean, uninterruptible: Boolean, - body: => A) extends Future[A] -{ + body: => A) extends Future[A] { private val result = Future.promise[A] private val thread = Isabelle_Thread.fork(name = name, group = group, pri = pri, daemon = daemon, - inherit_locals = inherit_locals, uninterruptible = uninterruptible) - { result.fulfill_result(Exn.capture(body)) } + inherit_locals = inherit_locals, uninterruptible = uninterruptible) { + result.fulfill_result(Exn.capture(body)) } def peek: Option[Exn.Result[A]] = result.peek def join_result: Exn.Result[A] = result.join_result diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Concurrent/isabelle_thread.scala --- a/src/Pure/Concurrent/isabelle_thread.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Concurrent/isabelle_thread.scala Fri Apr 01 23:26:19 2022 +0200 @@ -10,8 +10,7 @@ import java.util.concurrent.{ThreadPoolExecutor, TimeUnit, LinkedBlockingQueue} -object Isabelle_Thread -{ +object Isabelle_Thread { /* self-thread */ def self: Isabelle_Thread = @@ -28,8 +27,7 @@ private val counter = Counter.make() - def make_name(name: String = "", base: String = "thread"): String = - { + def make_name(name: String = "", base: String = "thread"): String = { val prefix = "Isabelle." val suffix = if (name.nonEmpty) name else base + counter() if (suffix.startsWith(prefix)) suffix else prefix + suffix @@ -46,8 +44,8 @@ group: ThreadGroup = current_thread_group, pri: Int = Thread.NORM_PRIORITY, daemon: Boolean = false, - inherit_locals: Boolean = false): Isabelle_Thread = - { + inherit_locals: Boolean = false + ): Isabelle_Thread = { new Isabelle_Thread(main, name = make_name(name = name), group = group, pri = pri, daemon = daemon, inherit_locals = inherit_locals) } @@ -58,8 +56,9 @@ pri: Int = Thread.NORM_PRIORITY, daemon: Boolean = false, inherit_locals: Boolean = false, - uninterruptible: Boolean = false)(body: => Unit): Isabelle_Thread = - { + uninterruptible: Boolean = false)( + body: => Unit + ): Isabelle_Thread = { val main: Runnable = if (uninterruptible) { () => Isabelle_Thread.uninterruptible { body } } else { () => body } @@ -73,14 +72,12 @@ /* thread pool */ - def max_threads(): Int = - { + def max_threads(): Int = { val m = Value.Int.unapply(System.getProperty("isabelle.threads", "0")) getOrElse 0 if (m > 0) m else (Runtime.getRuntime.availableProcessors max 1) min 8 } - lazy val pool: ThreadPoolExecutor = - { + lazy val pool: ThreadPoolExecutor = { val n = max_threads() val executor = new ThreadPoolExecutor(n, n, 2500L, TimeUnit.MILLISECONDS, new LinkedBlockingQueue[Runnable]) @@ -92,8 +89,7 @@ /* interrupt handlers */ - object Interrupt_Handler - { + object Interrupt_Handler { def apply(handle: Isabelle_Thread => Unit, name: String = "handler"): Interrupt_Handler = new Interrupt_Handler(handle, name) @@ -105,8 +101,7 @@ } class Interrupt_Handler private(handle: Isabelle_Thread => Unit, name: String) - extends Function[Isabelle_Thread, Unit] - { + extends Function[Isabelle_Thread, Unit] { def apply(thread: Isabelle_Thread): Unit = handle(thread) override def toString: String = name } @@ -131,8 +126,7 @@ class Isabelle_Thread private(main: Runnable, name: String, group: ThreadGroup, pri: Int, daemon: Boolean, inherit_locals: Boolean) - extends Thread(group, null, name, 0L, inherit_locals) -{ + extends Thread(group, null, name, 0L, inherit_locals) { thread => thread.setPriority(pri) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Concurrent/mailbox.scala --- a/src/Pure/Concurrent/mailbox.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Concurrent/mailbox.scala Fri Apr 01 23:26:19 2022 +0200 @@ -8,14 +8,12 @@ package isabelle -object Mailbox -{ +object Mailbox { def apply[A]: Mailbox[A] = new Mailbox[A]() } -class Mailbox[A] private() -{ +class Mailbox[A] private() { private val mailbox = Synchronized[List[A]](Nil) override def toString: String = mailbox.value.reverse.mkString("Mailbox(", ",", ")") diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Concurrent/par_list.scala --- a/src/Pure/Concurrent/par_list.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Concurrent/par_list.scala Fri Apr 01 23:26:19 2022 +0200 @@ -8,11 +8,12 @@ package isabelle -object Par_List -{ - def managed_results[A, B](f: A => B, xs: List[A], sequential: Boolean = false) - : List[Exn.Result[B]] = - { +object Par_List { + def managed_results[A, B]( + f: A => B, + xs: List[A], + sequential: Boolean = false + ) : List[Exn.Result[B]] = { if (sequential || xs.isEmpty || xs.tail.isEmpty) xs.map(x => Exn.capture { f(x) }) else { val state = Synchronized[(List[Future[B]], Boolean)]((Nil, false)) @@ -46,8 +47,7 @@ def map[A, B](f: A => B, xs: List[A], sequential: Boolean = false): List[B] = Exn.release_first(managed_results(f, xs, sequential = sequential)) - def get_some[A, B](f: A => Option[B], xs: List[A]): Option[B] = - { + def get_some[A, B](f: A => Option[B], xs: List[A]): Option[B] = { class Found(val res: B) extends Exception val results = managed_results( diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Concurrent/synchronized.scala --- a/src/Pure/Concurrent/synchronized.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Concurrent/synchronized.scala Fri Apr 01 23:26:19 2022 +0200 @@ -10,14 +10,12 @@ import scala.annotation.tailrec -object Synchronized -{ +object Synchronized { def apply[A](init: A): Synchronized[A] = new Synchronized(init) } -final class Synchronized[A] private(init: A) -{ +final class Synchronized[A] private(init: A) { /* state variable */ private var state: A = init @@ -38,8 +36,7 @@ notifyAll() Some(y) } - @tailrec def try_change(): Option[B] = - { + @tailrec def try_change(): Option[B] = { val x = state check(x) match { case None => diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/GUI/color_value.scala --- a/src/Pure/GUI/color_value.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/GUI/color_value.scala Fri Apr 01 23:26:19 2022 +0200 @@ -11,12 +11,10 @@ import java.util.Locale -object Color_Value -{ +object Color_Value { private var cache = Map.empty[String, Color] - def parse(s: String): Color = - { + def parse(s: String): Color = { val i = java.lang.Long.parseLong(s, 16) val r = ((i >> 24) & 0xFF).toInt val g = ((i >> 16) & 0xFF).toInt @@ -25,8 +23,7 @@ new Color(r, g, b, a) } - def print(c: Color): String = - { + def print(c: Color): String = { val r = java.lang.Integer.valueOf(c.getRed) val g = java.lang.Integer.valueOf(c.getGreen) val b = java.lang.Integer.valueOf(c.getBlue) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/GUI/desktop_app.scala --- a/src/Pure/GUI/desktop_app.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/GUI/desktop_app.scala Fri Apr 01 23:26:19 2022 +0200 @@ -9,8 +9,7 @@ import java.awt.Desktop -object Desktop_App -{ +object Desktop_App { def desktop_action(action: Desktop.Action, f: Desktop => Unit): Unit = if (Desktop.isDesktopSupported) { val desktop = Desktop.getDesktop diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/GUI/gui.scala --- a/src/Pure/GUI/gui.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/GUI/gui.scala Fri Apr 01 23:26:19 2022 +0200 @@ -17,8 +17,7 @@ import scala.swing.event.SelectionChanged -object GUI -{ +object GUI { /* Swing look-and-feel */ def init_laf(): Unit = com.formdev.flatlaf.FlatLightLaf.setup() @@ -28,8 +27,7 @@ def is_macos_laf: Boolean = Platform.is_macos && UIManager.getSystemLookAndFeelClassName() == current_laf - class Look_And_Feel(laf: LookAndFeel) extends Isabelle_System.Service - { + class Look_And_Feel(laf: LookAndFeel) extends Isabelle_System.Service { def info: UIManager.LookAndFeelInfo = new UIManager.LookAndFeelInfo(laf.getName, laf.getClass.getName) } @@ -37,8 +35,7 @@ lazy val look_and_feels: List[Look_And_Feel] = Isabelle_System.make_services(classOf[Look_And_Feel]) - def init_lafs(): Unit = - { + def init_lafs(): Unit = { val old_lafs = Set( "com.sun.java.swing.plaf.motif.MotifLookAndFeel", @@ -55,8 +52,7 @@ /* plain focus traversal, notably for text fields */ - def plain_focus_traversal(component: Component): Unit = - { + def plain_focus_traversal(component: Component): Unit = { val dummy_button = new JButton def apply(id: Int): Unit = component.setFocusTraversalKeys(id, dummy_button.getFocusTraversalKeys(id)) @@ -67,9 +63,12 @@ /* simple dialogs */ - def scrollable_text(raw_txt: String, width: Int = 60, height: Int = 20, editable: Boolean = false) - : ScrollPane = - { + def scrollable_text( + raw_txt: String, + width: Int = 60, + height: Int = 20, + editable: Boolean = false + ) : ScrollPane = { val txt = Output.clean_yxml(raw_txt) val text = new TextArea(txt) if (width > 0) text.columns = width @@ -78,9 +77,13 @@ new ScrollPane(text) } - private def simple_dialog(kind: Int, default_title: String, - parent: Component, title: String, message: Iterable[Any]): Unit = - { + private def simple_dialog( + kind: Int, + default_title: String, + parent: Component, + title: String, + message: Iterable[Any] + ): Unit = { GUI_Thread.now { val java_message = message.iterator.map({ case x: scala.swing.Component => x.peer case x => x }). @@ -113,8 +116,8 @@ private val Zoom_Factor = "([0-9]+)%?".r abstract class Zoom_Box extends ComboBox[String]( - List("50%", "70%", "85%", "100%", "125%", "150%", "175%", "200%", "300%", "400%")) - { + List("50%", "70%", "85%", "100%", "125%", "150%", "175%", "200%", "300%", "400%") + ) { def changed: Unit def factor: Int = parse(selection.item) @@ -128,8 +131,7 @@ private def print(i: Int): String = i.toString + "%" - def set_item(i: Int): Unit = - { + def set_item(i: Int): Unit = { peer.getEditor match { case null => case editor => editor.setItem(print(i)) @@ -170,10 +172,8 @@ /* location within multi-screen environment */ - final case class Screen_Location(point: Point, bounds: Rectangle) - { - def relative(parent: Component, size: Dimension): Point = - { + final case class Screen_Location(point: Point, bounds: Rectangle) { + def relative(parent: Component, size: Dimension): Point = { val w = size.width val h = size.height @@ -190,8 +190,7 @@ } } - def screen_location(component: Component, point: Point): Screen_Location = - { + def screen_location(component: Component, point: Point): Screen_Location = { val screen_point = new Point(point.x, point.y) if (component != null) SwingUtilities.convertPointToScreen(screen_point, component) @@ -212,8 +211,7 @@ /* screen size */ - sealed case class Screen_Size(bounds: Rectangle, insets: Insets) - { + sealed case class Screen_Size(bounds: Rectangle, insets: Insets) { def full_screen_bounds: Rectangle = if (Platform.is_linux) { // avoid menu bar and docking areas @@ -234,8 +232,7 @@ else bounds } - def screen_size(component: Component): Screen_Size = - { + def screen_size(component: Component): Screen_Size = { val config = component.getGraphicsConfiguration val bounds = config.getBounds val insets = Toolkit.getDefaultToolkit.getScreenInsets(config) @@ -274,10 +271,8 @@ case _ => None } - def traverse_components(component: Component, apply: Component => Unit): Unit = - { - def traverse(comp: Component): Unit = - { + def traverse_components(component: Component, apply: Component => Unit): Unit = { + def traverse(comp: Component): Unit = { apply(comp) comp match { case cont: Container => @@ -310,43 +305,44 @@ /* Isabelle fonts */ - def imitate_font(font: Font, + def imitate_font( + font: Font, family: String = Isabelle_Fonts.sans, - scale: Double = 1.0): Font = - { + scale: Double = 1.0 + ): Font = { val font1 = new Font(family, font.getStyle, font.getSize) val rel_size = line_metrics(font).getHeight.toDouble / line_metrics(font1).getHeight new Font(family, font.getStyle, (scale * rel_size * font.getSize).toInt) } - def imitate_font_css(font: Font, + def imitate_font_css( + font: Font, family: String = Isabelle_Fonts.sans, - scale: Double = 1.0): String = - { + scale: Double = 1.0 + ): String = { val font1 = new Font(family, font.getStyle, font.getSize) val rel_size = line_metrics(font).getHeight.toDouble / line_metrics(font1).getHeight "font-family: " + family + "; font-size: " + (scale * rel_size * 100).toInt + "%;" } - def use_isabelle_fonts(): Unit = - { + def use_isabelle_fonts(): Unit = { val default_font = label_font() val ui = UIManager.getDefaults - for (prop <- List( - "ToggleButton.font", - "CheckBoxMenuItem.font", - "Label.font", - "Menu.font", - "MenuItem.font", - "PopupMenu.font", - "Table.font", - "TableHeader.font", - "TextArea.font", - "TextField.font", - "TextPane.font", - "ToolTip.font", - "Tree.font")) - { + for (prop <- + List( + "ToggleButton.font", + "CheckBoxMenuItem.font", + "Label.font", + "Menu.font", + "MenuItem.font", + "PopupMenu.font", + "Table.font", + "TableHeader.font", + "TextArea.font", + "TextField.font", + "TextPane.font", + "ToolTip.font", + "Tree.font")) { val font = ui.get(prop) match { case font: Font => font case _ => default_font } ui.put(prop, GUI.imitate_font(font)) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/GUI/gui_thread.scala --- a/src/Pure/GUI/gui_thread.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/GUI/gui_thread.scala Fri Apr 01 23:26:19 2022 +0200 @@ -10,18 +10,15 @@ import javax.swing.SwingUtilities -object GUI_Thread -{ +object GUI_Thread { /* context check */ - def assert[A](body: => A): A = - { + def assert[A](body: => A): A = { Predef.assert(SwingUtilities.isEventDispatchThread) body } - def require[A](body: => A): A = - { + def require[A](body: => A): A = { Predef.require(SwingUtilities.isEventDispatchThread, "GUI thread expected") body } @@ -29,8 +26,7 @@ /* event dispatch queue */ - def now[A](body: => A): A = - { + def now[A](body: => A): A = { if (SwingUtilities.isEventDispatchThread) body else { lazy val result = assert { Exn.capture(body) } @@ -39,14 +35,12 @@ } } - def later(body: => Unit): Unit = - { + def later(body: => Unit): Unit = { if (SwingUtilities.isEventDispatchThread) body else SwingUtilities.invokeLater(() => body) } - def future[A](body: => A): Future[A] = - { + def future[A](body: => A): Future[A] = { if (SwingUtilities.isEventDispatchThread) Future.value(body) else { val promise = Future.promise[A] diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/GUI/popup.scala --- a/src/Pure/GUI/popup.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/GUI/popup.scala Fri Apr 01 23:26:19 2022 +0200 @@ -15,10 +15,9 @@ layered: JLayeredPane, component: JComponent, location: Point, - size: Dimension) -{ - def show: Unit = - { + size: Dimension +) { + def show: Unit = { component.setLocation(location) component.setSize(size) component.setPreferredSize(size) @@ -28,8 +27,7 @@ layered.repaint(component.getBounds()) } - def hide: Unit = - { + def hide: Unit = { val bounds = component.getBounds() layered.remove(component) layered.repaint(bounds) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/GUI/wrap_panel.scala --- a/src/Pure/GUI/wrap_panel.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/GUI/wrap_panel.scala Fri Apr 01 23:26:19 2022 +0200 @@ -16,25 +16,21 @@ import scala.swing.{Panel, FlowPanel, Component, SequentialContainer, ScrollPane} -object Wrap_Panel -{ +object Wrap_Panel { val Alignment = FlowPanel.Alignment class Layout(align: Int = FlowLayout.CENTER, hgap: Int = 5, vgap: Int = 5) - extends FlowLayout(align: Int, hgap: Int, vgap: Int) - { + extends FlowLayout(align: Int, hgap: Int, vgap: Int) { override def preferredLayoutSize(target: Container): Dimension = layout_size(target, true) - override def minimumLayoutSize(target: Container): Dimension = - { + override def minimumLayoutSize(target: Container): Dimension = { val minimum = layout_size(target, false) minimum.width -= (getHgap + 1) minimum } - private def layout_size(target: Container, preferred: Boolean): Dimension = - { + private def layout_size(target: Container, preferred: Boolean): Dimension = { target.getTreeLock.synchronized { val target_width = if (target.getSize.width == 0) Integer.MAX_VALUE @@ -52,8 +48,7 @@ val dim = new Dimension(0, 0) var row_width = 0 var row_height = 0 - def add_row(): Unit = - { + def add_row(): Unit = { dim.width = dim.width max row_width if (dim.height > 0) dim.height += vgap dim.height += row_height @@ -64,8 +59,7 @@ m = target.getComponent(i) if m.isVisible d = if (preferred) m.getPreferredSize else m.getMinimumSize() - } - { + } { if (row_width + d.width > max_width) { add_row() row_width = 0 @@ -105,10 +99,10 @@ new Wrap_Panel(contents, alignment) } -class Wrap_Panel(contents0: List[Component] = Nil, - alignment: Wrap_Panel.Alignment.Value = Wrap_Panel.Alignment.Right) - extends Panel with SequentialContainer.Wrapper -{ +class Wrap_Panel( + contents0: List[Component] = Nil, + alignment: Wrap_Panel.Alignment.Value = Wrap_Panel.Alignment.Right) +extends Panel with SequentialContainer.Wrapper { override lazy val peer: JPanel = new JPanel(new Wrap_Panel.Layout(alignment.id)) with SuperMixin diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/antiquote.scala --- a/src/Pure/General/antiquote.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/antiquote.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Antiquote -{ +object Antiquote { sealed abstract class Antiquote { def source: String } case class Text(source: String) extends Antiquote case class Control(source: String) extends Antiquote @@ -19,8 +18,7 @@ object Parsers extends Parsers - trait Parsers extends Scan.Parsers - { + trait Parsers extends Scan.Parsers { private val txt: Parser[String] = rep1(many1(s => !Symbol.is_control(s) && !Symbol.is_open(s) && s != "@") | "@" <~ guard(opt_term(one(s => s != "{")))) ^^ (x => x.mkString) @@ -45,8 +43,7 @@ /* read */ - def read(input: CharSequence): List[Antiquote] = - { + def read(input: CharSequence): List[Antiquote] = { Parsers.parseAll(Parsers.rep(Parsers.antiquote), Scan.char_reader(input)) match { case Parsers.Success(xs, _) => xs case Parsers.NoSuccess(_, next) => @@ -55,8 +52,7 @@ } } - def read_antiq_body(input: CharSequence): Option[String] = - { + def read_antiq_body(input: CharSequence): Option[String] = { read(input) match { case List(Antiq(source)) => Some(source.substring(2, source.length - 1)) case _ => None diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/bytes.scala --- a/src/Pure/General/bytes.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/bytes.scala Fri Apr 01 23:26:19 2022 +0200 @@ -15,12 +15,10 @@ import org.tukaani.xz.{XZInputStream, XZOutputStream} -object Bytes -{ +object Bytes { val empty: Bytes = new Bytes(Array[Byte](), 0, 0) - def apply(s: CharSequence): Bytes = - { + def apply(s: CharSequence): Bytes = { val str = s.toString if (str.isEmpty) empty else { @@ -44,21 +42,18 @@ /* base64 */ - def base64(s: String): Bytes = - { + def base64(s: String): Bytes = { val a = Base64.getDecoder.decode(s) new Bytes(a, 0, a.length) } - object Decode_Base64 extends Scala.Fun("decode_base64") - { + object Decode_Base64 extends Scala.Fun("decode_base64") { val here = Scala_Project.here def invoke(args: List[Bytes]): List[Bytes] = List(base64(Library.the_single(args).text)) } - object Encode_Base64 extends Scala.Fun("encode_base64") - { + object Encode_Base64 extends Scala.Fun("encode_base64") { val here = Scala_Project.here def invoke(args: List[Bytes]): List[Bytes] = List(Bytes(Library.the_single(args).base64)) @@ -85,8 +80,7 @@ new Bytes(out.toByteArray, 0, out.size) } - def read(file: JFile): Bytes = - { + def read(file: JFile): Bytes = { val length = file.length val limit = if (length < 0 || length > Integer.MAX_VALUE) Integer.MAX_VALUE else length.toInt using(new FileInputStream(file))(read_stream(_, limit = limit)) @@ -108,12 +102,10 @@ final class Bytes private( protected val bytes: Array[Byte], protected val offset: Int, - val length: Int) extends CharSequence -{ + val length: Int) extends CharSequence { /* equality */ - override def equals(that: Any): Boolean = - { + override def equals(that: Any): Boolean = { that match { case other: Bytes => if (this eq other) true @@ -123,8 +115,7 @@ } } - private lazy val hash: Int = - { + private lazy val hash: Int = { var h = 0 for (i <- offset until offset + length) { val b = bytes(i).asInstanceOf[Int] & 0xFF @@ -146,8 +137,7 @@ for (i <- (offset until (offset + length)).iterator) yield bytes(i) - def array: Array[Byte] = - { + def array: Array[Byte] = { val a = new Array[Byte](length) System.arraycopy(bytes, offset, a, 0, length) a @@ -155,16 +145,14 @@ def text: String = UTF8.decode_permissive(this) - def base64: String = - { + def base64: String = { val b = if (offset == 0 && length == bytes.length) bytes else Bytes(bytes, offset, length).bytes Base64.getEncoder.encodeToString(b) } - def maybe_base64: (Boolean, String) = - { + def maybe_base64: (Boolean, String) = { val s = text if (this == Bytes(s)) (false, s) else (true, base64) } @@ -191,8 +179,7 @@ if (0 <= i && i < length) (bytes(offset + i).asInstanceOf[Int] & 0xFF).asInstanceOf[Char] else throw new IndexOutOfBoundsException - def subSequence(i: Int, j: Int): Bytes = - { + def subSequence(i: Int, j: Int): Bytes = { if (0 <= i && i <= j && j <= length) new Bytes(bytes, offset + i, j - i) else throw new IndexOutOfBoundsException } @@ -217,16 +204,16 @@ def uncompress(cache: XZ.Cache = XZ.Cache()): Bytes = using(new XZInputStream(stream(), cache))(Bytes.read_stream(_, hint = length)) - def compress(options: XZ.Options = XZ.options(), cache: XZ.Cache = XZ.Cache()): Bytes = - { + def compress(options: XZ.Options = XZ.options(), cache: XZ.Cache = XZ.Cache()): Bytes = { val result = new ByteArrayOutputStream(length) using(new XZOutputStream(result, options, cache))(write_stream(_)) new Bytes(result.toByteArray, 0, result.size) } - def maybe_compress(options: XZ.Options = XZ.options(), cache: XZ.Cache = XZ.Cache()) - : (Boolean, Bytes) = - { + def maybe_compress( + options: XZ.Options = XZ.options(), + cache: XZ.Cache = XZ.Cache() + ) : (Boolean, Bytes) = { val compressed = compress(options = options, cache = cache) if (compressed.length < length) (true, compressed) else (false, this) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/cache.scala --- a/src/Pure/General/cache.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/cache.scala Fri Apr 01 23:26:19 2022 +0200 @@ -11,8 +11,7 @@ import java.lang.ref.WeakReference -object Cache -{ +object Cache { val default_max_string = 100 val default_initial_size = 131071 @@ -24,8 +23,7 @@ val none: Cache = make(max_string = 0) } -class Cache(max_string: Int, initial_size: Int) -{ +class Cache(max_string: Int, initial_size: Int) { val no_cache: Boolean = max_string == 0 private val table: JMap[Any, WeakReference[Any]] = @@ -35,8 +33,7 @@ override def toString: String = if (no_cache) "Cache.none" else "Cache(size = " + table.size + ")" - protected def lookup[A](x: A): Option[A] = - { + protected def lookup[A](x: A): Option[A] = { if (table == null) None else { val ref = table.get(x) @@ -45,8 +42,7 @@ } } - protected def store[A](x: A): A = - { + protected def store[A](x: A): A = { if (table == null || x == null) x else { table.put(x, new WeakReference[Any](x)) @@ -54,8 +50,7 @@ } } - protected def cache_string(x: String): String = - { + protected def cache_string(x: String): String = { if (x == null) null else if (x == "") "" else if (x == "true") "true" diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/codepoint.scala --- a/src/Pure/General/codepoint.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/codepoint.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,17 +7,13 @@ package isabelle -object Codepoint -{ +object Codepoint { def string(c: Int): String = new String(Array(c), 0, 1) - private class Iterator_Offset[A](s: String, result: (Int, Text.Offset) => A) - extends Iterator[A] - { + private class Iterator_Offset[A](s: String, result: (Int, Text.Offset) => A) extends Iterator[A] { var offset = 0 def hasNext: Boolean = offset < s.length - def next(): A = - { + def next(): A = { val c = s.codePointAt(offset) val i = offset offset += Character.charCount(c) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/comment.scala --- a/src/Pure/General/comment.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/comment.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,10 +7,8 @@ package isabelle -object Comment -{ - object Kind extends Enumeration - { +object Comment { + object Kind extends Enumeration { val COMMENT = Value("formal comment") val CANCEL = Value("canceled text") val LATEX = Value("embedded LaTeX") @@ -26,8 +24,7 @@ lazy val symbols_blanks: Set[Symbol.Symbol] = Set(Symbol.comment, Symbol.comment_decoded) - def content(source: String): String = - { + def content(source: String): String = { def err(): Nothing = error("Malformed formal comment: " + quote(source)) Symbol.explode(source) match { @@ -39,8 +36,7 @@ } } - trait Parsers extends Scan.Parsers - { + trait Parsers extends Scan.Parsers { def comment_prefix: Parser[String] = one(symbols_blanks) ~ many(Symbol.is_blank) ^^ { case a ~ b => a + b.mkString } | one(symbols) @@ -48,8 +44,7 @@ def comment_cartouche: Parser[String] = comment_prefix ~ cartouche ^^ { case a ~ b => a + b } - def comment_cartouche_line(ctxt: Scan.Line_Context): Parser[(String, Scan.Line_Context)] = - { + def comment_cartouche_line(ctxt: Scan.Line_Context): Parser[(String, Scan.Line_Context)] = { def cartouche_context(d: Int): Scan.Line_Context = if (d == 0) Scan.Finished else Scan.Cartouche_Comment(d) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/completion.scala --- a/src/Pure/General/completion.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/completion.scala Fri Apr 01 23:26:19 2022 +0200 @@ -15,8 +15,7 @@ import scala.math.Ordering -object Completion -{ +object Completion { /** completion result **/ sealed case class Item( @@ -28,8 +27,7 @@ move: Int, immediate: Boolean) - object Result - { + object Result { def empty(range: Text.Range): Result = Result(range, "", false, Nil) def merge(history: History, results: Option[Result]*): Option[Result] = results.foldLeft(None: Option[Result]) { @@ -56,12 +54,10 @@ private val COMPLETION_HISTORY = Path.explode("$ISABELLE_HOME_USER/etc/completion_history") - object History - { + object History { val empty: History = new History() - def load(): History = - { + def load(): History = { def ignore_error(msg: String): Unit = Output.warning("Ignoring bad content of file " + COMPLETION_HISTORY + (if (msg == "") "" else "\n" + msg)) @@ -82,16 +78,14 @@ } } - final class History private(rep: SortedMap[String, Int] = SortedMap.empty) - { + final class History private(rep: SortedMap[String, Int] = SortedMap.empty) { override def toString: String = rep.mkString("Completion.History(", ",", ")") def frequency(name: String): Int = default_frequency(Symbol.encode(name)) getOrElse rep.getOrElse(name, 0) - def + (entry: (String, Int)): History = - { + def + (entry: (String, Int)): History = { val (name, freq) = entry if (name == "") this else new History(rep + (name -> (frequency(name) + freq))) @@ -103,8 +97,7 @@ frequency(item2.name) compare frequency(item1.name) } - def save(): Unit = - { + def save(): Unit = { Isabelle_System.make_directory(COMPLETION_HISTORY.dir) File.write_backup(COMPLETION_HISTORY, { @@ -114,13 +107,11 @@ } } - class History_Variable - { + class History_Variable { private var history = History.empty def value: History = synchronized { history } - def load(): Unit = - { + def load(): Unit = { val h = History.load() synchronized { history = h } } @@ -155,12 +146,9 @@ def report_theories(pos: Position.T, thys: List[String], total: Int = 0): String = report_names(pos, thys.map(name => (name, (Markup.THEORY, name))), total) - object Semantic - { - object Info - { - def apply(pos: Position.T, semantic: Semantic): XML.Elem = - { + object Semantic { + object Info { + def apply(pos: Position.T, semantic: Semantic): XML.Elem = { val elem = semantic match { case No_Completion => XML.Elem(Markup(Markup.NO_COMPLETION, pos), Nil) @@ -178,8 +166,7 @@ info.info match { case XML.Elem(Markup(Markup.COMPLETION, _), body) => try { - val (total, names) = - { + val (total, names) = { import XML.Decode._ pair(int, list(pair(string, pair(string, string))))(body) } @@ -195,14 +182,13 @@ sealed abstract class Semantic case object No_Completion extends Semantic - case class Names(total: Int, names: List[(String, (String, String))]) extends Semantic - { + case class Names(total: Int, names: List[(String, (String, String))]) extends Semantic { def complete( range: Text.Range, history: Completion.History, unicode: Boolean, - original: String): Option[Completion.Result] = - { + original: String + ): Option[Completion.Result] = { def decode(s: String): String = if (unicode) Symbol.decode(s) else s val items = for { @@ -238,8 +224,7 @@ /* language context */ - object Language_Context - { + object Language_Context { val outer = Language_Context("", true, false) val inner = Language_Context(Markup.Language.UNKNOWN, true, false) val ML_outer = Language_Context(Markup.Language.ML, false, true) @@ -247,8 +232,7 @@ val SML_outer = Language_Context(Markup.Language.SML, false, false) } - sealed case class Language_Context(language: String, symbols: Boolean, antiquotes: Boolean) - { + sealed case class Language_Context(language: String, symbols: Boolean, antiquotes: Boolean) { def is_outer: Boolean = language == "" } @@ -264,8 +248,7 @@ /* word parsers */ - object Word_Parsers extends RegexParsers - { + object Word_Parsers extends RegexParsers { override val whiteSpace = "".r private val symboloid_regex: Regex = """\\([A-Za-z0-9_']+|<\^?[A-Za-z0-9_']+>?)""".r @@ -282,8 +265,7 @@ def is_word(s: CharSequence): Boolean = word_regex.pattern.matcher(s).matches def is_word_char(c: Char): Boolean = Symbol.is_ascii_letdig(c) || c == '.' - def read_symbol(in: CharSequence): Option[String] = - { + def read_symbol(in: CharSequence): Option[String] = { val reverse_in = new Library.Reverse(in) parse(reverse_symbol ^^ (_.reverse), reverse_in) match { case Success(result, _) => Some(result) @@ -291,8 +273,7 @@ } } - def read_word(in: CharSequence): Option[(String, String)] = - { + def read_word(in: CharSequence): Option[(String, String)] = { val reverse_in = new Library.Reverse(in) val parser = (reverse_symbol | reverse_symb | reverse_escape) ^^ (x => (x.reverse, "")) | @@ -343,8 +324,8 @@ words_lex: Scan.Lexicon = Scan.Lexicon.empty, words_map: Multi_Map[String, String] = Multi_Map.empty, abbrevs_lex: Scan.Lexicon = Scan.Lexicon.empty, - abbrevs_map: Multi_Map[String, (String, String)] = Multi_Map.empty) -{ + abbrevs_map: Multi_Map[String, (String, String)] = Multi_Map.empty +) { /* keywords */ private def is_symbol(name: String): Boolean = Symbol.symbols.defined(name) @@ -355,8 +336,7 @@ def add_keyword(keyword: String): Completion = new Completion(keywords + keyword, words_lex, words_map, abbrevs_lex, abbrevs_map) - def add_keywords(names: List[String]): Completion = - { + def add_keywords(names: List[String]): Completion = { val keywords1 = names.foldLeft(keywords) { case (ks, k) => if (ks(k)) ks else ks + k } if (keywords eq keywords1) this else new Completion(keywords1, words_lex, words_map, abbrevs_lex, abbrevs_map) @@ -365,11 +345,9 @@ /* symbols and abbrevs */ - def add_symbols: Completion = - { + def add_symbols: Completion = { val words = - Symbol.symbols.entries.flatMap(entry => - { + Symbol.symbols.entries.flatMap { entry => val sym = entry.symbol val word = "\\" + entry.name val seps = @@ -380,7 +358,7 @@ } List(sym -> sym, word -> sym) ::: seps.map(sep => word -> (sym + sep + "\\\u0007\\")) - }) + } new Completion( keywords, @@ -423,12 +401,11 @@ start: Text.Offset, text: CharSequence, caret: Int, - language_context: Completion.Language_Context): Option[Completion.Result] = - { + language_context: Completion.Language_Context + ): Option[Completion.Result] = { val length = text.length - val abbrevs_result = - { + val abbrevs_result = { val reverse_in = new Library.Reverse(text.subSequence(0, caret)) Scan.Parsers.parse(Scan.Parsers.literal(abbrevs_lex), reverse_in) match { case Scan.Parsers.Success(reverse_abbr, _) => diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/csv.scala --- a/src/Pure/General/csv.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/csv.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,10 +7,8 @@ package isabelle -object CSV -{ - def print_field(field: Any): String = - { +object CSV { + def print_field(field: Any): String = { val str = field.toString if (str.exists("\",\r\n".contains(_))) { (for (c <- str) yield { if (c == '"') "\"\"" else c.toString }).mkString("\"", "", "\"") @@ -18,13 +16,11 @@ else str } - sealed case class Record(fields: Any*) - { + sealed case class Record(fields: Any*) { override def toString: String = fields.iterator.map(print_field).mkString(",") } - sealed case class File(name: String, header: List[String], records: List[Record]) - { + sealed case class File(name: String, header: List[String], records: List[Record]) { override def toString: String = (Record(header:_*) :: records).mkString("\r\n") def file_name: String = name + ".csv" diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/date.scala --- a/src/Pure/General/date.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/date.scala Fri Apr 01 23:26:19 2022 +0200 @@ -15,14 +15,11 @@ import scala.annotation.tailrec -object Date -{ +object Date { /* format */ - object Format - { - def make(fmts: List[DateTimeFormatter], tune: String => String = s => s): Format = - { + object Format { + def make(fmts: List[DateTimeFormatter], tune: String => String = s => s): Format = { require(fmts.nonEmpty, "no date formats") new Format { @@ -41,8 +38,7 @@ val alt_date: Format = Format("uuuuMMdd") } - abstract class Format private - { + abstract class Format private { def apply(date: Date): String def parse(str: String): Date @@ -50,19 +46,20 @@ try { Some(parse(str)) } catch { case _: DateTimeParseException => None } } - object Formatter - { + object Formatter { def pattern(pat: String): DateTimeFormatter = DateTimeFormatter.ofPattern(pat) def variants(pats: List[String], locs: List[Locale] = Nil): List[DateTimeFormatter] = - pats.flatMap(pat => { + pats.flatMap { pat => val fmt = pattern(pat) if (locs.isEmpty) List(fmt) else locs.map(fmt.withLocale) - }) + } - @tailrec def try_variants(fmts: List[DateTimeFormatter], str: String, - last_exn: Option[DateTimeParseException] = None): TemporalAccessor = - { + @tailrec def try_variants( + fmts: List[DateTimeFormatter], + str: String, + last_exn: Option[DateTimeParseException] = None + ): TemporalAccessor = { fmts match { case Nil => throw last_exn.getOrElse(new DateTimeParseException("Failed to parse date", str, 0)) @@ -76,14 +73,12 @@ /* ordering */ - object Ordering extends scala.math.Ordering[Date] - { + object Ordering extends scala.math.Ordering[Date] { def compare(date1: Date, date2: Date): Int = date1.instant.compareTo(date2.instant) } - object Rev_Ordering extends scala.math.Ordering[Date] - { + object Rev_Ordering extends scala.math.Ordering[Date] { def compare(date1: Date, date2: Date): Int = date2.instant.compareTo(date1.instant) } @@ -104,8 +99,7 @@ def apply(t: Time, zone: ZoneId = timezone()): Date = instant(t.instant, zone) } -sealed case class Date(rep: ZonedDateTime) -{ +sealed case class Date(rep: ZonedDateTime) { def midnight: Date = new Date(ZonedDateTime.of(rep.toLocalDate, LocalTime.MIDNIGHT, rep.getZone)) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/exn.scala --- a/src/Pure/General/exn.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/exn.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,12 +7,10 @@ package isabelle -object Exn -{ +object Exn { /* user errors */ - class User_Error(message: String) extends RuntimeException(message) - { + class User_Error(message: String) extends RuntimeException(message) { override def equals(that: Any): Boolean = that match { case other: User_Error => message == other.getMessage @@ -23,8 +21,7 @@ override def toString: String = "\n" + Output.error_message_text(message) } - object ERROR - { + object ERROR { def apply(message: String): User_Error = new User_Error(message) def unapply(exn: Throwable): Option[String] = user_message(exn) } @@ -40,8 +37,7 @@ /* exceptions as values */ - sealed abstract class Result[A] - { + sealed abstract class Result[A] { def user_error: Result[A] = this match { case Exn(ERROR(msg)) => Exn(ERROR(msg)) @@ -86,10 +82,8 @@ try { Res(e) } catch { case exn: Throwable if !is_interrupt(exn) => Exn[A](exn) } - object Interrupt - { - object ERROR - { + object Interrupt { + object ERROR { def unapply(exn: Throwable): Option[String] = if (is_interrupt(exn)) Some(message(exn)) else user_message(exn) } @@ -106,16 +100,13 @@ /* message */ def user_message(exn: Throwable): Option[String] = - if (exn.isInstanceOf[User_Error] || exn.getClass == classOf[RuntimeException]) - { + if (exn.isInstanceOf[User_Error] || exn.getClass == classOf[RuntimeException]) { Some(proper_string(exn.getMessage) getOrElse "Error") } - else if (exn.isInstanceOf[java.sql.SQLException]) - { + else if (exn.isInstanceOf[java.sql.SQLException]) { Some(proper_string(exn.getMessage) getOrElse "SQL error") } - else if (exn.isInstanceOf[java.io.IOException]) - { + else if (exn.isInstanceOf[java.io.IOException]) { val msg = exn.getMessage Some(if (msg == null || msg == "") "I/O error" else "I/O error: " + msg) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/file.scala --- a/src/Pure/General/file.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/file.scala Fri Apr 01 23:26:19 2022 +0200 @@ -22,8 +22,7 @@ import scala.collection.mutable -object File -{ +object File { /* standard path (Cygwin or Posix) */ def standard_path(path: Path): String = path.expand.implode @@ -66,8 +65,7 @@ /* relative paths */ - def relative_path(base: Path, other: Path): Option[Path] = - { + def relative_path(base: Path, other: Path): Option[Path] = { val base_path = base.java_path val other_path = other.java_path if (other_path.startsWith(base_path)) @@ -95,8 +93,7 @@ /* directory content */ - def read_dir(dir: Path): List[String] = - { + def read_dir(dir: Path): List[String] = { if (!dir.is_dir) error("No such directory: " + dir.toString) val files = dir.file.listFiles if (files == null) Nil @@ -114,8 +111,8 @@ start: JFile, pred: JFile => Boolean = _ => true, include_dirs: Boolean = false, - follow_links: Boolean = false): List[JFile] = - { + follow_links: Boolean = false + ): List[JFile] = { val result = new mutable.ListBuffer[JFile] def check(file: JFile): Unit = if (pred(file)) result += file @@ -126,13 +123,17 @@ else EnumSet.noneOf(classOf[FileVisitOption]) Files.walkFileTree(start.toPath, options, Integer.MAX_VALUE, new SimpleFileVisitor[JPath] { - override def preVisitDirectory(path: JPath, attrs: BasicFileAttributes): FileVisitResult = - { + override def preVisitDirectory( + path: JPath, + attrs: BasicFileAttributes + ): FileVisitResult = { if (include_dirs) check(path.toFile) FileVisitResult.CONTINUE } - override def visitFile(path: JPath, attrs: BasicFileAttributes): FileVisitResult = - { + override def visitFile( + path: JPath, + attrs: BasicFileAttributes + ): FileVisitResult = { val file = path.toFile if (include_dirs || !file.isDirectory) check(file) FileVisitResult.CONTINUE @@ -151,8 +152,7 @@ def read(path: Path): String = read(path.file) - def read_stream(reader: BufferedReader): String = - { + def read_stream(reader: BufferedReader): String = { val output = new StringBuilder(100) var c = -1 while ({ c = reader.read; c != -1 }) output += c.toChar @@ -174,16 +174,14 @@ /* read lines */ - def read_line(reader: BufferedReader): Option[String] = - { + def read_line(reader: BufferedReader): Option[String] = { val line = try { reader.readLine} catch { case _: IOException => null } Option(line).map(Library.trim_line) } - def read_lines(reader: BufferedReader, progress: String => Unit): List[String] = - { + def read_lines(reader: BufferedReader, progress: String => Unit): List[String] = { val result = new mutable.ListBuffer[String] var line: Option[String] = None while ({ line = read_line(reader); line.isDefined }) { @@ -201,8 +199,10 @@ new BufferedWriter(new OutputStreamWriter(new FileOutputStream(file), UTF8.charset)) def write_file( - file: JFile, text: String, make_stream: OutputStream => OutputStream): Unit = - { + file: JFile, + text: String, + make_stream: OutputStream => OutputStream + ): Unit = { val stream = make_stream(new FileOutputStream(file)) using(new BufferedWriter(new OutputStreamWriter(stream, UTF8.charset)))(_.append(text)) } @@ -221,14 +221,12 @@ write_xz(path.file, text, options) def write_xz(path: Path, text: String): Unit = write_xz(path, text, XZ.options()) - def write_backup(path: Path, text: String): Unit = - { + def write_backup(path: Path, text: String): Unit = { if (path.is_file) Isabelle_System.move_file(path, path.backup) write(path, text) } - def write_backup2(path: Path, text: String): Unit = - { + def write_backup2(path: Path, text: String): Unit = { if (path.is_file) Isabelle_System.move_file(path, path.backup2) write(path, text) } @@ -245,8 +243,11 @@ /* change */ - def change(path: Path, init: Boolean = false, strict: Boolean = false)(f: String => String): Unit = - { + def change( + path: Path, + init: Boolean = false, + strict: Boolean = false + )(f: String => String): Unit = { if (!path.is_file && init) write(path, "") val x = read(path) val y = f(x) @@ -280,14 +281,12 @@ /* permissions */ - def is_executable(path: Path): Boolean = - { + def is_executable(path: Path): Boolean = { if (Platform.is_windows) Isabelle_System.bash("test -x " + bash_path(path)).check.ok else path.file.canExecute } - def set_executable(path: Path, flag: Boolean): Unit = - { + def set_executable(path: Path, flag: Boolean): Unit = { if (Platform.is_windows && flag) Isabelle_System.chmod("a+x", path) else if (Platform.is_windows) Isabelle_System.chmod("a-x", path) else path.file.setExecutable(flag, false) @@ -296,31 +295,26 @@ /* content */ - object Content - { + object Content { def apply(path: Path, content: Bytes): Content = new Content_Bytes(path, content) def apply(path: Path, content: String): Content = new Content_String(path, content) def apply(path: Path, content: XML.Body): Content_XML = new Content_XML(path, content) } - trait Content - { + trait Content { def path: Path def write(dir: Path): Unit } - final class Content_Bytes private[File](val path: Path, content: Bytes) extends Content - { + final class Content_Bytes private[File](val path: Path, content: Bytes) extends Content { def write(dir: Path): Unit = Bytes.write(dir + path, content) } - final class Content_String private[File](val path: Path, content: String) extends Content - { + final class Content_String private[File](val path: Path, content: String) extends Content { def write(dir: Path): Unit = File.write(dir + path, content) } - final class Content_XML private[File](val path: Path, content: XML.Body) - { + final class Content_XML private[File](val path: Path, content: XML.Body) { def output(out: XML.Body => String): Content_String = new Content_String(path, out(content)) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/file_watcher.scala --- a/src/Pure/General/file_watcher.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/file_watcher.scala Fri Apr 01 23:26:19 2022 +0200 @@ -16,8 +16,8 @@ import scala.jdk.CollectionConverters._ -class File_Watcher private[File_Watcher] // dummy template -{ +class File_Watcher private[File_Watcher] { + // dummy template def register(dir: JFile): Unit = {} def register_parent(file: JFile): Unit = {} def deregister(dir: JFile): Unit = {} @@ -25,10 +25,8 @@ def shutdown(): Unit = {} } -object File_Watcher -{ - val none: File_Watcher = new File_Watcher - { +object File_Watcher { + val none: File_Watcher = new File_Watcher { override def toString: String = "File_Watcher.none" } @@ -42,8 +40,7 @@ dirs: Map[JFile, WatchKey] = Map.empty, changed: Set[JFile] = Set.empty) - class Impl private[File_Watcher](handle: Set[JFile] => Unit, delay: Time) extends File_Watcher - { + class Impl private[File_Watcher](handle: Set[JFile] => Unit, delay: Time) extends File_Watcher { private val state = Synchronized(File_Watcher.State()) private val watcher = FileSystems.getDefault.newWatchService() @@ -62,8 +59,7 @@ st.copy(dirs = st.dirs + (dir -> key)) }) - override def register_parent(file: JFile): Unit = - { + override def register_parent(file: JFile): Unit = { val dir = file.getParentFile if (dir != null && dir.isDirectory) register(dir) } @@ -85,38 +81,35 @@ /* changed directory entries */ - private val delay_changed = Delay.last(delay) - { + private val delay_changed = Delay.last(delay) { val changed = state.change_result(st => (st.changed, st.copy(changed = Set.empty))) handle(changed) } - private val watcher_thread = Isabelle_Thread.fork(name = "file_watcher", daemon = true) - { + private val watcher_thread = Isabelle_Thread.fork(name = "file_watcher", daemon = true) { try { while (true) { val key = watcher.take val has_changed = - state.change_result(st => - { - val (remove, changed) = - st.dirs.collectFirst({ case (dir, key1) if key == key1 => dir }) match { - case Some(dir) => - val events: Iterable[WatchEvent[JPath]] = - key.pollEvents.asInstanceOf[JList[WatchEvent[JPath]]].asScala - val remove = if (key.reset) None else Some(dir) - val changed = - events.iterator.foldLeft(Set.empty[JFile]) { - case (set, event) => set + dir.toPath.resolve(event.context).toFile - } - (remove, changed) - case None => - key.pollEvents - key.reset - (None, Set.empty[JFile]) - } - (changed.nonEmpty, st.copy(dirs = st.dirs -- remove, changed = st.changed ++ changed)) - }) + state.change_result { st => + val (remove, changed) = + st.dirs.collectFirst({ case (dir, key1) if key == key1 => dir }) match { + case Some(dir) => + val events: Iterable[WatchEvent[JPath]] = + key.pollEvents.asInstanceOf[JList[WatchEvent[JPath]]].asScala + val remove = if (key.reset) None else Some(dir) + val changed = + events.iterator.foldLeft(Set.empty[JFile]) { + case (set, event) => set + dir.toPath.resolve(event.context).toFile + } + (remove, changed) + case None => + key.pollEvents + key.reset + (None, Set.empty[JFile]) + } + (changed.nonEmpty, st.copy(dirs = st.dirs -- remove, changed = st.changed ++ changed)) + } if (has_changed) delay_changed.invoke() } } @@ -126,8 +119,7 @@ /* shutdown */ - override def shutdown(): Unit = - { + override def shutdown(): Unit = { watcher_thread.interrupt() watcher_thread.join() delay_changed.revoke() diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/graph.scala --- a/src/Pure/General/graph.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/graph.scala Fri Apr 01 23:26:19 2022 +0200 @@ -11,27 +11,25 @@ import scala.annotation.tailrec -object Graph -{ - class Duplicate[Key](val key: Key) extends Exception - { +object Graph { + class Duplicate[Key](val key: Key) extends Exception { override def toString: String = "Graph.Duplicate(" + key.toString + ")" } - class Undefined[Key](val key: Key) extends Exception - { + class Undefined[Key](val key: Key) extends Exception { override def toString: String = "Graph.Undefined(" + key.toString + ")" } - class Cycles[Key](val cycles: List[List[Key]]) extends Exception - { + class Cycles[Key](val cycles: List[List[Key]]) extends Exception { override def toString: String = cycles.mkString("Graph.Cycles(", ",\n", ")") } def empty[Key, A](implicit ord: Ordering[Key]): Graph[Key, A] = new Graph[Key, A](SortedMap.empty(ord)) - def make[Key, A](entries: List[((Key, A), List[Key])], - symmetric: Boolean = false)(implicit ord: Ordering[Key]): Graph[Key, A] = - { + def make[Key, A]( + entries: List[((Key, A), List[Key])], + symmetric: Boolean = false)( + implicit ord: Ordering[Key] + ): Graph[Key, A] = { val graph1 = entries.foldLeft(empty[Key, A](ord)) { case (graph, ((x, info), _)) => graph.new_node(x, info) @@ -54,22 +52,21 @@ /* XML data representation */ def encode[Key, A](key: XML.Encode.T[Key], info: XML.Encode.T[A]): XML.Encode.T[Graph[Key, A]] = - (graph: Graph[Key, A]) => { + { (graph: Graph[Key, A]) => import XML.Encode._ list(pair(pair(key, info), list(key)))(graph.dest) } def decode[Key, A](key: XML.Decode.T[Key], info: XML.Decode.T[A])( implicit ord: Ordering[Key]): XML.Decode.T[Graph[Key, A]] = - (body: XML.Body) => { + { (body: XML.Body) => import XML.Decode._ make(list(pair(pair(key, info), list(key)))(body))(ord) } } -final class Graph[Key, A] private(rep: SortedMap[Key, (A, (SortedSet[Key], SortedSet[Key]))]) -{ +final class Graph[Key, A] private(rep: SortedMap[Key, (A, (SortedSet[Key], SortedSet[Key]))]) { type Keys = SortedSet[Key] type Entry = (A, (Keys, Keys)) @@ -121,8 +118,8 @@ def reachable_length( count: Key => Long, next: Key => Keys, - xs: List[Key]): Map[Key, Long] = - { + xs: List[Key] + ): Map[Key, Long] = { def reach(length: Long)(rs: Map[Key, Long], x: Key): Map[Key, Long] = if (rs.get(x) match { case Some(n) => n >= length case None => false }) rs else next(x).foldLeft(rs + (x -> length))(reach(length + count(x))) @@ -138,10 +135,9 @@ limit: Long, count: Key => Long, next: Key => Keys, - xs: List[Key]): Keys = - { - def reach(res: (Long, Keys), x: Key): (Long, Keys) = - { + xs: List[Key] + ): Keys = { + def reach(res: (Long, Keys), x: Key): (Long, Keys) = { val (n, rs) = res if (rs.contains(x)) res else next(x).foldLeft((n + count(x), rs + x))(reach) @@ -160,10 +156,12 @@ } /*reachable nodes with result in topological order (for acyclic graphs)*/ - private def reachable(next: Key => Keys, xs: List[Key], rev: Boolean = false): (List[List[Key]], Keys) = - { - def reach(x: Key, reached: (List[Key], Keys)): (List[Key], Keys) = - { + private def reachable( + next: Key => Keys, + xs: List[Key], + rev: Boolean = false + ): (List[List[Key]], Keys) = { + def reach(x: Key, reached: (List[Key], Keys)): (List[Key], Keys) = { val (rs, r_set) = reached if (r_set(x)) reached else { @@ -173,8 +171,7 @@ (x :: rs1, r_set1) } } - def reachs(reached: (List[List[Key]], Keys), x: Key): (List[List[Key]], Keys) = - { + def reachs(reached: (List[List[Key]], Keys), x: Key): (List[List[Key]], Keys) = { val (rss, r_set) = reached val (rs, r_set1) = reach(x, (Nil, r_set)) (rs :: rss, r_set1) @@ -218,8 +215,7 @@ /* node operations */ - def new_node(x: Key, info: A): Graph[Key, A] = - { + def new_node(x: Key, info: A): Graph[Key, A] = { if (defined(x)) throw new Graph.Duplicate(x) else new Graph[Key, A](rep + (x -> (info, (empty_keys, empty_keys)))) } @@ -235,8 +231,7 @@ map + (y -> (i, if (fst) (preds - x, succs) else (preds, succs - x))) } - def del_node(x: Key): Graph[Key, A] = - { + def del_node(x: Key): Graph[Key, A] = { val (preds, succs) = get_entry(x)._2 new Graph[Key, A]( succs.foldLeft(preds.foldLeft(rep - x)(del_adjacent(false, x)))(del_adjacent(true, x))) @@ -271,8 +266,7 @@ /* irreducible paths -- Hasse diagram */ - private def irreducible_preds(x_set: Keys, path: List[Key], z: Key): List[Key] = - { + private def irreducible_preds(x_set: Keys, path: List[Key], z: Key): List[Key] = { def red(x: Key)(x1: Key) = is_edge(x, x1) && x1 != z @tailrec def irreds(xs0: List[Key], xs1: List[Key]): List[Key] = xs0 match { @@ -286,8 +280,7 @@ irreds(imm_preds(z).toList, Nil) } - def irreducible_paths(x: Key, y: Key): List[List[Key]] = - { + def irreducible_paths(x: Key, y: Key): List[List[Key]] = { val (_, x_set) = reachable(imm_succs, List(x)) def paths(path: List[Key])(ps: List[List[Key]], z: Key): List[List[Key]] = if (x == z) (z :: path) :: ps @@ -298,8 +291,7 @@ /* transitive closure and reduction */ - private def transitive_step(z: Key): Graph[Key, A] = - { + private def transitive_step(z: Key): Graph[Key, A] = { val (preds, succs) = get_entry(z)._2 var graph = this for (x <- preds; y <- succs) graph = graph.add_edge(x, y) @@ -308,8 +300,7 @@ def transitive_closure: Graph[Key, A] = keys_iterator.foldLeft(this)(_.transitive_step(_)) - def transitive_reduction_acyclic: Graph[Key, A] = - { + def transitive_reduction_acyclic: Graph[Key, A] = { val trans = this.transitive_closure if (trans.iterator.exists({ case (x, (_, (_, succs))) => succs.contains(x) })) error("Cyclic graph") diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/graph_display.scala --- a/src/Pure/General/graph_display.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/graph_display.scala Fri Apr 01 23:26:19 2022 +0200 @@ -6,8 +6,7 @@ package isabelle -object Graph_Display -{ +object Graph_Display { /* graph entries */ type Entry = ((String, (String, XML.Body)), List[String]) // ident, name, content, parents @@ -15,12 +14,10 @@ /* graph structure */ - object Node - { + object Node { val dummy: Node = Node("", "") - object Ordering extends scala.math.Ordering[Node] - { + object Ordering extends scala.math.Ordering[Node] { def compare(node1: Node, node2: Node): Int = node1.name compare node2.name match { case 0 => node1.ident compare node2.ident @@ -28,8 +25,7 @@ } } } - sealed case class Node(name: String, ident: String) - { + sealed case class Node(name: String, ident: String) { def is_dummy: Boolean = this == Node.dummy override def toString: String = name } @@ -40,8 +36,7 @@ val empty_graph: Graph = isabelle.Graph.empty(Node.Ordering) - def build_graph(entries: List[Entry]): Graph = - { + def build_graph(entries: List[Entry]): Graph = { val node = entries.foldLeft(Map.empty[String, Node]) { case (m, ((ident, (name, _)), _)) => m + (ident -> Node(name, ident)) @@ -65,8 +60,8 @@ def make_graph[A]( graph: isabelle.Graph[String, A], isolated: Boolean = false, - name: (String, A) => String = (x: String, a: A) => x): Graph = - { + name: (String, A) => String = (x: String, a: A) => x + ): Graph = { val entries = (for { (x, (a, (ps, _))) <- graph.iterator if isolated || !graph.is_isolated(x) } yield ((x, (name(x, a), Nil)), ps.toList)).toList diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/graphics_file.scala --- a/src/Pure/General/graphics_file.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/graphics_file.scala Fri Apr 01 23:26:19 2022 +0200 @@ -18,13 +18,16 @@ import com.lowagie.text.pdf.{PdfWriter, BaseFont, FontMapper, DefaultFontMapper} -object Graphics_File -{ +object Graphics_File { /* PNG */ def write_png( - file: JFile, paint: Graphics2D => Unit, width: Int, height: Int, dpi: Int = 72): Unit = - { + file: JFile, + paint: Graphics2D => Unit, + width: Int, + height: Int, + dpi: Int = 72 + ): Unit = { val scale = dpi / 72.0f val w = (width * scale).round val h = (height * scale).round @@ -42,8 +45,7 @@ /* PDF */ - private def font_mapper(): FontMapper = - { + private def font_mapper(): FontMapper = { val mapper = new DefaultFontMapper for (entry <- Isabelle_Fonts.fonts()) { val params = new DefaultFontMapper.BaseFontParameters(File.platform_path(entry.path)) @@ -55,12 +57,10 @@ mapper } - def write_pdf(file: JFile, paint: Graphics2D => Unit, width: Int, height: Int): Unit = - { + def write_pdf(file: JFile, paint: Graphics2D => Unit, width: Int, height: Int): Unit = { import com.lowagie.text.{Document, Rectangle} - using(new BufferedOutputStream(new FileOutputStream(file)))(out => - { + using(new BufferedOutputStream(new FileOutputStream(file))) { out => val document = new Document() try { document.setPageSize(new Rectangle(width.toFloat, height.toFloat)) @@ -77,7 +77,7 @@ cb.addTemplate(tp, 1, 0, 0, 1, 0, 0) } finally { document.close() } - }) + } } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/http.scala --- a/src/Pure/General/http.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/http.scala Fri Apr 01 23:26:19 2022 +0200 @@ -13,12 +13,10 @@ import com.sun.net.httpserver.{HttpExchange, HttpHandler, HttpServer} -object HTTP -{ +object HTTP { /** content **/ - object Content - { + object Content { val mime_type_bytes: String = "application/octet-stream" val mime_type_text: String = "text/plain; charset=utf-8" val mime_type_html: String = "text/html; charset=utf-8" @@ -34,8 +32,7 @@ elapsed_time: Time = Time.zero): Content = new Content(bytes, file_name, mime_type, encoding, elapsed_time) - def read(file: JFile): Content = - { + def read(file: JFile): Content = { val bytes = Bytes.read(file) val file_name = file.getName val mime_type = Option(Files.probeContentType(file.toPath)).getOrElse(default_mime_type) @@ -50,8 +47,8 @@ val file_name: String, val mime_type: String, val encoding: String, - val elapsed_time: Time) - { + val elapsed_time: Time + ) { def text: String = new String(bytes.array, encoding) def json: JSON.T = JSON.parse(text) } @@ -62,14 +59,14 @@ val NEWLINE: String = "\r\n" - object Client - { + object Client { val default_timeout: Time = Time.seconds(180) - def open_connection(url: URL, + def open_connection( + url: URL, timeout: Time = default_timeout, - user_agent: String = ""): HttpURLConnection = - { + user_agent: String = "" + ): HttpURLConnection = { url.openConnection match { case connection: HttpURLConnection => if (0 < timeout.ms && timeout.ms <= Integer.MAX_VALUE) { @@ -83,13 +80,11 @@ } } - def get_content(connection: HttpURLConnection): Content = - { + def get_content(connection: HttpURLConnection): Content = { val Charset = """.*\bcharset="?([\S^"]+)"?.*""".r val start = Time.now() - using(connection.getInputStream)(stream => - { + using(connection.getInputStream) { stream => val bytes = Bytes.read_stream(stream, hint = connection.getContentLength) val stop = Time.now() @@ -103,16 +98,18 @@ } Content(bytes, file_name = file_name, mime_type = mime_type, encoding = encoding, elapsed_time = stop - start) - }) + } } def get(url: URL, timeout: Time = default_timeout, user_agent: String = ""): Content = get_content(open_connection(url, timeout = timeout, user_agent = user_agent)) - def post(url: URL, parameters: List[(String, Any)], + def post( + url: URL, + parameters: List[(String, Any)], timeout: Time = default_timeout, - user_agent: String = ""): Content = - { + user_agent: String = "" + ): Content = { val connection = open_connection(url, timeout = timeout, user_agent = user_agent) connection.setRequestMethod("POST") connection.setDoOutput(true) @@ -121,21 +118,18 @@ connection.setRequestProperty( "Content-Type", "multipart/form-data; boundary=" + quote(boundary)) - using(connection.getOutputStream)(out => - { + using(connection.getOutputStream) { out => def output(s: String): Unit = out.write(UTF8.bytes(s)) def output_newline(n: Int = 1): Unit = (1 to n).foreach(_ => output(NEWLINE)) def output_boundary(end: Boolean = false): Unit = output("--" + boundary + (if (end) "--" else "") + NEWLINE) def output_name(name: String): Unit = output("Content-Disposition: form-data; name=" + quote(name)) - def output_value(value: Any): Unit = - { + def output_value(value: Any): Unit = { output_newline(2) output(value.toString) } - def output_content(content: Content): Unit = - { + def output_content(content: Content): Unit = { proper_string(content.file_name).foreach(s => output("; filename=" + quote(s))) output_newline() proper_string(content.mime_type).foreach(s => output("Content-Type: " + s)) @@ -158,7 +152,7 @@ } output_boundary(end = true) out.flush() - }) + } get_content(connection) } @@ -177,8 +171,8 @@ val server_name: String, val service_name: String, val uri: URI, - val input: Bytes) - { + val input: Bytes + ) { def home: String = url_path(server_name, service_name) def root: String = home + "/" def query: String = home + "?" @@ -208,8 +202,7 @@ /* response */ - object Response - { + object Response { def apply( bytes: Bytes = Bytes.empty, content_type: String = Content.mime_type_bytes): Response = @@ -224,12 +217,10 @@ def read(path: Path): Response = content(Content.read(path)) } - class Response private[HTTP](val output: Bytes, val content_type: String) - { + class Response private[HTTP](val output: Bytes, val content_type: String) { override def toString: String = output.toString - def write(http: HttpExchange, code: Int): Unit = - { + def write(http: HttpExchange, code: Int): Unit = { http.getResponseHeaders.set("Content-Type", content_type) http.sendResponseHeaders(code, output.length.toLong) using(http.getResponseBody)(output.write_stream) @@ -239,8 +230,7 @@ /* service */ - abstract class Service(val name: String, method: String = "GET") - { + abstract class Service(val name: String, method: String = "GET") { override def toString: String = name def apply(request: Request): Option[Response] @@ -248,7 +238,7 @@ def context(server_name: String): String = proper_string(url_path(server_name, name)).getOrElse("/") - def handler(server_name: String): HttpHandler = (http: HttpExchange) => { + def handler(server_name: String): HttpHandler = { (http: HttpExchange) => val uri = http.getRequestURI val input = using(http.getRequestBody)(Bytes.read_stream(_)) if (http.getRequestMethod == method) { @@ -270,8 +260,7 @@ /* server */ - class Server private[HTTP](val name: String, val http_server: HttpServer) - { + class Server private[HTTP](val name: String, val http_server: HttpServer) { def += (service: Service): Unit = http_server.createContext(service.context(name), service.handler(name)) def -= (service: Service): Unit = @@ -288,8 +277,8 @@ def server( name: String = UUID.random().toString, - services: List[Service] = isabelle_services): Server = - { + services: List[Service] = isabelle_services + ): Server = { val http_server = HttpServer.create(new InetSocketAddress(isabelle.Server.localhost, 0), 0) http_server.setExecutor(null) @@ -310,8 +299,7 @@ object Welcome_Service extends Welcome() - class Welcome(name: String = "") extends Service(name) - { + class Welcome(name: String = "") extends Service(name) { def apply(request: Request): Option[Response] = if (request.toplevel) { Some(Response.text("Welcome to " + Isabelle_System.identification())) @@ -324,8 +312,7 @@ object Fonts_Service extends Fonts() - class Fonts(name: String = "fonts") extends Service(name) - { + class Fonts(name: String = "fonts") extends Service(name) { private lazy val html_fonts: List[Isabelle_Fonts.Entry] = Isabelle_Fonts.fonts(hidden = true) @@ -346,8 +333,7 @@ object PDFjs_Service extends PDFjs() - class PDFjs(name: String = "pdfjs") extends Service(name) - { + class PDFjs(name: String = "pdfjs") extends Service(name) { def apply(request: Request): Option[Response] = for { p <- request.uri_path @@ -361,8 +347,7 @@ object Docs_Service extends Docs() - class Docs(name: String = "docs") extends PDFjs(name) - { + class Docs(name: String = "docs") extends PDFjs(name) { private val doc_contents = isabelle.Doc.main_contents() // example: .../docs/web/viewer.html?file=system.pdf diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/json.scala --- a/src/Pure/General/json.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/json.scala Fri Apr 01 23:26:19 2022 +0200 @@ -15,13 +15,11 @@ import scala.util.parsing.input.CharArrayReader.EofCh -object JSON -{ +object JSON { type T = Any type S = String - object Object - { + object Object { type Entry = (String, JSON.T) type T = Map[String, JSON.T] @@ -40,13 +38,11 @@ /* lexer */ - object Kind extends Enumeration - { + object Kind extends Enumeration { val KEYWORD, STRING, NUMBER, ERROR = Value } - sealed case class Token(kind: Kind.Value, text: String) - { + sealed case class Token(kind: Kind.Value, text: String) { def is_keyword: Boolean = kind == Kind.KEYWORD def is_keyword(name: String): Boolean = kind == Kind.KEYWORD && text == name def is_string: Boolean = kind == Kind.STRING @@ -54,8 +50,7 @@ def is_error: Boolean = kind == Kind.ERROR } - object Lexer extends Scanners with Scan.Parsers - { + object Lexer extends Scanners with Scan.Parsers { override type Elem = Char type Token = JSON.Token @@ -129,8 +124,7 @@ /* parser */ - trait Parser extends Parsers - { + trait Parser extends Parsers { type Elem = Token def $$$(name: String): Parser[Token] = elem(name, _.is_keyword(name)) @@ -149,8 +143,7 @@ json_object | (json_array | (number | (string | ($$$("true") ^^^ true | ($$$("false") ^^^ false | ($$$("null") ^^^ null)))))) - def parse(input: CharSequence, strict: Boolean): T = - { + def parse(input: CharSequence, strict: Boolean): T = { val scanner = new Lexer.Scanner(Scan.char_reader(input)) phrase(if (strict) json_object | json_array else json_value)(scanner) match { case Success(json, _) => json @@ -166,20 +159,17 @@ def parse(s: S, strict: Boolean = true): T = Parser.parse(s, strict) - object Format - { + object Format { def unapply(s: S): Option[T] = try { Some(parse(s, strict = false)) } catch { case ERROR(_) => None } def apply_lines(json: List[T]): S = json.map(apply).mkString("[", ",\n", "]") - def apply(json: T): S = - { + def apply(json: T): S = { val result = new StringBuilder - def string(s: String): Unit = - { + def string(s: String): Unit = { result += '"' result ++= s.iterator.map { @@ -197,8 +187,7 @@ result += '"' } - def array(list: List[T]): Unit = - { + def array(list: List[T]): Unit = { result += '[' Library.separate(None, list.map(Some(_))).foreach({ case None => result += ',' @@ -207,8 +196,7 @@ result += ']' } - def object_(obj: Object.T): Unit = - { + def object_(obj: Object.T): Unit = { result += '{' Library.separate(None, obj.toList.map(Some(_))).foreach({ case None => result += ',' @@ -220,8 +208,7 @@ result += '}' } - def json_format(x: T): Unit = - { + def json_format(x: T): Unit = { x match { case null => result ++= "null" case _: Int | _: Long | _: Boolean => result ++= x.toString @@ -243,10 +230,8 @@ /* typed values */ - object Value - { - object UUID - { + object Value { + object UUID { def unapply(json: T): Option[isabelle.UUID.T] = json match { case x: java.lang.String => isabelle.UUID.unapply(x) @@ -254,8 +239,7 @@ } } - object String - { + object String { def unapply(json: T): Option[java.lang.String] = json match { case x: java.lang.String => Some(x) @@ -263,8 +247,7 @@ } } - object String0 - { + object String0 { def unapply(json: T): Option[java.lang.String] = json match { case null => Some("") @@ -273,8 +256,7 @@ } } - object Double - { + object Double { def unapply(json: T): Option[scala.Double] = json match { case x: scala.Double => Some(x) @@ -284,8 +266,7 @@ } } - object Long - { + object Long { def unapply(json: T): Option[scala.Long] = json match { case x: scala.Double if x.toLong.toDouble == x => Some(x.toLong) @@ -295,8 +276,7 @@ } } - object Int - { + object Int { def unapply(json: T): Option[scala.Int] = json match { case x: scala.Double if x.toInt.toDouble == x => Some(x.toInt) @@ -306,8 +286,7 @@ } } - object Boolean - { + object Boolean { def unapply(json: T): Option[scala.Boolean] = json match { case x: scala.Boolean => Some(x) @@ -315,8 +294,7 @@ } } - object List - { + object List { def unapply[A](json: T, unapply: T => Option[A]): Option[List[A]] = json match { case xs: List[T] => @@ -326,8 +304,7 @@ } } - object Seconds - { + object Seconds { def unapply(json: T): Option[Time] = Double.unapply(json).map(Time.seconds) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/json_api.scala --- a/src/Pure/General/json_api.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/json_api.scala Fri Apr 01 23:26:19 2022 +0200 @@ -9,15 +9,13 @@ import java.net.URL -object JSON_API -{ +object JSON_API { val mime_type = "application/vnd.api+json" def api_error(msg: String): Nothing = error("JSON API error: " + msg) def api_errors(msgs: List[String]): Nothing = error(("JSON API errors:" :: msgs).mkString("\n ")) - class Service(val url: URL) - { + class Service(val url: URL) { override def toString: String = url.toString def get(route: String): HTTP.Content = @@ -27,8 +25,7 @@ Root(get(if (route.isEmpty) "" else "/" + route).json) } - sealed case class Root(json: JSON.T) - { + sealed case class Root(json: JSON.T) { def get_links: Option[Links] = JSON.value(json, "links").map(Links) def get_next: Option[Service] = get_links.flatMap(_.get_next) @@ -48,8 +45,7 @@ def check_resource: Resource = check_data.resource def check_resources: List[Resource] = check_data.resources - def iterator: Iterator[Root] = - { + def iterator: Iterator[Root] = { val init = Some(this) new Iterator[Root] { private var state: Option[Root] = init @@ -70,8 +66,7 @@ override def toString: String = "Root(" + (if (ok) "ok" else "errors") + ")" } - sealed case class Links(json: JSON.T) - { + sealed case class Links(json: JSON.T) { def get_next: Option[Service] = for { JSON.Value.String(next) <- JSON.value(json, "next") @@ -82,8 +77,7 @@ if (get_next.isEmpty) "Links()" else "Links(next)" } - sealed case class Data(data: JSON.T, included: List[Resource]) - { + sealed case class Data(data: JSON.T, included: List[Resource]) { def is_multi: Boolean = JSON.Value.List.unapply(data, Some(_)).nonEmpty def resource: Resource = @@ -98,13 +92,11 @@ if (is_multi) "Data(resources)" else "Data(resource)" } - object Resource - { + object Resource { def some(json: JSON.T): Some[Resource] = Some(Resource(json)) } - sealed case class Resource(json: JSON.T) - { + sealed case class Resource(json: JSON.T) { val id: JSON.T = JSON.value(json, "id") getOrElse api_error("missing id") val type_ : JSON.T = JSON.value(json, "type") getOrElse api_error("missing type") val fields: JSON.Object.T = diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/linear_set.scala --- a/src/Pure/General/linear_set.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/linear_set.scala Fri Apr 01 23:26:19 2022 +0200 @@ -13,8 +13,7 @@ import scala.collection.{IterableFactory, IterableFactoryDefaults} -object Linear_Set extends IterableFactory[Linear_Set] -{ +object Linear_Set extends IterableFactory[Linear_Set] { private val empty_val: Linear_Set[Nothing] = new Linear_Set[Nothing](None, None, Map(), Map()) override def empty[A]: Linear_Set[A] = empty_val.asInstanceOf[Linear_Set[A]] @@ -22,8 +21,7 @@ entries.iterator.foldLeft(empty[A])(_.incl(_)) override def newBuilder[A]: mutable.Builder[A, Linear_Set[A]] = new Builder[A] - private class Builder[A] extends mutable.Builder[A, Linear_Set[A]] - { + private class Builder[A] extends mutable.Builder[A, Linear_Set[A]] { private var res = empty[A] override def clear(): Unit = { res = empty[A] } override def addOne(elem: A): this.type = { res = res.incl(elem); this } @@ -37,11 +35,12 @@ final class Linear_Set[A] private( - start: Option[A], end: Option[A], val nexts: Map[A, A], prevs: Map[A, A]) + start: Option[A], + end: Option[A], + val nexts: Map[A, A], prevs: Map[A, A]) extends Iterable[A] with SetOps[A, Linear_Set, Linear_Set[A]] - with IterableFactoryDefaults[A, Linear_Set] -{ + with IterableFactoryDefaults[A, Linear_Set] { /* relative addressing */ def next(elem: A): Option[A] = diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/logger.scala --- a/src/Pure/General/logger.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/logger.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Logger -{ +object Logger { def make(log_file: Option[Path]): Logger = log_file match { case Some(file) => new File_Logger(file) case None => No_Logger } @@ -16,21 +15,18 @@ new Logger { def apply(msg: => String): Unit = progress.echo(msg) } } -trait Logger -{ +trait Logger { def apply(msg: => String): Unit def timeit[A](message: String = "", enabled: Boolean = true)(e: => A): A = Timing.timeit(message, enabled, apply(_))(e) } -object No_Logger extends Logger -{ +object No_Logger extends Logger { def apply(msg: => String): Unit = {} } -class File_Logger(path: Path) extends Logger -{ +class File_Logger(path: Path) extends Logger { def apply(msg: => String): Unit = synchronized { File.append(path, msg + "\n") } override def toString: String = path.toString diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/long_name.scala --- a/src/Pure/General/long_name.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/long_name.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Long_Name -{ +object Long_Name { val separator = "." val separator_char = '.' diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/mailman.scala --- a/src/Pure/General/mailman.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/mailman.scala Fri Apr 01 23:26:19 2022 +0200 @@ -14,8 +14,7 @@ import scala.util.matching.Regex.Match -object Mailman -{ +object Mailman { /* mailing list messages */ def is_address(s: String): Boolean = @@ -215,8 +214,8 @@ title: String, author_info: List[String], body: String, - tags: List[String]) - { + tags: List[String] + ) { if (author_info.isEmpty || author_info.exists(_.isEmpty)) { error("Bad author information in " + quote(name)) } @@ -224,12 +223,10 @@ override def toString: String = name } - object Messages - { + object Messages { type Graph = isabelle.Graph[String, Message] - def apply(msgs: List[Message]): Messages = - { + def apply(msgs: List[Message]): Messages = { def make_node(g: Graph, node: String, msg: Message): Graph = if (g.defined(node) && Date.Ordering.compare(g.get_node(node).date, msg.date) >= 0) g else g.default_node(node, msg) @@ -248,8 +245,7 @@ })) } - def find(dir: Path): Messages = - { + def find(dir: Path): Messages = { val msgs = for { archive <- List(Isabelle_Users, Isabelle_Dev) @@ -258,8 +254,7 @@ Messages(msgs) } - sealed case class Cluster(author_info: List[String]) - { + sealed case class Cluster(author_info: List[String]) { val (addresses, names) = author_info.partition(is_address) val name: String = @@ -277,20 +272,17 @@ def unique: Boolean = names.length == 1 && addresses.length == 1 def multi: Boolean = names.length > 1 || addresses.length > 1 - def print: String = - { + def print: String = { val entries = names ::: (if (addresses.isEmpty) Nil else List("")) ::: addresses entries.mkString("\n * ", "\n ", "") } } } - class Messages private(val sorted: List[Message], val graph: Messages.Graph) - { + class Messages private(val sorted: List[Message], val graph: Messages.Graph) { override def toString: String = "Messages(" + sorted.size + ")" - object Node_Ordering extends scala.math.Ordering[String] - { + object Node_Ordering extends scala.math.Ordering[String] { override def compare(a: String, b: String): Int = Date.Rev_Ordering.compare(graph.get_node(a).date, graph.get_node(b).date) } @@ -303,8 +295,7 @@ def get_address(msg: Message): String = get_cluster(msg).get_address getOrElse error("No author address for " + quote(msg.name)) - def check(check_all: Boolean, check_multi: Boolean = false): Unit = - { + def check(check_all: Boolean, check_multi: Boolean = false): Unit = { val clusters = sorted.map(get_cluster).distinct.sortBy(_.name_lowercase) if (check_all) { @@ -330,16 +321,15 @@ abstract class Archive( url: URL, name: String = "", - tag: String = "") - { + tag: String = "" + ) { def message_regex: Regex def message_content(name: String, lines: List[String]): Message def message_match(lines: List[String], re: Regex): Option[Match] = lines.flatMap(re.findFirstMatchIn(_)).headOption - def make_name(str: String): String = - { + def make_name(str: String): String = { val s = str.trim.replaceAll("""\s+""", " ").replaceAll(" at ", "@") .replace("mailbroy.informatik.tu-muenchen.de", "in.tum.de") @@ -355,8 +345,7 @@ private val main_html = Url.read(main_url) - val list_name: String = - { + val list_name: String = { val title = """The ([^</>]*) Archives""".r.findFirstMatchIn(main_html).map(_.group(1)) (proper_string(name) orElse title).getOrElse(error("Failed to determine mailing list name")) @@ -379,8 +368,7 @@ msg <- message_regex.findAllMatchIn(html).map(_.group(1)) } yield href + "/" + msg).toList - def get(target_dir: Path, href: String, progress: Progress = new Progress): Option[Path] = - { + def get(target_dir: Path, href: String, progress: Progress = new Progress): Option[Path] = { val dir = target_dir + Path.basic(list_name) val path = dir + Path.explode(href) val url = new URL(main_url, href) @@ -413,8 +401,7 @@ download_text(target_dir, progress = progress) ::: download_msg(target_dir, progress = progress) - def make_title(str: String): String = - { + def make_title(str: String): String = { val Trim1 = ("""\s*\Q[""" + list_tag + """]\E\s*(.*)""").r val Trim2 = """(?i:(?:re|fw|fwd)\s*:\s*)(.*)""".r val Trim3 = """\[\s*(.*?)\s*\]""".r @@ -431,8 +418,7 @@ def get_messages(): List[Message] = for (href <- hrefs_msg) yield message_content(href, split_lines(read_text(href))) - def find_messages(dir: Path): List[Message] = - { + def find_messages(dir: Path): List[Message] = { for { file <- File.find_files(dir.file, file => file.getName.endsWith(".html")) rel_path <- File.relative_path(dir, File.path(file)) @@ -444,10 +430,8 @@ } } - private class Message_Chunk(bg: String = "", en: String = "") - { - def unapply(lines: List[String]): Option[List[String]] = - { + private class Message_Chunk(bg: String = "", en: String = "") { + def unapply(lines: List[String]): Option[List[String]] = { val res1 = if (bg.isEmpty) Some(lines) else { @@ -479,8 +463,8 @@ object Isabelle_Users extends Archive( Url("https://lists.cam.ac.uk/pipermail/cl-isabelle-users"), - name = "isabelle-users", tag = "isabelle") - { + name = "isabelle-users", tag = "isabelle" + ) { override def message_regex: Regex = """
  • """.r private object Head extends @@ -489,8 +473,7 @@ private object Body extends Message_Chunk(bg = "", en = "") - private object Date_Format - { + private object Date_Format { private val Trim1 = """\w+,\s*(.*)""".r private val Trim2 = """(.*?)\s*\(.*\)""".r private val Format = @@ -499,8 +482,7 @@ "d MMM uuuu H:m:s z", "d MMM yy H:m:s Z", "d MMM yy H:m:s z") - def unapply(s: String): Option[Date] = - { + def unapply(s: String): Option[Date] = { val s0 = s.replaceAll("""\s+""", " ") val s1 = s0 match { @@ -516,14 +498,12 @@ } } - override def make_name(str: String): String = - { + override def make_name(str: String): String = { val s = Library.perhaps_unsuffix(" via Cl-isabelle-users", super.make_name(str)) if (s == "cl-isabelle-users@lists.cam.ac.uk") "" else s } - object Address - { + object Address { private def anchor(s: String): String = """""" + s + """""" private def angl(s: String): String = """<""" + s + """>""" private def quot(s: String): String = """"""" + s + """"""" @@ -559,8 +539,7 @@ } } - override def message_content(name: String, lines: List[String]): Message = - { + override def message_content(name: String, lines: List[String]): Message = { def err(msg: String = ""): Nothing = error("Malformed message: " + name + (if (msg.isEmpty) "" else "\n" + msg)) @@ -604,21 +583,18 @@ /* isabelle-dev mailing list */ - object Isabelle_Dev extends Archive(Url("https://mailmanbroy.in.tum.de/pipermail/isabelle-dev")) - { + object Isabelle_Dev extends Archive(Url("https://mailmanbroy.in.tum.de/pipermail/isabelle-dev")) { override def message_regex: Regex = """
  • """.r private object Head extends Message_Chunk(en = "") private object Body extends Message_Chunk(bg = "", en = "") - private object Date_Format - { + private object Date_Format { val Format = Date.Format("E MMM d H:m:s z uuuu") def unapply(s: String): Option[Date] = Format.unapply(s.replaceAll("""\s+""", " ")) } - override def message_content(name: String, lines: List[String]): Message = - { + override def message_content(name: String, lines: List[String]): Message = { def err(msg: String = ""): Nothing = error("Malformed message: " + name + (if (msg.isEmpty) "" else "\n" + msg)) @@ -633,8 +609,7 @@ case None => err("Missing Date") } - val (title, author_address) = - { + val (title, author_address) = { message_match(head, """TITLE="([^"]+)">(.*)""".r) match { case Some(m) => (make_title(HTML.input(m.group(1))), make_name(HTML.input(m.group(2)))) case None => err("Missing TITLE") diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/mercurial.scala --- a/src/Pure/General/mercurial.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/mercurial.scala Fri Apr 01 23:26:19 2022 +0200 @@ -12,19 +12,16 @@ import scala.collection.mutable -object Mercurial -{ +object Mercurial { type Graph = isabelle.Graph[String, Unit] /* HTTP server */ - object Server - { + object Server { def apply(root: String): Server = new Server(root) - def start(root: Path): Server = - { + def start(root: Path): Server = { val hg = repository(root) val server_process = Future.promise[Bash.Process] @@ -46,8 +43,7 @@ } } - class Server private(val root: String) extends AutoCloseable - { + class Server private(val root: String) extends AutoCloseable { override def toString: String = root def close(): Unit = () @@ -70,17 +66,15 @@ def download_archive(rev: String = "tip", progress: Progress = new Progress): HTTP.Content = Isabelle_System.download(archive(rev = rev), progress = progress) - def download_dir(dir: Path, rev: String = "tip", progress: Progress = new Progress): Unit = - { + def download_dir(dir: Path, rev: String = "tip", progress: Progress = new Progress): Unit = { Isabelle_System.new_directory(dir) - Isabelle_System.with_tmp_file("rev", ext = ".tar.gz")(archive_path => - { + Isabelle_System.with_tmp_file("rev", ext = ".tar.gz") { archive_path => val content = download_archive(rev = rev, progress = progress) Bytes.write(archive_path, content.bytes) progress.echo("Unpacking " + rev + ".tar.gz") Isabelle_System.gnutar("-xzf " + File.bash_path(archive_path), dir = dir, original_owner = true, strip = 1).check - }) + } } } @@ -100,14 +94,12 @@ private val Archive_Node = """^node: (\S{12}).*$""".r private val Archive_Tag = """^tag: (\S+).*$""".r - sealed case class Archive_Info(lines: List[String]) - { + sealed case class Archive_Info(lines: List[String]) { def id: Option[String] = lines.collectFirst({ case Archive_Node(a) => a }) def tags: List[String] = for (Archive_Tag(tag) <- lines if tag != "tip") yield tag } - def archive_info(root: Path): Option[Archive_Info] = - { + def archive_info(root: Path): Option[Archive_Info] = { val path = root + Path.explode(".hg_archival.txt") if (path.is_file) Some(Archive_Info(Library.trim_split_lines(File.read(path)))) else None } @@ -125,15 +117,13 @@ ssh.is_dir(root + Path.explode(".hg")) && new Repository(root, ssh).command("root").ok - def repository(root: Path, ssh: SSH.System = SSH.Local): Repository = - { + def repository(root: Path, ssh: SSH.System = SSH.Local): Repository = { val hg = new Repository(root, ssh) hg.command("root").check hg } - def find_repository(start: Path, ssh: SSH.System = SSH.Local): Option[Repository] = - { + def find_repository(start: Path, ssh: SSH.System = SSH.Local): Option[Repository] = { @tailrec def find(root: Path): Option[Repository] = if (is_repository(root, ssh)) Some(repository(root, ssh = ssh)) else if (root.is_root) None @@ -142,9 +132,12 @@ find(ssh.expand_path(start)) } - private def make_repository(root: Path, cmd: String, args: String, ssh: SSH.System = SSH.Local) - : Repository = - { + private def make_repository( + root: Path, + cmd: String, + args: String, + ssh: SSH.System = SSH.Local + ) : Repository = { val hg = new Repository(root, ssh) ssh.make_directory(hg.root.dir) hg.command(cmd, args, repository = false).check @@ -159,14 +152,12 @@ make_repository(root, "clone", options + " " + Bash.string(source) + " " + ssh.bash_path(root) + opt_rev(rev), ssh = ssh) - def setup_repository(source: String, root: Path, ssh: SSH.System = SSH.Local): Repository = - { + def setup_repository(source: String, root: Path, ssh: SSH.System = SSH.Local): Repository = { if (ssh.is_dir(root)) { val hg = repository(root, ssh = ssh); hg.pull(remote = source); hg } else clone_repository(source, root, options = "--noupdate", ssh = ssh) } - class Repository private[Mercurial](root_path: Path, ssh: SSH.System = SSH.Local) - { + class Repository private[Mercurial](root_path: Path, ssh: SSH.System = SSH.Local) { hg => val root: Path = ssh.expand_path(root_path) @@ -174,18 +165,23 @@ override def toString: String = ssh.hg_url + root.implode - def command_line(name: String, args: String = "", options: String = "", - repository: Boolean = true): String = - { + def command_line( + name: String, + args: String = "", + options: String = "", + repository: Boolean = true + ): String = { "export LANG=C HGPLAIN=\n\"${HG:-hg}\" --config " + Bash.string("defaults." + name + "=") + (if (repository) " --repository " + ssh.bash_path(root) else "") + " --noninteractive " + name + " " + options + " " + args } def command( - name: String, args: String = "", options: String = "", - repository: Boolean = true): Process_Result = - { + name: String, + args: String = "", + options: String = "", + repository: Boolean = true + ): Process_Result = { ssh.execute(command_line(name, args = args, options = options, repository = repository)) } @@ -203,8 +199,7 @@ def id(rev: String = "tip"): String = identify(rev, options = "-i") - def tags(rev: String = "tip"): String = - { + def tags(rev: String = "tip"): String = { val result = identify(rev, options = "-t") Library.space_explode(' ', result).filterNot(_ == "tip").mkString(" ") } @@ -221,8 +216,11 @@ def parent(): String = log(rev = "p1()", template = "{node|short}") def push( - remote: String = "", rev: String = "", force: Boolean = false, options: String = ""): Unit = - { + remote: String = "", + rev: String = "", + force: Boolean = false, + options: String = "" + ): Unit = { hg.command("push", opt_rev(rev) + opt_flag("--force", force) + optional(remote), options). check_rc(rc => rc == 0 | rc == 1) } @@ -231,8 +229,11 @@ hg.command("pull", opt_rev(rev) + optional(remote), options).check def update( - rev: String = "", clean: Boolean = false, check: Boolean = false, options: String = ""): Unit = - { + rev: String = "", + clean: Boolean = false, + check: Boolean = false, + options: String = "" + ): Unit = { hg.command("update", opt_rev(rev) + opt_flag("--clean", clean) + opt_flag("--check", check), options).check } @@ -240,8 +241,7 @@ def known_files(): List[String] = hg.command("status", options = "--modified --added --clean --no-status").check.out_lines - def graph(): Graph = - { + def graph(): Graph = { val Node = """^node: (\w{12}) (\w{12}) (\w{12})""".r val log_result = log(template = """node: {node|short} {p1node|short} {p2node|short}\n""") @@ -258,13 +258,11 @@ /* check files */ - def check_files(files: List[Path], ssh: SSH.System = SSH.Local): (List[Path], List[Path]) = - { + def check_files(files: List[Path], ssh: SSH.System = SSH.Local): (List[Path], List[Path]) = { val outside = new mutable.ListBuffer[Path] val unknown = new mutable.ListBuffer[Path] - @tailrec def check(paths: List[Path]): Unit = - { + @tailrec def check(paths: List[Path]): Unit = { paths match { case path :: rest => find_repository(path, ssh) match { @@ -287,15 +285,13 @@ /* setup remote vs. local repository */ - private def edit_hgrc(local_hg: Repository, path_name: String, source: String): Unit = - { + private def edit_hgrc(local_hg: Repository, path_name: String, source: String): Unit = { val hgrc = local_hg.root + Path.explode(".hg/hgrc") def header(line: String): Boolean = line.startsWith("[paths]") val Entry = """^(\S+)\s*=\s*(.*)$""".r val new_entry = path_name + " = " + source - def commit(lines: List[String]): Boolean = - { + def commit(lines: List[String]): Boolean = { File.write(hgrc, cat_lines(lines)) true } @@ -332,8 +328,8 @@ remote_name: String = "", path_name: String = default_path_name, remote_exists: Boolean = false, - progress: Progress = new Progress): Unit = - { + progress: Progress = new Progress + ): Unit = { /* local repository */ Isabelle_System.make_directory(local_path) @@ -401,13 +397,13 @@ val isabelle_tool = Isabelle_Tool("hg_setup", "setup remote vs. local Mercurial repository", - Scala_Project.here, args => - { - var remote_name = "" - var path_name = default_path_name - var remote_exists = false + Scala_Project.here, + { args => + var remote_name = "" + var path_name = default_path_name + var remote_exists = false - val getopts = Getopts(""" + val getopts = Getopts(""" Usage: isabelle hg_setup [OPTIONS] REMOTE LOCAL_DIR Options are: @@ -418,21 +414,21 @@ Setup a remote vs. local Mercurial repository: REMOTE either refers to a Phabricator server "user@host" or SSH file server "ssh://user@host/path". """, - "n:" -> (arg => remote_name = arg), - "p:" -> (arg => path_name = arg), - "r" -> (_ => remote_exists = true)) + "n:" -> (arg => remote_name = arg), + "p:" -> (arg => path_name = arg), + "r" -> (_ => remote_exists = true)) - val more_args = getopts(args) + val more_args = getopts(args) - val (remote, local_path) = - more_args match { - case List(arg1, arg2) => (arg1, Path.explode(arg2)) - case _ => getopts.usage() - } + val (remote, local_path) = + more_args match { + case List(arg1, arg2) => (arg1, Path.explode(arg2)) + case _ => getopts.usage() + } - val progress = new Console_Progress + val progress = new Console_Progress - hg_setup(remote, local_path, remote_name = remote_name, path_name = path_name, - remote_exists = remote_exists, progress = progress) - }) + hg_setup(remote, local_path, remote_name = remote_name, path_name = path_name, + remote_exists = remote_exists, progress = progress) + }) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/multi_map.scala --- a/src/Pure/General/multi_map.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/multi_map.scala Fri Apr 01 23:26:19 2022 +0200 @@ -11,8 +11,7 @@ import scala.collection.immutable.{Iterable, MapOps} -object Multi_Map extends MapFactory[Multi_Map] -{ +object Multi_Map extends MapFactory[Multi_Map] { private val empty_val: Multi_Map[Any, Nothing] = new Multi_Map[Any, Nothing](Map.empty) def empty[A, B]: Multi_Map[A, B] = empty_val.asInstanceOf[Multi_Map[A, B]] @@ -20,8 +19,7 @@ entries.iterator.foldLeft(empty[A, B]) { case (m, (a, b)) => m.insert(a, b) } override def newBuilder[A, B]: mutable.Builder[(A, B), Multi_Map[A, B]] = new Builder[A, B] - private class Builder[A, B] extends mutable.Builder[(A, B), Multi_Map[A, B]] - { + private class Builder[A, B] extends mutable.Builder[(A, B), Multi_Map[A, B]] { private var res = empty[A, B] override def clear(): Unit = { res = empty[A, B] } override def addOne(p: (A, B)): this.type = { res = res.insert(p._1, p._2); this } @@ -32,23 +30,20 @@ final class Multi_Map[A, +B] private(protected val rep: Map[A, List[B]]) extends Iterable[(A, B)] with MapOps[A, B, Multi_Map, Multi_Map[A, B]] - with MapFactoryDefaults[A, B, Multi_Map, Iterable] -{ + with MapFactoryDefaults[A, B, Multi_Map, Iterable] { /* Multi_Map operations */ def iterator_list: Iterator[(A, List[B])] = rep.iterator def get_list(a: A): List[B] = rep.getOrElse(a, Nil) - def insert[B1 >: B](a: A, b: B1): Multi_Map[A, B1] = - { + def insert[B1 >: B](a: A, b: B1): Multi_Map[A, B1] = { val bs = get_list(a) if (bs.contains(b)) this else new Multi_Map(rep + (a -> (b :: bs))) } - def remove[B1 >: B](a: A, b: B1): Multi_Map[A, B1] = - { + def remove[B1 >: B](a: A, b: B1): Multi_Map[A, B1] = { val bs = get_list(a) if (bs.contains(b)) { bs.filterNot(_ == b) match { diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/output.scala --- a/src/Pure/General/output.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/output.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Output -{ +object Output { def clean_yxml(msg: String): String = try { XML.content(Protocol_Message.clean_reports(YXML.parse_body(msg))) } catch { case ERROR(_) => msg } @@ -21,24 +20,21 @@ def error_prefix(s: String): String = Library.prefix_lines("*** ", s) def error_message_text(msg: String): String = error_prefix(clean_yxml(msg)) - def writeln(msg: String, stdout: Boolean = false, include_empty: Boolean = false): Unit = - { + def writeln(msg: String, stdout: Boolean = false, include_empty: Boolean = false): Unit = { if (msg.nonEmpty || include_empty) { if (stdout) Console.print(writeln_text(msg) + "\n") else Console.err.print(writeln_text(msg) + "\n") } } - def warning(msg: String, stdout: Boolean = false, include_empty: Boolean = false): Unit = - { + def warning(msg: String, stdout: Boolean = false, include_empty: Boolean = false): Unit = { if (msg.nonEmpty || include_empty) { if (stdout) Console.print(warning_text(msg) + "\n") else Console.err.print(warning_text(msg) + "\n") } } - def error_message(msg: String, stdout: Boolean = false, include_empty: Boolean = false): Unit = - { + def error_message(msg: String, stdout: Boolean = false, include_empty: Boolean = false): Unit = { if (msg.nonEmpty || include_empty) { if (stdout) Console.print(error_message_text(msg) + "\n") else Console.err.print(error_message_text(msg) + "\n") diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/path.scala --- a/src/Pure/General/path.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/path.scala Fri Apr 01 23:26:19 2022 +0200 @@ -15,8 +15,7 @@ import scala.util.matching.Regex -object Path -{ +object Path { /* path elements */ sealed abstract class Elem @@ -95,8 +94,7 @@ /* explode */ - def explode(str: String): Path = - { + def explode(str: String): Path = { def explode_elem(s: String): Elem = try { if (s == "..") Parent @@ -148,8 +146,7 @@ /* case-insensitive names */ - def check_case_insensitive(paths: List[Path]): Unit = - { + def check_case_insensitive(paths: List[Path]): Unit = { val table = paths.foldLeft(Multi_Map.empty[String, String]) { case (tab, path) => val name = path.expand.implode @@ -164,8 +161,9 @@ } -final class Path private(protected val elems: List[Path.Elem]) // reversed elements -{ +final class Path private( + protected val elems: List[Path.Elem] // reversed elements +) { override def hashCode: Int = elems.hashCode override def equals(that: Any): Boolean = that match { @@ -237,14 +235,12 @@ def patch: Path = ext("patch") def shasum: Path = ext("shasum") - def backup: Path = - { + def backup: Path = { val (prfx, s) = split_path prfx + Path.basic(s + "~") } - def backup2: Path = - { + def backup2: Path = { val (prfx, s) = split_path prfx + Path.basic(s + "~~") } @@ -254,8 +250,7 @@ private val Ext = new Regex("(.*)\\.([^.]*)") - def split_ext: (Path, String) = - { + def split_ext: (Path, String) = { val (prefix, base) = split_path base match { case Ext(b, e) => (prefix + Path.basic(b), e) @@ -271,8 +266,7 @@ /* expand */ - def expand_env(env: JMap[String, String]): Path = - { + def expand_env(env: JMap[String, String]): Path = { def eval(elem: Path.Elem): List[Path.Elem] = elem match { case Path.Variable(s) => @@ -293,8 +287,7 @@ /* implode wrt. given directories */ - def implode_symbolic: String = - { + def implode_symbolic: String = { val directories = Library.space_explode(':', Isabelle_System.getenv("ISABELLE_DIRECTORIES")).reverse val full_name = expand.implode diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/position.scala --- a/src/Pure/General/position.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/position.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Position -{ +object Position { type T = Properties.T val none: T = Nil @@ -28,8 +27,7 @@ val Def_Theory = new Properties.String(Markup.DEF_THEORY) - object Line_File - { + object Line_File { def apply(line: Int, file: String): T = (if (line > 0) Line(line) else Nil) ::: (if (file != "") File(file) else Nil) @@ -42,8 +40,7 @@ } } - object Def_Line_File - { + object Def_Line_File { def unapply(pos: T): Option[(Int, String)] = (pos, pos) match { case (Def_Line(i), Def_File(name)) => Some((i, name)) @@ -52,8 +49,7 @@ } } - object Range - { + object Range { def apply(range: Symbol.Range): T = Offset(range.start) ::: End_Offset(range.stop) def unapply(pos: T): Option[Symbol.Range] = (pos, pos) match { @@ -63,8 +59,7 @@ } } - object Def_Range - { + object Def_Range { def apply(range: Symbol.Range): T = Def_Offset(range.start) ::: Def_End_Offset(range.stop) def unapply(pos: T): Option[Symbol.Range] = (pos, pos) match { @@ -74,8 +69,7 @@ } } - object Item_Id - { + object Item_Id { def unapply(pos: T): Option[(Long, Symbol.Range)] = pos match { case Id(id) => @@ -86,8 +80,7 @@ } } - object Item_Def_Id - { + object Item_Def_Id { def unapply(pos: T): Option[(Long, Symbol.Range)] = pos match { case Def_Id(id) => @@ -98,8 +91,7 @@ } } - object Item_File - { + object Item_File { def unapply(pos: T): Option[(String, Int, Symbol.Range)] = pos match { case Line_File(line, name) => @@ -110,8 +102,7 @@ } } - object Item_Def_File - { + object Item_Def_File { def unapply(pos: T): Option[(String, Int, Symbol.Range)] = pos match { case Def_Line_File(line, name) => @@ -125,8 +116,7 @@ /* here: user output */ - def here(props: T, delimited: Boolean = true): String = - { + def here(props: T, delimited: Boolean = true): String = { val pos = props.filter(Markup.position_property) if (pos.isEmpty) "" else { @@ -145,8 +135,7 @@ /* JSON representation */ - object JSON - { + object JSON { def apply(pos: T): isabelle.JSON.Object.T = isabelle.JSON.Object.empty ++ Line.unapply(pos).map(Line.name -> _) ++ diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/pretty.scala --- a/src/Pure/General/pretty.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/pretty.scala Fri Apr 01 23:26:19 2022 +0200 @@ -9,8 +9,7 @@ import scala.annotation.tailrec -object Pretty -{ +object Pretty { /* XML constructors */ val space: XML.Body = List(XML.Text(Symbol.space)) @@ -36,14 +35,12 @@ /* text metric -- standardized to width of space */ - abstract class Metric - { + abstract class Metric { val unit: Double def apply(s: String): Double } - object Default_Metric extends Metric - { + object Default_Metric extends Metric { val unit = 1.0 def apply(s: String): Double = s.length.toDouble } @@ -57,22 +54,22 @@ private case class Block( markup: Option[(Markup, Option[XML.Body])], consistent: Boolean, indent: Int, body: List[Tree], length: Double) extends Tree - private case class Break(force: Boolean, width: Int, indent: Int) extends Tree - { def length: Double = width.toDouble } + private case class Break(force: Boolean, width: Int, indent: Int) extends Tree { + def length: Double = width.toDouble + } private case class Str(string: String, length: Double) extends Tree private val FBreak = Break(true, 1, 0) private def make_block( - markup: Option[(Markup, Option[XML.Body])], - consistent: Boolean, - indent: Int, - body: List[Tree]): Tree = - { + markup: Option[(Markup, Option[XML.Body])], + consistent: Boolean, + indent: Int, + body: List[Tree] + ): Tree = { val indent1 = force_nat(indent) - @tailrec def body_length(prts: List[Tree], len: Double): Double = - { + @tailrec def body_length(prts: List[Tree], len: Double): Double = { val (line, rest) = Library.take_prefix[Tree]({ case Break(true, _, _) => false case _ => true }, prts) val len1 = (line.foldLeft(0.0) { case (l, t) => l + t.length }) max len @@ -88,8 +85,7 @@ /* formatted output */ - private sealed case class Text(tx: XML.Body = Nil, pos: Double = 0.0, nl: Int = 0) - { + private sealed case class Text(tx: XML.Body = Nil, pos: Double = 0.0, nl: Int = 0) { def newline: Text = copy(tx = fbrk :: tx, pos = 0.0, nl = nl + 1) def string(s: String, len: Double): Text = copy(tx = if (s == "") tx else XML.Text(s) :: tx, pos = pos + len) @@ -121,8 +117,8 @@ def formatted(input: XML.Body, margin: Double = default_margin, breakgain: Double = default_breakgain, - metric: Metric = Default_Metric): XML.Body = - { + metric: Metric = Default_Metric + ): XML.Body = { val emergencypos = (margin / 2).round.toInt def make_tree(inp: XML.Body): List[Tree] = diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/properties.scala --- a/src/Pure/General/properties.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/properties.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,20 +7,17 @@ package isabelle -object Properties -{ +object Properties { /* entries */ type Entry = (java.lang.String, java.lang.String) type T = List[Entry] - object Eq - { + object Eq { def apply(a: java.lang.String, b: java.lang.String): java.lang.String = a + "=" + b def apply(entry: Entry): java.lang.String = apply(entry._1, entry._2) - def unapply(str: java.lang.String): Option[Entry] = - { + def unapply(str: java.lang.String): Option[Entry] = { val i = str.indexOf('=') if (i <= 0) None else Some((str.substring(0, i), str.substring(i + 1))) } @@ -32,8 +29,7 @@ def get(props: T, name: java.lang.String): Option[java.lang.String] = props.collectFirst({ case (x, y) if x == name => y }) - def put(props: T, entry: Entry): T = - { + def put(props: T, entry: Entry): T = { val (x, y) = entry def update(ps: T): T = ps match { @@ -54,8 +50,8 @@ def compress(ps: List[T], options: XZ.Options = XZ.options(), - cache: XZ.Cache = XZ.Cache()): Bytes = - { + cache: XZ.Cache = XZ.Cache() + ): Bytes = { if (ps.isEmpty) Bytes.empty else { Bytes(YXML.string_of_body(XML.Encode.list(XML.Encode.properties)(ps))). @@ -63,8 +59,7 @@ } } - def uncompress(bs: Bytes, cache: XML.Cache = XML.Cache.none): List[T] = - { + def uncompress(bs: Bytes, cache: XML.Cache = XML.Cache.none): List[T] = { if (bs.is_empty) Nil else { val ps = @@ -86,16 +81,14 @@ /* entry types */ - class String(val name: java.lang.String) - { + class String(val name: java.lang.String) { def apply(value: java.lang.String): T = List((name, value)) def unapply(props: T): Option[java.lang.String] = props.find(_._1 == name).map(_._2) def get(props: T): java.lang.String = unapply(props).getOrElse("") } - class Boolean(val name: java.lang.String) - { + class Boolean(val name: java.lang.String) { def apply(value: scala.Boolean): T = List((name, Value.Boolean(value))) def unapply(props: T): Option[scala.Boolean] = props.find(_._1 == name) match { @@ -105,8 +98,7 @@ def get(props: T): scala.Boolean = unapply(props).getOrElse(false) } - class Int(val name: java.lang.String) - { + class Int(val name: java.lang.String) { def apply(value: scala.Int): T = List((name, Value.Int(value))) def unapply(props: T): Option[scala.Int] = props.find(_._1 == name) match { @@ -116,8 +108,7 @@ def get(props: T): scala.Int = unapply(props).getOrElse(0) } - class Long(val name: java.lang.String) - { + class Long(val name: java.lang.String) { def apply(value: scala.Long): T = List((name, Value.Long(value))) def unapply(props: T): Option[scala.Long] = props.find(_._1 == name) match { @@ -127,8 +118,7 @@ def get(props: T): scala.Long = unapply(props).getOrElse(0) } - class Double(val name: java.lang.String) - { + class Double(val name: java.lang.String) { def apply(value: scala.Double): T = List((name, Value.Double(value))) def unapply(props: T): Option[scala.Double] = props.find(_._1 == name) match { diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/rdf.scala --- a/src/Pure/General/rdf.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/rdf.scala Fri Apr 01 23:26:19 2022 +0200 @@ -9,8 +9,7 @@ package isabelle -object RDF -{ +object RDF { /* document */ val rdf: XML.Namespace = XML.Namespace("rdf", "http://www.w3.org/1999/02/22-rdf-syntax-ns#") @@ -20,8 +19,8 @@ def document(body: XML.Body, namespaces: List[XML.Namespace] = default_namespaces, - attributes: XML.Attributes = Nil): XML.Elem = - { + attributes: XML.Attributes = Nil + ): XML.Elem = { XML.Elem(Markup(rdf("RDF"), namespaces.map(_.attribute) ::: attributes), body) } @@ -29,8 +28,11 @@ /* multiple triples vs. compact description */ sealed case class Triple( - subject: String, predicate: String, `object`: XML.Body = Nil, resource: String = "") - { + subject: String, + predicate: String, + `object`: XML.Body = Nil, + resource: String = "" + ) { require(`object`.isEmpty || resource.isEmpty) def property: XML.Elem = @@ -79,8 +81,8 @@ /* predicates */ - object Property // binary relation with plain value - { + object Property { + // binary relation with plain value def title: String = dcterms("title") def creator: String = dcterms("creator") def contributor: String = dcterms("contributor") diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/scan.scala --- a/src/Pure/General/scan.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/scan.scala Fri Apr 01 23:26:19 2022 +0200 @@ -17,8 +17,7 @@ import java.net.URL -object Scan -{ +object Scan { /** context of partial line-oriented scans **/ abstract class Line_Context @@ -35,8 +34,7 @@ object Parsers extends Parsers - trait Parsers extends RegexParsers - { + trait Parsers extends RegexParsers { override val whiteSpace: Regex = "".r @@ -49,10 +47,8 @@ /* repeated symbols */ def repeated(pred: Symbol.Symbol => Boolean, min_count: Int, max_count: Int): Parser[String] = - new Parser[String] - { - def apply(in: Input) = - { + new Parser[String] { + def apply(in: Input) = { val start = in.offset val end = in.source.length val matcher = new Symbol.Matcher(in.source) @@ -91,19 +87,16 @@ /* quoted strings */ - private def quoted_body(quote: Symbol.Symbol): Parser[String] = - { + private def quoted_body(quote: Symbol.Symbol): Parser[String] = { rep(many1(sym => sym != quote && sym != "\\") | "\\" + quote | "\\\\" | ("""\\\d\d\d""".r ^? { case x if x.substring(1, 4).toInt <= 255 => x })) ^^ (_.mkString) } - def quoted(quote: Symbol.Symbol): Parser[String] = - { + def quoted(quote: Symbol.Symbol): Parser[String] = { quote ~ quoted_body(quote) ~ quote ^^ { case x ~ y ~ z => x + y + z } }.named("quoted") - def quoted_content(quote: Symbol.Symbol, source: String): String = - { + def quoted_content(quote: Symbol.Symbol, source: String): String = { require(parseAll(quoted(quote), source).successful, "no quoted text") val body = source.substring(1, source.length - 1) if (body.exists(_ == '\\')) { @@ -115,8 +108,7 @@ else body } - def quoted_line(quote: Symbol.Symbol, ctxt: Line_Context): Parser[(String, Line_Context)] = - { + def quoted_line(quote: Symbol.Symbol, ctxt: Line_Context): Parser[(String, Line_Context)] = { ctxt match { case Finished => quote ~ quoted_body(quote) ~ opt_term(quote) ^^ @@ -136,12 +128,10 @@ /* nested text cartouches */ - def cartouche_depth(depth: Int): Parser[(String, Int)] = new Parser[(String, Int)] - { + def cartouche_depth(depth: Int): Parser[(String, Int)] = new Parser[(String, Int)] { require(depth >= 0, "bad cartouche depth") - def apply(in: Input) = - { + def apply(in: Input) = { val start = in.offset val end = in.source.length val matcher = new Symbol.Matcher(in.source) @@ -165,8 +155,7 @@ def cartouche: Parser[String] = cartouche_depth(0) ^? { case (x, d) if d == 0 => x } - def cartouche_line(ctxt: Line_Context): Parser[(String, Line_Context)] = - { + def cartouche_line(ctxt: Line_Context): Parser[(String, Line_Context)] = { def cartouche_context(d: Int): Line_Context = if (d == 0) Finished else Cartouche(d) @@ -182,8 +171,7 @@ val recover_cartouche: Parser[String] = cartouche_depth(0) ^^ (_._1) - def cartouche_content(source: String): String = - { + def cartouche_content(source: String): String = { def err(): Nothing = error("Malformed text cartouche: " + quote(source)) val source1 = Library.try_unprefix(Symbol.open_decoded, source) orElse @@ -195,18 +183,15 @@ /* nested comments */ - private def comment_depth(depth: Int): Parser[(String, Int)] = new Parser[(String, Int)] - { + private def comment_depth(depth: Int): Parser[(String, Int)] = new Parser[(String, Int)] { require(depth >= 0, "bad comment depth") val comment_text: Parser[List[String]] = rep1(many1(sym => sym != "*" && sym != "(") | """\*(?!\))|\((?!\*)""".r) - def apply(in: Input) = - { + def apply(in: Input) = { var rest = in - def try_parse[A](p: Parser[A]): Boolean = - { + def try_parse[A](p: Parser[A]): Boolean = { parse(p ^^^ (()), rest) match { case Success(_, next) => { rest = next; true } case _ => false @@ -228,8 +213,7 @@ def comment: Parser[String] = comment_depth(0) ^? { case (x, d) if d == 0 => x } - def comment_line(ctxt: Line_Context): Parser[(String, Line_Context)] = - { + def comment_line(ctxt: Line_Context): Parser[(String, Line_Context)] = { val depth = ctxt match { case Finished => 0 @@ -246,8 +230,7 @@ val recover_comment: Parser[String] = comment_depth(0) ^^ (_._1) - def comment_content(source: String): String = - { + def comment_content(source: String): String = { require(parseAll(comment, source).successful, "no comment") source.substring(2, source.length - 2) } @@ -262,10 +245,8 @@ /* keyword */ - def literal(lexicon: Lexicon): Parser[String] = new Parser[String] - { - def apply(in: Input) = - { + def literal(lexicon: Lexicon): Parser[String] = new Parser[String] { + def apply(in: Input) = { val result = lexicon.scan(in) if (result.isEmpty) Failure("keyword expected", in) else Success(result, in.drop(result.length)) @@ -277,8 +258,7 @@ /** Lexicon -- position tree **/ - object Lexicon - { + object Lexicon { /* representation */ private sealed case class Tree(branches: Map[Char, (String, Tree)]) @@ -288,8 +268,7 @@ def apply(elems: String*): Lexicon = empty ++ elems } - final class Lexicon private(rep: Lexicon.Tree) - { + final class Lexicon private(rep: Lexicon.Tree) { /* auxiliary operations */ private def dest(tree: Lexicon.Tree, result: List[String]): List[String] = @@ -297,11 +276,10 @@ case (res, (_, (s, tr))) => if (s.isEmpty) dest(tr, res) else dest(tr, s :: res) } - private def lookup(str: CharSequence): Option[(Boolean, Lexicon.Tree)] = - { + private def lookup(str: CharSequence): Option[(Boolean, Lexicon.Tree)] = { val len = str.length - @tailrec def look(tree: Lexicon.Tree, tip: Boolean, i: Int): Option[(Boolean, Lexicon.Tree)] = - { + @tailrec + def look(tree: Lexicon.Tree, tip: Boolean, i: Int): Option[(Boolean, Lexicon.Tree)] = { if (i < len) { tree.branches.get(str.charAt(i)) match { case Some((s, tr)) => look(tr, s.nonEmpty, i + 1) @@ -376,14 +354,12 @@ /* scan */ - def scan(in: Reader[Char]): String = - { + def scan(in: Reader[Char]): String = { val source = in.source val offset = in.offset val len = source.length - offset - @tailrec def scan_tree(tree: Lexicon.Tree, result: String, i: Int): String = - { + @tailrec def scan_tree(tree: Lexicon.Tree, result: String, i: Int): String = { if (i < len) { tree.branches.get(source.charAt(offset + i)) match { case Some((s, tr)) => scan_tree(tr, if (s.isEmpty) result else s, i + 1) @@ -400,9 +376,7 @@ /** read stream without decoding: efficient length operation **/ - private class Restricted_Seq(seq: IndexedSeq[Char], start: Int, end: Int) - extends CharSequence - { + private class Restricted_Seq(seq: IndexedSeq[Char], start: Int, end: Int) extends CharSequence { def charAt(i: Int): Char = if (0 <= i && i < length) seq(start + i) else throw new IndexOutOfBoundsException @@ -413,8 +387,7 @@ if (0 <= i && i <= j && j <= length) new Restricted_Seq(seq, start + i, start + j) else throw new IndexOutOfBoundsException - override def toString: String = - { + override def toString: String = { val buf = new StringBuilder(length) for (offset <- start until end) buf.append(seq(offset)) buf.toString @@ -423,26 +396,23 @@ abstract class Byte_Reader extends Reader[Char] with AutoCloseable - private def make_byte_reader(stream: InputStream, stream_length: Int): Byte_Reader = - { + private def make_byte_reader(stream: InputStream, stream_length: Int): Byte_Reader = { val buffered_stream = new BufferedInputStream(stream) - val seq = new PagedSeq( - (buf: Array[Char], offset: Int, length: Int) => - { - var i = 0 - var c = 0 - var eof = false - while (!eof && i < length) { - c = buffered_stream.read - if (c == -1) eof = true - else { buf(offset + i) = c.toChar; i += 1 } - } - if (i > 0) i else -1 - }) + val seq = + new PagedSeq({ (buf: Array[Char], offset: Int, length: Int) => + var i = 0 + var c = 0 + var eof = false + while (!eof && i < length) { + c = buffered_stream.read + if (c == -1) eof = true + else { buf(offset + i) = c.toChar; i += 1 } + } + if (i > 0) i else -1 + }) val restricted_seq = new Restricted_Seq(seq, 0, stream_length) - class Paged_Reader(override val offset: Int) extends Byte_Reader - { + class Paged_Reader(override val offset: Int) extends Byte_Reader { override lazy val source: CharSequence = restricted_seq def first: Char = if (seq.isDefinedAt(offset)) seq(offset) else '\u001a' def rest: Paged_Reader = if (seq.isDefinedAt(offset)) new Paged_Reader(offset + 1) else this @@ -457,8 +427,7 @@ def byte_reader(file: JFile): Byte_Reader = make_byte_reader(new FileInputStream(file), file.length.toInt) - def byte_reader(url: URL): Byte_Reader = - { + def byte_reader(url: URL): Byte_Reader = { val connection = url.openConnection val stream = connection.getInputStream val stream_length = connection.getContentLength diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/sha1.scala --- a/src/Pure/General/sha1.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/sha1.scala Fri Apr 01 23:26:19 2022 +0200 @@ -13,10 +13,8 @@ import isabelle.setup.{Build => Setup_Build} -object SHA1 -{ - final class Digest private[SHA1](rep: String) - { +object SHA1 { + final class Digest private[SHA1](rep: String) { override def toString: String = rep override def hashCode: Int = rep.hashCode override def equals(that: Any): Boolean = @@ -29,24 +27,22 @@ def fake_digest(rep: String): Digest = new Digest(rep) - def make_digest(body: MessageDigest => Unit): Digest = - { + def make_digest(body: MessageDigest => Unit): Digest = { val digest_body = new Setup_Build.Digest_Body { def apply(sha: MessageDigest): Unit = body(sha)} new Digest(Setup_Build.make_digest(digest_body)) } def digest(file: JFile): Digest = - make_digest(sha => using(new FileInputStream(file))(stream => - { - val buf = new Array[Byte](65536) - var m = 0 - var cont = true - while (cont) { - m = stream.read(buf, 0, buf.length) - if (m != -1) sha.update(buf, 0, m) - cont = (m != -1) - } - })) + make_digest(sha => using(new FileInputStream(file)) { stream => + val buf = new Array[Byte](65536) + var m = 0 + var cont = true + while (cont) { + m = stream.read(buf, 0, buf.length) + if (m != -1) sha.update(buf, 0, m) + cont = (m != -1) + } + }) def digest(path: Path): Digest = digest(path.file) def digest(bytes: Array[Byte]): Digest = make_digest(_.update(bytes)) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/sql.scala --- a/src/Pure/General/sql.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/sql.scala Fri Apr 01 23:26:19 2022 +0200 @@ -13,8 +13,7 @@ import scala.collection.mutable -object SQL -{ +object SQL { /** SQL language **/ type Source = String @@ -59,8 +58,7 @@ /* types */ - object Type extends Enumeration - { + object Type extends Enumeration { val Boolean = Value("BOOLEAN") val Int = Value("INTEGER") val Long = Value("BIGINT") @@ -84,8 +82,7 @@ /* columns */ - object Column - { + object Column { def bool(name: String, strict: Boolean = false, primary_key: Boolean = false): Column = Column(name, Type.Boolean, strict, primary_key) def int(name: String, strict: Boolean = false, primary_key: Boolean = false): Column = @@ -103,9 +100,12 @@ } sealed case class Column( - name: String, T: Type.Value, strict: Boolean = false, primary_key: Boolean = false, - expr: SQL.Source = "") - { + name: String, + T: Type.Value, + strict: Boolean = false, + primary_key: Boolean = false, + expr: SQL.Source = "" + ) { def make_primary_key: Column = copy(primary_key = true) def apply(table: Table): Column = @@ -130,8 +130,7 @@ /* tables */ - sealed case class Table(name: String, columns: List[Column], body: Source = "") - { + sealed case class Table(name: String, columns: List[Column], body: Source = "") { private val columns_index: Map[String, Int] = columns.iterator.map(_.name).zipWithIndex.toMap @@ -148,8 +147,7 @@ def query_named: Source = query + " AS " + SQL.ident(name) - def create(strict: Boolean = false, sql_type: Type.Value => Source): Source = - { + def create(strict: Boolean = false, sql_type: Type.Value => Source): Source = { val primary_key = columns.filter(_.primary_key).map(_.name) match { case Nil => Nil @@ -189,63 +187,49 @@ /* statements */ - class Statement private[SQL](val db: Database, val rep: PreparedStatement) - extends AutoCloseable - { + class Statement private[SQL](val db: Database, val rep: PreparedStatement) extends AutoCloseable { stmt => - object bool - { + object bool { def update(i: Int, x: Boolean): Unit = rep.setBoolean(i, x) - def update(i: Int, x: Option[Boolean]): Unit = - { + def update(i: Int, x: Option[Boolean]): Unit = { if (x.isDefined) update(i, x.get) else rep.setNull(i, java.sql.Types.BOOLEAN) } } - object int - { + object int { def update(i: Int, x: Int): Unit = rep.setInt(i, x) - def update(i: Int, x: Option[Int]): Unit = - { + def update(i: Int, x: Option[Int]): Unit = { if (x.isDefined) update(i, x.get) else rep.setNull(i, java.sql.Types.INTEGER) } } - object long - { + object long { def update(i: Int, x: Long): Unit = rep.setLong(i, x) - def update(i: Int, x: Option[Long]): Unit = - { + def update(i: Int, x: Option[Long]): Unit = { if (x.isDefined) update(i, x.get) else rep.setNull(i, java.sql.Types.BIGINT) } } - object double - { + object double { def update(i: Int, x: Double): Unit = rep.setDouble(i, x) - def update(i: Int, x: Option[Double]): Unit = - { + def update(i: Int, x: Option[Double]): Unit = { if (x.isDefined) update(i, x.get) else rep.setNull(i, java.sql.Types.DOUBLE) } } - object string - { + object string { def update(i: Int, x: String): Unit = rep.setString(i, x) def update(i: Int, x: Option[String]): Unit = update(i, x.orNull) } - object bytes - { - def update(i: Int, bytes: Bytes): Unit = - { + object bytes { + def update(i: Int, bytes: Bytes): Unit = { if (bytes == null) rep.setBytes(i, null) else rep.setBinaryStream(i, bytes.stream(), bytes.length) } def update(i: Int, bytes: Option[Bytes]): Unit = update(i, bytes.orNull) } - object date - { + object date { def update(i: Int, date: Date): Unit = db.update_date(stmt, i, date) def update(i: Int, date: Option[Date]): Unit = update(i, date.orNull) } @@ -259,14 +243,12 @@ /* results */ - class Result private[SQL](val stmt: Statement, val rep: ResultSet) - { + class Result private[SQL](val stmt: Statement, val rep: ResultSet) { res => def next(): Boolean = rep.next() - def iterator[A](get: Result => A): Iterator[A] = new Iterator[A] - { + def iterator[A](get: Result => A): Iterator[A] = new Iterator[A] { private var _next: Boolean = res.next() def hasNext: Boolean = _next def next(): A = { val x = get(res); _next = res.next(); x } @@ -276,13 +258,11 @@ def int(column: Column): Int = rep.getInt(column.name) def long(column: Column): Long = rep.getLong(column.name) def double(column: Column): Double = rep.getDouble(column.name) - def string(column: Column): String = - { + def string(column: Column): String = { val s = rep.getString(column.name) if (s == null) "" else s } - def bytes(column: Column): Bytes = - { + def bytes(column: Column): Bytes = { val bs = rep.getBytes(column.name) if (bs == null) Bytes.empty else Bytes(bs) } @@ -291,8 +271,7 @@ def timing(c1: Column, c2: Column, c3: Column): Timing = Timing(Time.ms(long(c1)), Time.ms(long(c2)), Time.ms(long(c3))) - def get[A](column: Column, f: Column => A): Option[A] = - { + def get[A](column: Column, f: Column => A): Option[A] = { val x = f(column) if (rep.wasNull) None else Some(x) } @@ -308,8 +287,7 @@ /* database */ - trait Database extends AutoCloseable - { + trait Database extends AutoCloseable { db => @@ -324,8 +302,7 @@ def close(): Unit = connection.close() - def transaction[A](body: => A): A = - { + def transaction[A](body: => A): A = { val auto_commit = connection.getAutoCommit try { connection.setAutoCommit(false) @@ -357,8 +334,7 @@ /* tables and views */ - def tables: List[String] = - { + def tables: List[String] = { val result = new mutable.ListBuffer[String] val rs = connection.getMetaData.getTables(null, null, "%", null) while (rs.next) { result += rs.getString(3) } @@ -373,8 +349,7 @@ strict: Boolean = false, unique: Boolean = false): Unit = using_statement(table.create_index(name, columns, strict, unique))(_.execute()) - def create_view(table: Table, strict: Boolean = false): Unit = - { + def create_view(table: Table, strict: Boolean = false): Unit = { if (strict || !tables.contains(table.name)) { val sql = "CREATE VIEW " + table + " AS " + { table.query; table.body } using_statement(sql)(_.execute()) @@ -387,13 +362,11 @@ /** SQLite **/ -object SQLite -{ +object SQLite { // see https://www.sqlite.org/lang_datefunc.html val date_format: Date.Format = Date.Format("uuuu-MM-dd HH:mm:ss.SSS x") - lazy val init_jdbc: Unit = - { + lazy val init_jdbc: Unit = { val lib_path = Path.explode("$ISABELLE_SQLITE_HOME/" + Platform.jvm_platform) val lib_name = File.find_files(lib_path.file) match { @@ -406,8 +379,7 @@ Class.forName("org.sqlite.JDBC") } - def open_database(path: Path): Database = - { + def open_database(path: Path): Database = { init_jdbc val path0 = path.expand val s0 = File.platform_path(path0) @@ -416,8 +388,7 @@ new Database(path0.toString, connection) } - class Database private[SQLite](name: String, val connection: Connection) extends SQL.Database - { + class Database private[SQLite](name: String, val connection: Connection) extends SQL.Database { override def toString: String = name def sql_type(T: SQL.Type.Value): SQL.Source = SQL.sql_type_sqlite(T) @@ -440,8 +411,7 @@ /** PostgreSQL **/ -object PostgreSQL -{ +object PostgreSQL { val default_port = 5432 lazy val init_jdbc: Unit = Class.forName("org.postgresql.Driver") @@ -453,8 +423,8 @@ host: String = "", port: Int = 0, ssh: Option[SSH.Session] = None, - ssh_close: Boolean = false): Database = - { + ssh_close: Boolean = false + ): Database = { init_jdbc if (user == "") error("Undefined database user") @@ -487,9 +457,10 @@ } class Database private[PostgreSQL]( - name: String, val connection: Connection, port_forwarding: Option[SSH.Port_Forwarding]) - extends SQL.Database - { + name: String, + val connection: Connection, + port_forwarding: Option[SSH.Port_Forwarding] + ) extends SQL.Database { override def toString: String = name def sql_type(T: SQL.Type.Value): SQL.Source = SQL.sql_type_postgresql(T) @@ -499,8 +470,7 @@ if (date == null) stmt.rep.setObject(i, null) else stmt.rep.setObject(i, OffsetDateTime.from(date.to(Date.timezone_utc).rep)) - def date(res: SQL.Result, column: SQL.Column): Date = - { + def date(res: SQL.Result, column: SQL.Column): Date = { val obj = res.rep.getObject(column.name, classOf[OffsetDateTime]) if (obj == null) null else Date.instant(obj.toInstant) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/ssh.scala --- a/src/Pure/General/ssh.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/ssh.scala Fri Apr 01 23:26:19 2022 +0200 @@ -18,12 +18,10 @@ JSchException} -object SSH -{ +object SSH { /* target machine: user@host syntax */ - object Target - { + object Target { val User_Host: Regex = "^([^@]+)@(.+)$".r def parse(s: String): (String, String) = @@ -63,8 +61,7 @@ /* init context */ - def init_context(options: Options): Context = - { + def init_context(options: Options): Context = { val config_dir = Path.explode(options.string("ssh_config_dir")) if (!config_dir.is_dir) error("Bad ssh config directory: " + config_dir) @@ -100,16 +97,18 @@ proxy_host = proxy_host, proxy_user = proxy_user, proxy_port = proxy_port, permissive = permissive) - class Context private[SSH](val options: Options, val jsch: JSch) - { + class Context private[SSH](val options: Options, val jsch: JSch) { def update_options(new_options: Options): Context = new Context(new_options, jsch) - private def connect_session(host: String, user: String = "", port: Int = 0, + private def connect_session( + host: String, + user: String = "", + port: Int = 0, host_key_permissive: Boolean = false, nominal_host: String = "", nominal_user: String = "", - on_close: () => Unit = () => ()): Session = - { + on_close: () => Unit = () => () + ): Session = { val session = jsch.getSession(proper_string(user).orNull, host, make_port(port)) session.setUserInfo(No_User_Info) @@ -131,10 +130,15 @@ } def open_session( - host: String, user: String = "", port: Int = 0, actual_host: String = "", - proxy_host: String = "", proxy_user: String = "", proxy_port: Int = 0, - permissive: Boolean = false): Session = - { + host: String, + user: String = "", + port: Int = 0, + actual_host: String = "", + proxy_host: String = "", + proxy_user: String = "", + proxy_port: Int = 0, + permissive: Boolean = false + ): Session = { val connect_host = proper_string(actual_host) getOrElse host if (proxy_host == "") connect_session(host = connect_host, user = user, port = port) else { @@ -148,7 +152,7 @@ connect_session(host = fw.local_host, port = fw.local_port, host_key_permissive = permissive, nominal_host = host, nominal_user = user, user = user, - on_close = () => { fw.close(); proxy.close() }) + on_close = { () => fw.close(); proxy.close() }) } catch { case exn: Throwable => fw.close(); proxy.close(); throw exn } } @@ -158,17 +162,14 @@ /* logging */ - def logging(verbose: Boolean = true, debug: Boolean = false): Unit = - { + def logging(verbose: Boolean = true, debug: Boolean = false): Unit = { JSch.setLogger(if (verbose) new Logger(debug) else null) } - private class Logger(debug: Boolean) extends JSch_Logger - { + private class Logger(debug: Boolean) extends JSch_Logger { def isEnabled(level: Int): Boolean = level != JSch_Logger.DEBUG || debug - def log(level: Int, msg: String): Unit = - { + def log(level: Int, msg: String): Unit = { level match { case JSch_Logger.ERROR | JSch_Logger.FATAL => Output.error_message(msg) case JSch_Logger.WARN => Output.warning(msg) @@ -180,8 +181,7 @@ /* user info */ - object No_User_Info extends UserInfo - { + object No_User_Info extends UserInfo { def getPassphrase: String = null def getPassword: String = null def promptPassword(msg: String): Boolean = false @@ -193,11 +193,15 @@ /* port forwarding */ - object Port_Forwarding - { - def open(ssh: Session, ssh_close: Boolean, - local_host: String, local_port: Int, remote_host: String, remote_port: Int): Port_Forwarding = - { + object Port_Forwarding { + def open( + ssh: Session, + ssh_close: Boolean, + local_host: String, + local_port: Int, + remote_host: String, + remote_port: Int + ): Port_Forwarding = { val port = ssh.session.setPortForwardingL(local_host, local_port, remote_host, remote_port) new Port_Forwarding(ssh, ssh_close, local_host, port, remote_host, remote_port) } @@ -209,13 +213,12 @@ val local_host: String, val local_port: Int, val remote_host: String, - val remote_port: Int) extends AutoCloseable - { + val remote_port: Int + ) extends AutoCloseable { override def toString: String = local_host + ":" + local_port + ":" + remote_host + ":" + remote_port - def close(): Unit = - { + def close(): Unit = { ssh.session.delPortForwardingL(local_host, local_port) if (ssh_close) ssh.close() } @@ -226,8 +229,7 @@ type Attrs = SftpATTRS - sealed case class Dir_Entry(name: String, is_dir: Boolean) - { + sealed case class Dir_Entry(name: String, is_dir: Boolean) { def is_file: Boolean = !is_dir } @@ -236,8 +238,7 @@ private val exec_wait_delay = Time.seconds(0.3) - class Exec private[SSH](session: Session, channel: ChannelExec) extends AutoCloseable - { + class Exec private[SSH](session: Session, channel: ChannelExec) extends AutoCloseable { override def toString: String = "exec " + session.toString def close(): Unit = channel.disconnect @@ -258,16 +259,14 @@ def result( progress_stdout: String => Unit = (_: String) => (), progress_stderr: String => Unit = (_: String) => (), - strict: Boolean = true): Process_Result = - { + strict: Boolean = true + ): Process_Result = { stdin.close() - def read_lines(stream: InputStream, progress: String => Unit): List[String] = - { + def read_lines(stream: InputStream, progress: String => Unit): List[String] = { val result = new mutable.ListBuffer[String] val line_buffer = new ByteArrayOutputStream(100) - def line_flush(): Unit = - { + def line_flush(): Unit = { val line = Library.trim_line(line_buffer.toString(UTF8.charset_name)) progress(line) result += line @@ -292,8 +291,7 @@ val out_lines = Future.thread("ssh_stdout") { read_lines(stdout, progress_stdout) } val err_lines = Future.thread("ssh_stderr") { read_lines(stderr, progress_stderr) } - def terminate(): Unit = - { + def terminate(): Unit = { close() out_lines.join err_lines.join @@ -319,8 +317,8 @@ val session: JSch_Session, on_close: () => Unit, val nominal_host: String, - val nominal_user: String) extends System - { + val nominal_user: String + ) extends System { def update_options(new_options: Options): Session = new Session(new_options, session, on_close, nominal_host, nominal_user) @@ -350,8 +348,7 @@ override def close(): Unit = { sftp.disconnect; session.disconnect; on_close() } - val settings: JMap[String, String] = - { + val settings: JMap[String, String] = { val home = sftp.getHome JMap.of("HOME", home, "USER_HOME", home) } @@ -379,8 +376,7 @@ try { sftp.lstat(remote_path(path)).isLink } catch { case _: SftpException => false } - override def make_directory(path: Path): Path = - { + override def make_directory(path: Path): Path = { if (!is_dir(path)) { execute( "perl -e \"use File::Path make_path; make_path('" + remote_path(path) + "');\"") @@ -389,8 +385,7 @@ path } - def read_dir(path: Path): List[Dir_Entry] = - { + def read_dir(path: Path): List[Dir_Entry] = { if (!is_dir(path)) error("No such directory: " + path.toString) val dir_name = remote_path(path) @@ -416,13 +411,12 @@ start: Path, pred: Path => Boolean = _ => true, include_dirs: Boolean = false, - follow_links: Boolean = false): List[Path] = - { + follow_links: Boolean = false + ): List[Path] = { val result = new mutable.ListBuffer[Path] def check(path: Path): Unit = { if (pred(path)) result += path } - def find(dir: Path): Unit = - { + def find(dir: Path): Unit = { if (include_dirs) check(dir) if (follow_links || !is_link(dir)) { for (entry <- read_dir(dir)) { @@ -454,8 +448,7 @@ /* exec channel */ - def exec(command: String): Exec = - { + def exec(command: String): Exec = { val channel = session.openChannel("exec").asInstanceOf[ChannelExec] channel.setCommand("export USER_HOME=\"$HOME\"\n" + command) new Exec(this, channel) @@ -481,8 +474,7 @@ def tmp_dir(): String = execute("mktemp -d -t tmp.XXXXXXXXXX").check.out - override def with_tmp_dir[A](body: Path => A): A = - { + override def with_tmp_dir[A](body: Path => A): A = { val remote_dir = tmp_dir() try { body(Path.explode(remote_dir)) } finally { rm_tree(remote_dir) } } @@ -491,8 +483,7 @@ /* system operations */ - trait System extends AutoCloseable - { + trait System extends AutoCloseable { def close(): Unit = () def hg_url: String = "" diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/symbol.scala --- a/src/Pure/General/symbol.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/symbol.scala Fri Apr 01 23:26:19 2022 +0200 @@ -12,8 +12,7 @@ import scala.annotation.tailrec -object Symbol -{ +object Symbol { type Symbol = String // counting Isabelle symbols, starting from 1 @@ -28,8 +27,7 @@ private val static_spaces = space * 4000 - def spaces(n: Int): String = - { + def spaces(n: Int): String = { require(n >= 0, "negative spaces") if (n < static_spaces.length) static_spaces.substring(0, n) else space * n @@ -69,8 +67,7 @@ def is_ascii_identifier(s: String): Boolean = s.nonEmpty && is_ascii_letter(s(0)) && s.forall(is_ascii_letdig) - def ascii(c: Char): Symbol = - { + def ascii(c: Char): Symbol = { if (c > 127) error("Non-ASCII character: " + quote(c.toString)) else char_symbol(c) } @@ -95,8 +92,7 @@ def is_newline(s: Symbol): Boolean = s == "\n" || s == "\r" || s == "\r\n" - class Matcher(text: CharSequence) - { + class Matcher(text: CharSequence) { private def ok(i: Int): Boolean = 0 <= i && i < text.length private def char(i: Int): Char = if (ok(i)) text.charAt(i) else 0 private def maybe_char(c: Char, i: Int): Int = if (char(i) == c) i + 1 else i @@ -106,8 +102,7 @@ private def maybe_ascii_id(i: Int): Int = if (is_ascii_letter(char(i))) many_ascii_letdig(i + 1) else i - def match_length(i: Int): Int = - { + def match_length(i: Int): Int = { val a = char(i) val b = char(i + 1) @@ -129,13 +124,11 @@ /* iterator */ def iterator(text: CharSequence): Iterator[Symbol] = - new Iterator[Symbol] - { + new Iterator[Symbol] { private val matcher = new Matcher(text) private var i = 0 def hasNext: Boolean = i < text.length - def next(): Symbol = - { + def next(): Symbol = { val s = matcher.match_symbol(i) i += s.length s @@ -158,14 +151,12 @@ /* decoding offsets */ - object Index - { + object Index { private sealed case class Entry(chr: Int, sym: Int) val empty: Index = new Index(Nil) - def apply(text: CharSequence): Index = - { + def apply(text: CharSequence): Index = { val matcher = new Matcher(text) val buf = new mutable.ListBuffer[Entry] var chr = 0 @@ -180,17 +171,14 @@ } } - final class Index private(entries: List[Index.Entry]) - { + final class Index private(entries: List[Index.Entry]) { private val hash: Int = entries.hashCode private val index: Array[Index.Entry] = entries.toArray - def decode(symbol_offset: Offset): Text.Offset = - { + def decode(symbol_offset: Offset): Text.Offset = { val sym = symbol_offset - 1 val end = index.length - @tailrec def bisect(a: Int, b: Int): Int = - { + @tailrec def bisect(a: Int, b: Int): Int = { if (a < b) { val c = (a + b) / 2 if (sym < index(c).sym) bisect(a, c) @@ -216,8 +204,7 @@ /* symbolic text chunks -- without actual text */ - object Text_Chunk - { + object Text_Chunk { sealed abstract class Name case object Default extends Name case class Id(id: Document_ID.Generic) extends Name @@ -227,8 +214,7 @@ new Text_Chunk(Text.Range(0, text.length), Index(text)) } - final class Text_Chunk private(val range: Text.Range, private val index: Index) - { + final class Text_Chunk private(val range: Text.Range, private val index: Index) { override def hashCode: Int = (range, index).hashCode override def equals(that: Any): Boolean = that match { @@ -242,8 +228,7 @@ def decode(symbol_offset: Offset): Text.Offset = index.decode(symbol_offset) def decode(symbol_range: Range): Text.Range = index.decode(symbol_range) - def incorporate(symbol_range: Range): Option[Text.Range] = - { + def incorporate(symbol_range: Range): Option[Text.Range] = { def in(r: Range): Option[Text.Range] = range.try_restrict(decode(r)) match { case Some(r1) if !r1.is_singularity => Some(r1) @@ -256,10 +241,8 @@ /* recoding text */ - private class Recoder(list: List[(String, String)]) - { - private val (min, max) = - { + private class Recoder(list: List[(String, String)]) { + private val (min, max) = { var min = '\uffff' var max = '\u0000' for ((x, _) <- list) { @@ -269,8 +252,7 @@ } (min, max) } - private val table = - { + private val table = { var tab = Map[String, String]() for ((x, y) <- list) { tab.get(x) match { @@ -281,8 +263,7 @@ } tab } - def recode(text: String): String = - { + def recode(text: String): String = { val len = text.length val matcher = new Symbol.Matcher(text) val result = new StringBuilder(len) @@ -304,8 +285,7 @@ /** defined symbols **/ - object Argument extends Enumeration - { + object Argument extends Enumeration { val none, cartouche, space_cartouche = Value def unapply(s: String): Option[Value] = @@ -313,8 +293,7 @@ catch { case _: NoSuchElementException => None} } - object Entry - { + object Entry { private val Name = new Regex("""\\<\^?([A-Za-z][A-Za-z0-9_']*)>""") private val Argument = new Properties.String("argument") private val Abbrev = new Properties.String("abbrev") @@ -322,8 +301,7 @@ private val Font = new Properties.String("font") private val Group = new Properties.String("group") - def apply(symbol: Symbol, props: Properties.T): Entry = - { + def apply(symbol: Symbol, props: Properties.T): Entry = { def err(msg: String): Nothing = error(msg + " for symbol " + quote(symbol)) val name = @@ -363,8 +341,8 @@ val code: Option[Int], val font: Option[String], val groups: List[String], - val abbrevs: List[String]) - { + val abbrevs: List[String] + ) { override def toString: String = symbol val decode: Option[String] = @@ -373,27 +351,22 @@ lazy val symbols: Symbols = Symbols.load() - object Symbols - { - def load(static: Boolean = false): Symbols = - { + object Symbols { + def load(static: Boolean = false): Symbols = { val paths = if (static) List(Path.explode("~~/etc/symbols")) else Path.split(Isabelle_System.getenv("ISABELLE_SYMBOLS")) make(cat_lines(for (path <- paths if path.is_file) yield File.read(path))) } - def make(symbols_spec: String): Symbols = - { + def make(symbols_spec: String): Symbols = { val No_Decl = new Regex("""(?xs) ^\s* (?: \#.* )? $ """) val Key = new Regex("""(?xs) (.+): """) - def read_decl(decl: String): (Symbol, Properties.T) = - { + def read_decl(decl: String): (Symbol, Properties.T) = { def err() = error("Bad symbol declaration: " + decl) - def read_props(props: List[String]): Properties.T = - { + def read_props(props: List[String]): Properties.T = { props match { case Nil => Nil case _ :: Nil => err() @@ -420,8 +393,7 @@ } } - class Symbols(val entries: List[Entry]) - { + class Symbols(val entries: List[Entry]) { override def toString: String = entries.mkString("Symbols(", ", ", ")") @@ -452,8 +424,7 @@ /* recoding */ - private val (decoder, encoder) = - { + private val (decoder, encoder) = { val mapping = for (entry <- entries; s <- entry.decode) yield entry.symbol -> s (new Recoder(mapping), new Recoder(for ((x, y) <- mapping) yield (y, x))) @@ -565,8 +536,7 @@ def encode_yxml(body: XML.Body): String = encode(YXML.string_of_body(body)) - def decode_strict(text: String): String = - { + def decode_strict(text: String): String = { val decoded = decode(text) if (encode(decoded) == text) decoded else { @@ -685,8 +655,7 @@ if (is_ascii(sym)) is_ascii_printable(sym(0)) else !is_control(sym) - object Metric extends Pretty.Metric - { + object Metric extends Pretty.Metric { val unit = 1.0 def apply(str: String): Double = (for (s <- iterator(str)) yield { diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/time.scala --- a/src/Pure/General/time.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/time.scala Fri Apr 01 23:26:19 2022 +0200 @@ -11,8 +11,7 @@ import java.time.Instant -object Time -{ +object Time { def seconds(s: Double): Time = new Time((s * 1000.0).round) def minutes(m: Double): Time = new Time((m * 60000.0).round) def ms(ms: Long): Time = new Time(ms) @@ -28,8 +27,7 @@ def instant(t: Instant): Time = ms(t.getEpochSecond * 1000L + t.getNano / 1000000L) } -final class Time private(val ms: Long) extends AnyVal -{ +final class Time private(val ms: Long) extends AnyVal { def proper_ms: Option[Long] = if (ms == 0) None else Some(ms) def seconds: Double = ms / 1000.0 @@ -54,8 +52,7 @@ def message: String = toString + "s" - def message_hms: String = - { + def message_hms: String = { val s = ms / 1000 String.format(Locale.ROOT, "%d:%02d:%02d", java.lang.Long.valueOf(s / 3600), diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/timing.scala --- a/src/Pure/General/timing.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/timing.scala Fri Apr 01 23:26:19 2022 +0200 @@ -10,15 +10,14 @@ import java.util.Locale -object Timing -{ +object Timing { val zero: Timing = Timing(Time.zero, Time.zero, Time.zero) def timeit[A]( message: String = "", enabled: Boolean = true, - output: String => Unit = Output.warning(_))(e: => A): A = - { + output: String => Unit = Output.warning(_) + )(e: => A): A = { if (enabled) { val start = Time.now() val result = Exn.capture(e) @@ -40,15 +39,13 @@ String.format(Locale.ROOT, ", factor %.2f", java.lang.Double.valueOf(f)) } -sealed case class Timing(elapsed: Time, cpu: Time, gc: Time) -{ +sealed case class Timing(elapsed: Time, cpu: Time, gc: Time) { def is_zero: Boolean = elapsed.is_zero && cpu.is_zero && gc.is_zero def is_relevant: Boolean = elapsed.is_relevant || cpu.is_relevant || gc.is_relevant def resources: Time = cpu + gc - def factor: Option[Double] = - { + def factor: Option[Double] = { val t1 = elapsed.seconds val t2 = resources.seconds if (t1 >= 3.0 && t2 >= 3.0) Some(t2 / t1) else None @@ -63,8 +60,7 @@ elapsed.message + " elapsed time, " + cpu.message + " cpu time, " + gc.message + " GC time" + Timing.factor_format(cpu.seconds / elapsed.seconds) - def message_resources: String = - { + def message_resources: String = { val factor_text = factor match { case Some(f) => Timing.factor_format(f) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/untyped.scala --- a/src/Pure/General/untyped.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/untyped.scala Fri Apr 01 23:26:19 2022 +0200 @@ -10,17 +10,14 @@ import java.lang.reflect.{Constructor, Method, Field} -object Untyped -{ - def constructor[C](c: Class[C], arg_types: Class[_]*): Constructor[C] = - { +object Untyped { + def constructor[C](c: Class[C], arg_types: Class[_]*): Constructor[C] = { val con = c.getDeclaredConstructor(arg_types: _*) con.setAccessible(true) con } - def method(c: Class[_], name: String, arg_types: Class[_]*): Method = - { + def method(c: Class[_], name: String, arg_types: Class[_]*): Method = { val m = c.getDeclaredMethod(name, arg_types: _*) m.setAccessible(true) m @@ -36,8 +33,7 @@ } } - def field(obj: AnyRef, x: String): Field = - { + def field(obj: AnyRef, x: String): Field = { val iterator = for { c <- classes(obj) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/url.scala --- a/src/Pure/General/url.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/url.scala Fri Apr 01 23:26:19 2022 +0200 @@ -14,8 +14,7 @@ import java.util.zip.GZIPInputStream -object Url -{ +object Url { /* special characters */ def escape_special(c: Char): String = @@ -32,8 +31,7 @@ /* make and check URLs */ - def apply(name: String): URL = - { + def apply(name: String): URL = { try { new URL(name) } catch { case _: MalformedURLException => error("Malformed URL " + quote(name)) } } @@ -52,8 +50,7 @@ def file_name(url: URL): String = Library.take_suffix[Char](c => c != '/' && c != '\\', url.getFile.toString.toList)._2.mkString - def trim_index(url: URL): URL = - { + def trim_index(url: URL): URL = { Library.try_unprefix("/index.html", url.toString) match { case Some(u) => Url(u) case None => diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/utf8.scala --- a/src/Pure/General/utf8.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/utf8.scala Fri Apr 01 23:26:19 2022 +0200 @@ -11,8 +11,7 @@ import scala.io.Codec -object UTF8 -{ +object UTF8 { /* charset */ val charset_name: String = "UTF-8" @@ -27,14 +26,12 @@ // see also https://en.wikipedia.org/wiki/UTF-8#Description // overlong encodings enable byte-stuffing of low-ASCII - def decode_permissive(text: CharSequence): String = - { + def decode_permissive(text: CharSequence): String = { val len = text.length val buf = new java.lang.StringBuilder(len) var code = -1 var rest = 0 - def flush(): Unit = - { + def flush(): Unit = { if (code != -1) { if (rest == 0 && Character.isValidCodePoint(code)) buf.appendCodePoint(code) @@ -43,14 +40,12 @@ rest = 0 } } - def init(x: Int, n: Int): Unit = - { + def init(x: Int, n: Int): Unit = { flush() code = x rest = n } - def push(x: Int): Unit = - { + def push(x: Int): Unit = { if (rest <= 0) init(x, -1) else { code <<= 6 diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/uuid.scala --- a/src/Pure/General/uuid.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/uuid.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object UUID -{ +object UUID { type T = java.util.UUID def random(): T = java.util.UUID.randomUUID() diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/value.scala --- a/src/Pure/General/value.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/value.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,10 +7,8 @@ package isabelle -object Value -{ - object Boolean - { +object Value { + object Boolean { def apply(x: scala.Boolean): java.lang.String = x.toString def unapply(s: java.lang.String): Option[scala.Boolean] = s match { @@ -22,8 +20,7 @@ unapply(s) getOrElse error("Bad boolean: " + quote(s)) } - object Nat - { + object Nat { def unapply(s: java.lang.String): Option[scala.Int] = s match { case Int(n) if n >= 0 => Some(n) @@ -33,8 +30,7 @@ unapply(s) getOrElse error("Bad natural number: " + quote(s)) } - object Int - { + object Int { def apply(x: scala.Int): java.lang.String = Library.signed_string_of_int(x) def unapply(s: java.lang.String): Option[scala.Int] = try { Some(Integer.parseInt(s)) } @@ -43,8 +39,7 @@ unapply(s) getOrElse error("Bad integer: " + quote(s)) } - object Long - { + object Long { def apply(x: scala.Long): java.lang.String = Library.signed_string_of_long(x) def unapply(s: java.lang.String): Option[scala.Long] = try { Some(java.lang.Long.parseLong(s)) } @@ -53,8 +48,7 @@ unapply(s) getOrElse error("Bad integer: " + quote(s)) } - object Double - { + object Double { def apply(x: scala.Double): java.lang.String = x.toString def unapply(s: java.lang.String): Option[scala.Double] = try { Some(java.lang.Double.parseDouble(s)) } @@ -63,10 +57,8 @@ unapply(s) getOrElse error("Bad real: " + quote(s)) } - object Seconds - { - def apply(t: Time): java.lang.String = - { + object Seconds { + def apply(t: Time): java.lang.String = { val s = t.seconds if (s.toInt.toDouble == s) s.toInt.toString else t.toString } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/word.scala --- a/src/Pure/General/word.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/word.scala Fri Apr 01 23:26:19 2022 +0200 @@ -10,8 +10,7 @@ import java.util.Locale -object Word -{ +object Word { /* directionality */ def bidi_detect(str: String): Boolean = @@ -43,8 +42,7 @@ case object Uppercase extends Case case object Capitalized extends Case - object Case - { + object Case { def apply(c: Case, str: String): String = c match { case Lowercase => lowercase(str) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/General/xz.scala --- a/src/Pure/General/xz.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/General/xz.scala Fri Apr 01 23:26:19 2022 +0200 @@ -10,14 +10,12 @@ import org.tukaani.xz.{LZMA2Options, ArrayCache, BasicArrayCache} -object XZ -{ +object XZ { /* options */ type Options = LZMA2Options - def options(preset: Int = 3): Options = - { + def options(preset: Int = 3): Options = { val opts = new LZMA2Options opts.setPreset(preset) opts @@ -28,8 +26,7 @@ type Cache = ArrayCache - object Cache - { + object Cache { def none: ArrayCache = ArrayCache.getDummyCache() def apply(): ArrayCache = ArrayCache.getDefaultCache() def make(): ArrayCache = new BasicArrayCache diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Isar/document_structure.scala --- a/src/Pure/Isar/document_structure.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Isar/document_structure.scala Fri Apr 01 23:26:19 2022 +0200 @@ -11,13 +11,13 @@ import scala.annotation.tailrec -object Document_Structure -{ +object Document_Structure { /** general structure **/ sealed abstract class Document { def length: Int } - case class Block(name: String, text: String, body: List[Document]) extends Document - { val length: Int = body.foldLeft(0)(_ + _.length) } + case class Block(name: String, text: String, body: List[Document]) extends Document { + val length: Int = body.foldLeft(0)(_ + _.length) + } case class Atom(length: Int) extends Document def is_theory_command(keywords: Keyword.Keywords, command: Command): Boolean = @@ -55,8 +55,8 @@ def parse_blocks( syntax: Outer_Syntax, node_name: Document.Node.Name, - text: CharSequence): List[Document] = - { + text: CharSequence + ): List[Document] = { def is_plain_theory(command: Command): Boolean = is_theory_command(syntax.keywords, command) && !command.span.is_begin && !command.span.is_end @@ -83,14 +83,12 @@ def flush(): Unit = if (is_plain_theory(stack.head._1)) close() - def result(): List[Document] = - { + def result(): List[Document] = { while (close()) { } stack.head._2.toList } - def add(command: Command): Unit = - { + def add(command: Command): Unit = { if (command.span.is_begin || is_plain_theory(command)) { flush(); open(command) } else if (command.span.is_end) { flush(); close() } @@ -109,8 +107,7 @@ /** section headings **/ - trait Item - { + trait Item { def name: String = "" def source: String = "" def heading_level: Option[Int] = None @@ -118,15 +115,13 @@ object No_Item extends Item - class Sections(keywords: Keyword.Keywords) - { + class Sections(keywords: Keyword.Keywords) { private def buffer(): mutable.ListBuffer[Document] = new mutable.ListBuffer[Document] private var stack: List[(Int, Item, mutable.ListBuffer[Document])] = List((0, No_Item, buffer())) - @tailrec private def close(level: Int => Boolean): Unit = - { + @tailrec private def close(level: Int => Boolean): Unit = { stack match { case (lev, item, body) :: (_, _, body2) :: _ if level(lev) => body2 += Block(item.name, item.source, body.toList) @@ -136,14 +131,12 @@ } } - def result(): List[Document] = - { + def result(): List[Document] = { close(_ => true) stack.head._3.toList } - def add(item: Item): Unit = - { + def add(item: Item): Unit = { item.heading_level match { case Some(i) => close(_ > i) @@ -157,8 +150,7 @@ /* outer syntax sections */ - private class Command_Item(keywords: Keyword.Keywords, command: Command) extends Item - { + private class Command_Item(keywords: Keyword.Keywords, command: Command) extends Item { override def name: String = command.span.name override def source: String = command.source override def heading_level: Option[Int] = Document_Structure.heading_level(keywords, command) @@ -167,8 +159,8 @@ def parse_sections( syntax: Outer_Syntax, node_name: Document.Node.Name, - text: CharSequence): List[Document] = - { + text: CharSequence + ): List[Document] = { val sections = new Sections(syntax.keywords) for { span <- syntax.parse_spans(text) } { @@ -182,14 +174,12 @@ /* ML sections */ - private class ML_Item(token: ML_Lex.Token, level: Option[Int]) extends Item - { + private class ML_Item(token: ML_Lex.Token, level: Option[Int]) extends Item { override def source: String = token.source override def heading_level: Option[Int] = level } - def parse_ml_sections(SML: Boolean, text: CharSequence): List[Document] = - { + def parse_ml_sections(SML: Boolean, text: CharSequence): List[Document] = { val sections = new Sections(Keyword.Keywords.empty) val nl = new ML_Item(ML_Lex.Token(ML_Lex.Kind.SPACE, "\n"), None) @@ -200,8 +190,7 @@ toks.filterNot(_.is_space) match { case List(tok) if tok.is_comment => val s = tok.source - if (Codepoint.iterator(s).exists(c => Character.isLetter(c) || Character.isDigit(c))) - { + if (Codepoint.iterator(s).exists(c => Character.isLetter(c) || Character.isDigit(c))) { if (s.startsWith("(**** ") && s.endsWith(" ****)")) Some(0) else if (s.startsWith("(*** ") && s.endsWith(" ***)")) Some(1) else if (s.startsWith("(** ") && s.endsWith(" **)")) Some(2) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Isar/keyword.scala --- a/src/Pure/Isar/keyword.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Isar/keyword.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Keyword -{ +object Keyword { /** keyword classification **/ /* kinds */ @@ -111,8 +110,8 @@ kind_pos: Position.T = Position.none, load_command: String = "", load_command_pos: Position.T = Position.none, - tags: List[String] = Nil) - { + tags: List[String] = Nil + ) { override def toString: String = kind + (if (load_command.isEmpty) "" else " (" + quote(load_command) + ")") + @@ -122,17 +121,15 @@ copy(kind = f(kind), load_command = f(load_command), tags = tags.map(f)) } - object Keywords - { + object Keywords { def empty: Keywords = new Keywords() } class Keywords private( val kinds: Map[String, String] = Map.empty, - val load_commands: Map[String, String] = Map.empty) - { - override def toString: String = - { + val load_commands: Map[String, String] = Map.empty + ) { + override def toString: String = { val entries = for ((name, kind) <- kinds.toList.sortBy(_._1)) yield { val load_decl = @@ -175,8 +172,7 @@ /* add keywords */ - def + (name: String, kind: String = "", load_command: String = ""): Keywords = - { + def + (name: String, kind: String = "", load_command: String = ""): Keywords = { val kinds1 = kinds + (name -> kind) val load_commands1 = if (kind == THY_LOAD) { diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Isar/line_structure.scala --- a/src/Pure/Isar/line_structure.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Isar/line_structure.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Line_Structure -{ +object Line_Structure { val init: Line_Structure = Line_Structure() } @@ -19,10 +18,9 @@ depth: Int = 0, span_depth: Int = 0, after_span_depth: Int = 0, - element_depth: Int = 0) -{ - def update(keywords: Keyword.Keywords, tokens: List[Token]): Line_Structure = - { + element_depth: Int = 0 +) { + def update(keywords: Keyword.Keywords, tokens: List[Token]): Line_Structure = { val improper1 = tokens.forall(tok => !tok.is_proper) val blank1 = tokens.forall(_.is_space) val command1 = tokens.exists(_.is_begin_or_command) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Isar/outer_syntax.scala --- a/src/Pure/Isar/outer_syntax.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Isar/outer_syntax.scala Fri Apr 01 23:26:19 2022 +0200 @@ -10,8 +10,7 @@ import scala.collection.mutable -object Outer_Syntax -{ +object Outer_Syntax { /* syntax */ val empty: Outer_Syntax = new Outer_Syntax() @@ -21,8 +20,7 @@ /* string literals */ - def quote_string(str: String): String = - { + def quote_string(str: String): String = { val result = new StringBuilder(str.length + 10) result += '"' for (s <- Symbol.iterator(str)) { @@ -47,8 +45,8 @@ val keywords: Keyword.Keywords = Keyword.Keywords.empty, val rev_abbrevs: Thy_Header.Abbrevs = Nil, val language_context: Completion.Language_Context = Completion.Language_Context.outer, - val has_tokens: Boolean = true) -{ + val has_tokens: Boolean = true +) { /** syntax content **/ override def toString: String = keywords.toString @@ -83,8 +81,7 @@ /* completion */ - private lazy val completion: Completion = - { + private lazy val completion: Completion = { val completion_keywords = (keywords.minor.iterator ++ keywords.major.iterator).toList val completion_abbrevs = completion_keywords.flatMap(name => @@ -109,8 +106,8 @@ start: Text.Offset, text: CharSequence, caret: Int, - context: Completion.Language_Context): Option[Completion.Result] = - { + context: Completion.Language_Context + ): Option[Completion.Result] = { completion.complete(history, unicode, explicit, start, text, caret, context) } @@ -145,8 +142,7 @@ def set_language_context(context: Completion.Language_Context): Outer_Syntax = new Outer_Syntax(keywords, rev_abbrevs, context, has_tokens) - def no_tokens: Outer_Syntax = - { + def no_tokens: Outer_Syntax = { require(keywords.is_empty, "bad syntax keywords") new Outer_Syntax( rev_abbrevs = rev_abbrevs, @@ -160,14 +156,12 @@ /* command spans */ - def parse_spans(toks: List[Token]): List[Command_Span.Span] = - { + def parse_spans(toks: List[Token]): List[Command_Span.Span] = { val result = new mutable.ListBuffer[Command_Span.Span] val content = new mutable.ListBuffer[Token] val ignored = new mutable.ListBuffer[Token] - def ship(content: List[Token]): Unit = - { + def ship(content: List[Token]): Unit = { val kind = if (content.forall(_.is_ignored)) Command_Span.Ignored_Span else if (content.exists(_.is_error)) Command_Span.Malformed_Span @@ -187,8 +181,7 @@ result += Command_Span.Span(kind, content) } - def flush(): Unit = - { + def flush(): Unit = { if (content.nonEmpty) { ship(content.toList); content.clear() } if (ignored.nonEmpty) { ship(ignored.toList); ignored.clear() } } @@ -197,8 +190,9 @@ if (tok.is_ignored) ignored += tok else if (keywords.is_before_command(tok) || tok.is_command && - (!content.exists(keywords.is_before_command) || content.exists(_.is_command))) - { flush(); content += tok } + (!content.exists(keywords.is_before_command) || content.exists(_.is_command))) { + flush(); content += tok + } else { content ++= ignored; ignored.clear(); content += tok } } flush() diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Isar/parse.scala --- a/src/Pure/Isar/parse.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Isar/parse.scala Fri Apr 01 23:26:19 2022 +0200 @@ -11,12 +11,10 @@ import scala.annotation.tailrec -object Parse -{ +object Parse { /* parsing tokens */ - trait Parser extends Parsers - { + trait Parser extends Parsers { type Elem = Token def filter_proper: Boolean = true @@ -27,8 +25,7 @@ private def proper_position: Parser[Position.T] = new Parser[Position.T] { - def apply(raw_input: Input) = - { + def apply(raw_input: Input) = { val in = proper(raw_input) val pos = in.pos match { @@ -44,8 +41,7 @@ def token(s: String, pred: Elem => Boolean): Parser[Elem] = new Parser[Elem] { - def apply(raw_input: Input) = - { + def apply(raw_input: Input) = { val in = proper(raw_input) if (in.atEnd) Failure(s + " expected,\nbut end-of-input was found", in) else { @@ -95,8 +91,7 @@ def parse[T](p: Parser[T], in: Token.Reader): ParseResult[T] = p(in) - def parse_all[T](p: Parser[T], in: Token.Reader): ParseResult[T] = - { + def parse_all[T](p: Parser[T], in: Token.Reader): ParseResult[T] = { val result = parse(p, in) val rest = proper(result.next) if (result.successful && !rest.atEnd) Error("bad input", rest) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Isar/token.scala --- a/src/Pure/Isar/token.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Isar/token.scala Fri Apr 01 23:26:19 2022 +0200 @@ -11,12 +11,10 @@ import scala.util.parsing.input -object Token -{ +object Token { /* tokens */ - object Kind extends Enumeration - { + object Kind extends Enumeration { /*immediate source*/ val COMMAND = Value("command") val KEYWORD = Value("keyword") @@ -46,10 +44,8 @@ object Parsers extends Parsers - trait Parsers extends Scan.Parsers with Comment.Parsers - { - private def delimited_token: Parser[Token] = - { + trait Parsers extends Scan.Parsers with Comment.Parsers { + private def delimited_token: Parser[Token] = { val string = quoted("\"") ^^ (x => Token(Token.Kind.STRING, x)) val alt_string = quoted("`") ^^ (x => Token(Token.Kind.ALT_STRING, x)) val cmt = comment ^^ (x => Token(Token.Kind.INFORMAL_COMMENT, x)) @@ -60,8 +56,7 @@ string | (alt_string | (cmt | (formal_cmt | (cart | ctrl)))) } - private def other_token(keywords: Keyword.Keywords): Parser[Token] = - { + private def other_token(keywords: Keyword.Keywords): Parser[Token] = { val letdigs1 = many1(Symbol.is_letdig) val sub = one(s => s == Symbol.sub_decoded || s == Symbol.sub) val id = @@ -109,9 +104,10 @@ def token(keywords: Keyword.Keywords): Parser[Token] = delimited_token | other_token(keywords) - def token_line(keywords: Keyword.Keywords, ctxt: Scan.Line_Context) - : Parser[(Token, Scan.Line_Context)] = - { + def token_line( + keywords: Keyword.Keywords, + ctxt: Scan.Line_Context + ) : Parser[(Token, Scan.Line_Context)] = { val string = quoted_line("\"", ctxt) ^^ { case (x, c) => (Token(Token.Kind.STRING, x), c) } val alt_string = @@ -135,9 +131,11 @@ case _ => error("Unexpected failure of tokenizing input:\n" + inp.toString) } - def explode_line(keywords: Keyword.Keywords, inp: CharSequence, context: Scan.Line_Context) - : (List[Token], Scan.Line_Context) = - { + def explode_line( + keywords: Keyword.Keywords, + inp: CharSequence, + context: Scan.Line_Context + ) : (List[Token], Scan.Line_Context) = { var in: input.Reader[Char] = Scan.char_reader(inp) val toks = new mutable.ListBuffer[Token] var ctxt = context @@ -197,8 +195,7 @@ /* token reader */ - object Pos - { + object Pos { val none: Pos = new Pos(0, 0, "", "") val start: Pos = new Pos(1, 1, "", "") def file(file: String): Pos = new Pos(1, 1, file, "") @@ -211,14 +208,12 @@ val offset: Symbol.Offset, val file: String, val id: String) - extends input.Position - { + extends input.Position { def column = 0 def lineContents = "" def advance(token: Token): Pos = advance(token.source) - def advance(source: String): Pos = - { + def advance(source: String): Pos = { var line1 = line var offset1 = offset for (s <- Symbol.iterator(source)) { @@ -245,8 +240,7 @@ abstract class Reader extends input.Reader[Token] - private class Token_Reader(tokens: List[Token], val pos: Pos) extends Reader - { + private class Token_Reader(tokens: List[Token], val pos: Pos) extends Reader { def first: Token = tokens.head def rest: Token_Reader = new Token_Reader(tokens.tail, pos.advance(first)) def atEnd: Boolean = tokens.isEmpty @@ -257,8 +251,7 @@ } -sealed case class Token(kind: Token.Kind.Value, source: String) -{ +sealed case class Token(kind: Token.Kind.Value, source: String) { def is_command: Boolean = kind == Token.Kind.COMMAND def is_command(name: String): Boolean = kind == Token.Kind.COMMAND && source == name def is_keyword: Boolean = kind == Token.Kind.KEYWORD @@ -318,8 +311,7 @@ else if (kind == Token.Kind.FORMAL_COMMENT) Comment.content(source) else source - def is_system_name: Boolean = - { + def is_system_name: Boolean = { val s = content is_name && Path.is_wellformed(s) && !s.exists(Symbol.is_ascii_blank) && diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/ML/ml_console.scala --- a/src/Pure/ML/ml_console.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/ML/ml_console.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,12 +7,10 @@ package isabelle -object ML_Console -{ +object ML_Console { /* command line entry point */ - def main(args: Array[String]): Unit = - { + def main(args: Array[String]): Unit = { Command_Line.tool { var dirs: List[Path] = Nil var include_sessions: List[String] = Nil diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/ML/ml_lex.scala --- a/src/Pure/ML/ml_lex.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/ML/ml_lex.scala Fri Apr 01 23:26:19 2022 +0200 @@ -11,8 +11,7 @@ import scala.util.parsing.input.Reader -object ML_Lex -{ +object ML_Lex { /** keywords **/ val keywords: Set[String] = @@ -38,8 +37,7 @@ /** tokens **/ - object Kind extends Enumeration - { + object Kind extends Enumeration { val KEYWORD = Value("keyword") val IDENT = Value("identifier") val LONG_IDENT = Value("long identifier") @@ -63,8 +61,7 @@ val ERROR = Value("bad input") } - sealed case class Token(kind: Kind.Value, source: String) - { + sealed case class Token(kind: Kind.Value, source: String) { def is_keyword: Boolean = kind == Kind.KEYWORD def is_delimiter: Boolean = is_keyword && !Symbol.is_ascii_identifier(source) def is_space: Boolean = kind == Kind.SPACE @@ -78,8 +75,7 @@ case object ML_String extends Scan.Line_Context case class Antiq(ctxt: Scan.Line_Context) extends Scan.Line_Context - private object Parsers extends Scan.Parsers with Antiquote.Parsers with Comment.Parsers - { + private object Parsers extends Scan.Parsers with Antiquote.Parsers with Comment.Parsers { /* string material */ private val blanks = many(character(Symbol.is_ascii_blank)) @@ -120,8 +116,7 @@ private val ml_string: Parser[Token] = "\"" ~ ml_string_body ~ "\"" ^^ { case x ~ y ~ z => Token(Kind.STRING, x + y + z) } - private def ml_string_line(ctxt: Scan.Line_Context): Parser[(Token, Scan.Line_Context)] = - { + private def ml_string_line(ctxt: Scan.Line_Context): Parser[(Token, Scan.Line_Context)] = { def result(x: String, c: Scan.Line_Context) = (Token(Kind.STRING, x), c) ctxt match { @@ -185,8 +180,7 @@ /* token */ - private def other_token: Parser[Token] = - { + private def other_token: Parser[Token] = { /* identifiers */ val letdigs = many(character(Symbol.is_ascii_letdig)) @@ -255,8 +249,7 @@ def token: Parser[Token] = ml_char | (ml_string | (ml_comment | (ml_formal_comment | other_token))) - def token_line(SML: Boolean, ctxt: Scan.Line_Context): Parser[(Token, Scan.Line_Context)] = - { + def token_line(SML: Boolean, ctxt: Scan.Line_Context): Parser[(Token, Scan.Line_Context)] = { val other = (ml_char | other_token) ^^ (x => (x, Scan.Finished)) if (SML) ml_string_line(ctxt) | (ml_comment_line(ctxt) | other) @@ -279,9 +272,11 @@ case _ => error("Unexpected failure of tokenizing input:\n" + input.toString) } - def tokenize_line(SML: Boolean, input: CharSequence, context: Scan.Line_Context) - : (List[Token], Scan.Line_Context) = - { + def tokenize_line( + SML: Boolean, + input: CharSequence, + context: Scan.Line_Context + ) : (List[Token], Scan.Line_Context) = { var in: Reader[Char] = Scan.char_reader(input) val toks = new mutable.ListBuffer[Token] var ctxt = context diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/ML/ml_process.scala --- a/src/Pure/ML/ml_process.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/ML/ml_process.scala Fri Apr 01 23:26:19 2022 +0200 @@ -11,8 +11,7 @@ import java.io.{File => JFile} -object ML_Process -{ +object ML_Process { def apply(options: Options, sessions_structure: Sessions.Structure, store: Sessions.Store, @@ -26,8 +25,8 @@ env: JMap[String, String] = Isabelle_System.settings(), redirect: Boolean = false, cleanup: () => Unit = () => (), - session_base: Option[Sessions.Base] = None): Bash.Process = - { + session_base: Option[Sessions.Base] = None + ): Bash.Process = { val logic_name = Isabelle_System.default_logic(logic) val heaps: List[String] = @@ -101,8 +100,7 @@ // ISABELLE_TMP val isabelle_tmp = Isabelle_System.tmp_dir("process") - val ml_runtime_options = - { + val ml_runtime_options = { val ml_options = Word.explode(Isabelle_System.getenv("ML_OPTIONS")) val ml_options1 = if (ml_options.exists(_.containsSlice("gcthreads"))) ml_options @@ -131,28 +129,27 @@ cwd = cwd, env = bash_env, redirect = redirect, - cleanup = () => - { - isabelle_process_options.delete - init_session.delete - Isabelle_System.rm_tree(isabelle_tmp) - cleanup() - }) + cleanup = { () => + isabelle_process_options.delete + init_session.delete + Isabelle_System.rm_tree(isabelle_tmp) + cleanup() + }) } /* Isabelle tool wrapper */ val isabelle_tool = Isabelle_Tool("process", "raw ML process (batch mode)", - Scala_Project.here, args => - { - var dirs: List[Path] = Nil - var eval_args: List[String] = Nil - var logic = Isabelle_System.getenv("ISABELLE_LOGIC") - var modes: List[String] = Nil - var options = Options.init() + Scala_Project.here, + { args => + var dirs: List[Path] = Nil + var eval_args: List[String] = Nil + var logic = Isabelle_System.getenv("ISABELLE_LOGIC") + var modes: List[String] = Nil + var options = Options.init() - val getopts = Getopts(""" + val getopts = Getopts(""" Usage: isabelle process [OPTIONS] Options are: @@ -166,27 +163,26 @@ Run the raw Isabelle ML process in batch mode. """, - "T:" -> (arg => - eval_args = eval_args ::: List("--eval", "use_thy " + ML_Syntax.print_string_bytes(arg))), - "d:" -> (arg => dirs = dirs ::: List(Path.explode(arg))), - "e:" -> (arg => eval_args = eval_args ::: List("--eval", arg)), - "f:" -> (arg => eval_args = eval_args ::: List("--use", arg)), - "l:" -> (arg => logic = arg), - "m:" -> (arg => modes = arg :: modes), - "o:" -> (arg => options = options + arg)) + "T:" -> (arg => + eval_args = eval_args ::: List("--eval", "use_thy " + ML_Syntax.print_string_bytes(arg))), + "d:" -> (arg => dirs = dirs ::: List(Path.explode(arg))), + "e:" -> (arg => eval_args = eval_args ::: List("--eval", arg)), + "f:" -> (arg => eval_args = eval_args ::: List("--use", arg)), + "l:" -> (arg => logic = arg), + "m:" -> (arg => modes = arg :: modes), + "o:" -> (arg => options = options + arg)) - val more_args = getopts(args) - if (args.isEmpty || more_args.nonEmpty) getopts.usage() + val more_args = getopts(args) + if (args.isEmpty || more_args.nonEmpty) getopts.usage() - val base_info = Sessions.base_info(options, logic, dirs = dirs).check - val store = Sessions.store(options) - val result = - ML_Process(options, base_info.sessions_structure, store, logic = logic, args = eval_args, - modes = modes, session_base = Some(base_info.base)) - .result( - progress_stdout = Output.writeln(_, stdout = true), - progress_stderr = Output.writeln(_)) + val base_info = Sessions.base_info(options, logic, dirs = dirs).check + val store = Sessions.store(options) + val result = + ML_Process(options, base_info.sessions_structure, store, logic = logic, args = eval_args, + modes = modes, session_base = Some(base_info.base)).result( + progress_stdout = Output.writeln(_, stdout = true), + progress_stderr = Output.writeln(_)) - sys.exit(result.rc) - }) + sys.exit(result.rc) + }) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/ML/ml_profiling.scala --- a/src/Pure/ML/ml_profiling.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/ML/ml_profiling.scala Fri Apr 01 23:26:19 2022 +0200 @@ -12,10 +12,8 @@ import scala.collection.immutable.SortedMap -object ML_Profiling -{ - sealed case class Entry(name: String, count: Long) - { +object ML_Profiling { + sealed case class Entry(name: String, count: Long) { def clean_name: Entry = copy(name = """-?\(\d+\).*$""".r.replaceAllIn(name, "")) def print: String = @@ -23,8 +21,7 @@ count.asInstanceOf[AnyRef], name.asInstanceOf[AnyRef]) } - sealed case class Report(kind: String, entries: List[Entry]) - { + sealed case class Report(kind: String, entries: List[Entry]) { def clean_name: Report = copy(entries = entries.map(_.clean_name)) def total: Entry = Entry("TOTAL", entries.iterator.map(_.count).sum) @@ -33,8 +30,7 @@ ("profile_" + kind + ":\n") + cat_lines((entries ::: List(total)).map(_.print)) } - def account(reports: List[Report]): List[Report] = - { + def account(reports: List[Report]): List[Report] = { val empty = SortedMap.empty[String, Long].withDefaultValue(0L) var results = SortedMap.empty[String, SortedMap[String, Long]].withDefaultValue(empty) for (report <- reports) { diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/ML/ml_statistics.scala --- a/src/Pure/ML/ml_statistics.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/ML/ml_statistics.scala Fri Apr 01 23:26:19 2022 +0200 @@ -17,8 +17,7 @@ import org.jfree.chart.plot.PlotOrientation -object ML_Statistics -{ +object ML_Statistics { /* properties */ val Now = new Properties.Double("now") @@ -31,8 +30,7 @@ val Heap_Free = new Properties.Long("size_heap_free_last_GC") val GC_Percent = new Properties.Int("GC_percent") - sealed case class Memory_Status(heap_size: Long, heap_free: Long, gc_percent: Int) - { + sealed case class Memory_Status(heap_size: Long, heap_free: Long, gc_percent: Int) { def heap_used: Long = (heap_size - heap_free) max 0 def heap_used_fraction: Double = if (heap_size == 0) 0.0 else heap_used.toDouble / heap_size @@ -40,8 +38,7 @@ if (1 <= gc_percent && gc_percent <= 100) Some((gc_percent - 1) * 0.01) else None } - def memory_status(props: Properties.T): Memory_Status = - { + def memory_status(props: Properties.T): Memory_Status = { val heap_size = Heap_Size.get(props) val heap_free = Heap_Free.get(props) val gc_percent = GC_Percent.get(props) @@ -54,10 +51,9 @@ def monitor(pid: Long, stats_dir: String = "", delay: Time = Time.seconds(0.5), - consume: Properties.T => Unit = Console.println): Unit = - { - def progress_stdout(line: String): Unit = - { + consume: Properties.T => Unit = Console.println + ): Unit = { + def progress_stdout(line: String): Unit = { val props = Library.space_explode(',', line).flatMap(Properties.Eq.unapply) if (props.nonEmpty) consume(props) } @@ -75,8 +71,7 @@ /* protocol handler */ - class Handler extends Session.Protocol_Handler - { + class Handler extends Session.Protocol_Handler { private var session: Session = null private var monitoring: Future[Unit] = Future.value(()) @@ -182,16 +177,17 @@ /* content interpretation */ - final case class Entry(time: Double, data: Map[String, Double]) - { + final case class Entry(time: Double, data: Map[String, Double]) { def get(field: String): Double = data.getOrElse(field, 0.0) } val empty: ML_Statistics = apply(Nil) - def apply(ml_statistics0: List[Properties.T], heading: String = "", - domain: String => Boolean = _ => true): ML_Statistics = - { + def apply( + ml_statistics0: List[Properties.T], + heading: String = "", + domain: String => Boolean = _ => true + ): ML_Statistics = { require(ml_statistics0.forall(props => Now.unapply(props).isDefined), "missing \"now\" field") val ml_statistics = ml_statistics0.sortBy(now) @@ -205,8 +201,7 @@ (x, _) <- props.iterator if x != Now.name && domain(x) } yield x) - val content = - { + val content = { var last_edge = Map.empty[String, (Double, Double, Double)] val result = new mutable.ListBuffer[ML_Statistics.Entry] for (props <- ml_statistics) { @@ -254,8 +249,8 @@ val fields: Set[String], val content: List[ML_Statistics.Entry], val time_start: Double, - val duration: Double) -{ + val duration: Double +) { override def toString: String = if (content.isEmpty) "ML_Statistics.empty" else "ML_Statistics(length = " + content.length + ", fields = " + fields.size + ")" @@ -266,8 +261,7 @@ def maximum(field: String): Double = content.foldLeft(0.0) { case (m, e) => m max e.get(field) } - def average(field: String): Double = - { + def average(field: String): Double = { @tailrec def sum(t0: Double, list: List[ML_Statistics.Entry], acc: Double): Double = list match { case Nil => acc @@ -285,8 +279,7 @@ /* charts */ - def update_data(data: XYSeriesCollection, selected_fields: List[String]): Unit = - { + def update_data(data: XYSeriesCollection, selected_fields: List[String]): Unit = { data.removeAllSeries() for (field <- selected_fields) { val series = new XYSeries(field) @@ -295,8 +288,7 @@ } } - def chart(title: String, selected_fields: List[String]): JFreeChart = - { + def chart(title: String, selected_fields: List[String]): JFreeChart = { val data = new XYSeriesCollection update_data(data, selected_fields) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/ML/ml_syntax.scala --- a/src/Pure/ML/ml_syntax.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/ML/ml_syntax.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object ML_Syntax -{ +object ML_Syntax { /* numbers */ private def signed(s: String): String = diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/PIDE/byte_message.scala --- a/src/Pure/PIDE/byte_message.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/PIDE/byte_message.scala Fri Apr 01 23:26:19 2022 +0200 @@ -9,8 +9,7 @@ import java.io.{ByteArrayOutputStream, OutputStream, InputStream, IOException} -object Byte_Message -{ +object Byte_Message { /* output operations */ def write(stream: OutputStream, bytes: List[Bytes]): Unit = @@ -20,8 +19,7 @@ try { stream.flush() } catch { case _: IOException => } - def write_line(stream: OutputStream, bytes: Bytes): Unit = - { + def write_line(stream: OutputStream, bytes: Bytes): Unit = { write(stream, List(bytes, Bytes.newline)) flush(stream) } @@ -32,15 +30,13 @@ def read(stream: InputStream, n: Int): Bytes = Bytes.read_stream(stream, limit = n) - def read_block(stream: InputStream, n: Int): (Option[Bytes], Int) = - { + def read_block(stream: InputStream, n: Int): (Option[Bytes], Int) = { val msg = read(stream, n) val len = msg.length (if (len == n) Some(msg) else None, len) } - def read_line(stream: InputStream): Option[Bytes] = - { + def read_line(stream: InputStream): Option[Bytes] = { val line = new ByteArrayOutputStream(100) var c = 0 while ({ c = stream.read; c != -1 && c != 10 }) line.write(c) @@ -60,8 +56,7 @@ private def make_header(ns: List[Int]): List[Bytes] = List(Bytes(ns.mkString(",")), Bytes.newline) - def write_message(stream: OutputStream, chunks: List[Bytes]): Unit = - { + def write_message(stream: OutputStream, chunks: List[Bytes]): Unit = { write(stream, make_header(chunks.map(_.length)) ::: chunks) flush(stream) } @@ -86,14 +81,12 @@ private def is_length(msg: Bytes): Boolean = !msg.is_empty && msg.iterator.forall(b => Symbol.is_ascii_digit(b.toChar)) - private def is_terminated(msg: Bytes): Boolean = - { + private def is_terminated(msg: Bytes): Boolean = { val len = msg.length len > 0 && Symbol.is_ascii_line_terminator(msg.charAt(len - 1)) } - def write_line_message(stream: OutputStream, msg: Bytes): Unit = - { + def write_line_message(stream: OutputStream, msg: Bytes): Unit = { if (is_length(msg) || is_terminated(msg)) error ("Bad content for line message:\n" ++ msg.text.take(100)) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/PIDE/command.scala --- a/src/Pure/PIDE/command.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/PIDE/command.scala Fri Apr 01 23:26:19 2022 +0200 @@ -11,14 +11,11 @@ import scala.collection.immutable.SortedMap -object Command -{ +object Command { /* blobs */ - object Blob - { - def read_file(name: Document.Node.Name, src_path: Path): Blob = - { + object Blob { + def read_file(name: Document.Node.Name, src_path: Path): Blob = { val bytes = Bytes.read(name.path) val chunk = Symbol.Text_Chunk(bytes.text) Blob(name, src_path, Some((bytes.sha1_digest, chunk))) @@ -28,16 +25,15 @@ sealed case class Blob( name: Document.Node.Name, src_path: Path, - content: Option[(SHA1.Digest, Symbol.Text_Chunk)]) - { + content: Option[(SHA1.Digest, Symbol.Text_Chunk)] + ) { def read_file: Bytes = Bytes.read(name.path) def chunk_file: Symbol.Text_Chunk.File = Symbol.Text_Chunk.File(name.node) } - object Blobs_Info - { + object Blobs_Info { val none: Blobs_Info = Blobs_Info(Nil) def errors(msgs: List[String]): Blobs_Info = @@ -52,8 +48,7 @@ /* results */ - object Results - { + object Results { type Entry = (Long, XML.Elem) val empty: Results = new Results(SortedMap.empty) def make(args: IterableOnce[Results.Entry]): Results = @@ -62,8 +57,7 @@ args.iterator.foldLeft(empty)(_ ++ _) } - final class Results private(private val rep: SortedMap[Long, XML.Elem]) - { + final class Results private(private val rep: SortedMap[Long, XML.Elem]) { def is_empty: Boolean = rep.isEmpty def defined(serial: Long): Boolean = rep.isDefinedAt(serial) def get(serial: Long): Option[XML.Elem] = rep.get(serial) @@ -90,16 +84,14 @@ /* exports */ - object Exports - { + object Exports { type Entry = (Long, Export.Entry) val empty: Exports = new Exports(SortedMap.empty) def merge(args: IterableOnce[Exports]): Exports = args.iterator.foldLeft(empty)(_ ++ _) } - final class Exports private(private val rep: SortedMap[Long, Export.Entry]) - { + final class Exports private(private val rep: SortedMap[Long, Export.Entry]) { def is_empty: Boolean = rep.isEmpty def iterator: Iterator[Exports.Entry] = rep.iterator @@ -124,16 +116,14 @@ /* markups */ - object Markup_Index - { + object Markup_Index { val markup: Markup_Index = Markup_Index(false, Symbol.Text_Chunk.Default) def blob(blob: Blob): Markup_Index = Markup_Index(false, blob.chunk_file) } sealed case class Markup_Index(status: Boolean, chunk_name: Symbol.Text_Chunk.Name) - object Markups - { + object Markups { type Entry = (Markup_Index, Markup_Tree) val empty: Markups = new Markups(Map.empty) def init(markup: Markup_Tree): Markups = new Markups(Map(Markup_Index.markup -> markup)) @@ -143,8 +133,7 @@ args.iterator.foldLeft(empty)(_ ++ _) } - final class Markups private(private val rep: Map[Markup_Index, Markup_Tree]) - { + final class Markups private(private val rep: Map[Markup_Index, Markup_Tree]) { def is_empty: Boolean = rep.isEmpty def apply(index: Markup_Index): Markup_Tree = @@ -153,8 +142,7 @@ def add(index: Markup_Index, markup: Text.Markup): Markups = new Markups(rep + (index -> (this(index) + markup))) - def + (entry: Markups.Entry): Markups = - { + def + (entry: Markups.Entry): Markups = { val (index, tree) = entry new Markups(rep + (index -> (this(index).merge(tree, Text.Range.full, Markup.Elements.full)))) } @@ -168,8 +156,7 @@ for (Markup_Index(_, Symbol.Text_Chunk.Id(id)) <- rep.keysIterator) yield id - def redirect(other_id: Document_ID.Generic): Markups = - { + def redirect(other_id: Document_ID.Generic): Markups = { val rep1 = (for { (Markup_Index(status, Symbol.Text_Chunk.Id(id)), markup) <- rep.iterator @@ -190,8 +177,7 @@ /* state */ - object State - { + object State { def get_result(states: List[State], serial: Long): Option[XML.Elem] = states.find(st => st.results.defined(serial)).map(st => st.results.get(serial).get) @@ -225,14 +211,13 @@ status: List[Markup] = Nil, results: Results = Results.empty, exports: Exports = Exports.empty, - markups: Markups = Markups.empty) - { + markups: Markups = Markups.empty + ) { def initialized: Boolean = status.exists(markup => markup.name == Markup.INITIALIZED) def consolidating: Boolean = status.exists(markup => markup.name == Markup.CONSOLIDATING) def consolidated: Boolean = status.exists(markup => markup.name == Markup.CONSOLIDATED) - lazy val maybe_consolidated: Boolean = - { + lazy val maybe_consolidated: Boolean = { var touched = false var forks = 0 var runs = 0 @@ -248,8 +233,7 @@ touched && forks == 0 && runs == 0 } - lazy val document_status: Document_Status.Command_Status = - { + lazy val document_status: Document_Status.Command_Status = { val warnings = if (results.iterator.exists(p => Protocol.is_warning(p._2) || Protocol.is_legacy(p._2))) List(Markup(Markup.WARNING, Nil)) @@ -263,8 +247,7 @@ def markup(index: Markup_Index): Markup_Tree = markups(index) - def redirect(other_command: Command): Option[State] = - { + def redirect(other_command: Command): Option[State] = { val markups1 = markups.redirect(other_command.id) if (markups1.is_empty) None else Some(new State(other_command, markups = markups1)) @@ -281,8 +264,10 @@ else None private def add_markup( - status: Boolean, chunk_name: Symbol.Text_Chunk.Name, m: Text.Markup): State = - { + status: Boolean, + chunk_name: Symbol.Text_Chunk.Name, + m: Text.Markup + ): State = { val markups1 = if (status || Document_Status.Command_Status.liberal_elements(m.info.name)) markups.add(Markup_Index(true, chunk_name), m) @@ -385,8 +370,8 @@ id: Document_ID.Command, node_name: Document.Node.Name, blobs_info: Blobs_Info, - span: Command_Span.Span): Command = - { + span: Command_Span.Span + ): Command = { val (source, span1) = span.compact_source new Command(id, node_name, blobs_info, span1, source, Results.empty, Markups.empty) } @@ -401,8 +386,8 @@ node_name: Document.Node.Name = Document.Node.Name.empty, blobs_info: Blobs_Info = Blobs_Info.none, results: Results = Results.empty, - markups: Markups = Markups.empty): Command = - { + markups: Markups = Markups.empty + ): Command = { val (source1, span1) = Command_Span.unparsed(source, theory).compact_source new Command(id, node_name, blobs_info, span1, source1, results, markups) } @@ -418,17 +403,16 @@ type Edit = (Option[Command], Option[Command]) - object Perspective - { + object Perspective { val empty: Perspective = Perspective(Nil) } - sealed case class Perspective(commands: List[Command]) // visible commands in canonical order - { + sealed case class Perspective( + commands: List[Command] // visible commands in canonical order + ) { def is_empty: Boolean = commands.isEmpty - def same(that: Perspective): Boolean = - { + def same(that: Perspective): Boolean = { val cmds1 = this.commands val cmds2 = that.commands require(!cmds1.exists(_.is_undefined), "cmds1 not defined") @@ -447,8 +431,8 @@ get_blob: Document.Node.Name => Option[Document.Blob], can_import: Document.Node.Name => Boolean, node_name: Document.Node.Name, - span: Command_Span.Span): Blobs_Info = - { + span: Command_Span.Span + ): Blobs_Info = { span.name match { // inlined errors case Thy_Header.THEORY => @@ -491,8 +475,8 @@ def build_blobs_info( syntax: Outer_Syntax, node_name: Document.Node.Name, - load_commands: List[Command_Span.Span]): Blobs_Info = - { + load_commands: List[Command_Span.Span] + ): Blobs_Info = { val blobs = for { span <- load_commands @@ -511,14 +495,14 @@ final class Command private( - val id: Document_ID.Command, - val node_name: Document.Node.Name, - val blobs_info: Command.Blobs_Info, - val span: Command_Span.Span, - val source: String, - val init_results: Command.Results, - val init_markups: Command.Markups) -{ + val id: Document_ID.Command, + val node_name: Document.Node.Name, + val blobs_info: Command.Blobs_Info, + val span: Command_Span.Span, + val source: String, + val init_results: Command.Results, + val init_markups: Command.Markups +) { override def toString: String = id.toString + "/" + span.kind.toString @@ -593,9 +577,9 @@ /* reported positions */ - def reported_position(pos: Position.T) - : Option[(Document_ID.Generic, Symbol.Text_Chunk.Name, Option[Symbol.Range])] = - { + def reported_position( + pos: Position.T + ) : Option[(Document_ID.Generic, Symbol.Text_Chunk.Name, Option[Symbol.Range])] = { pos match { case Position.Id(id) => val chunk_name = @@ -613,8 +597,8 @@ self_id: Document_ID.Generic => Boolean, chunk_name: Symbol.Text_Chunk.Name, chunk: Symbol.Text_Chunk, - message: XML.Elem): Set[Text.Range] = - { + message: XML.Elem + ): Set[Text.Range] = { def elem(props: Properties.T, set: Set[Text.Range]): Set[Text.Range] = reported_position(props) match { case Some((id, name, reported_range)) if self_id(id) && name == chunk_name => diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/PIDE/command_span.scala --- a/src/Pure/PIDE/command_span.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/PIDE/command_span.scala Fri Apr 01 23:26:19 2022 +0200 @@ -11,19 +11,16 @@ import scala.util.parsing.input.CharSequenceReader -object Command_Span -{ +object Command_Span { /* loaded files */ - object Loaded_Files - { + object Loaded_Files { val none: Loaded_Files = Loaded_Files(Nil, -1) } sealed case class Loaded_Files(files: List[String], index: Int) abstract class Load_Command(val name: String, val here: Scala_Project.Here) - extends Isabelle_System.Service - { + extends Isabelle_System.Service { override def toString: String = name def position: Position.T = here.position @@ -66,8 +63,7 @@ /* span */ - sealed case class Span(kind: Kind, content: List[Token]) - { + sealed case class Span(kind: Kind, content: List[Token]) { def is_theory: Boolean = kind == Theory_Span def name: String = @@ -96,8 +92,7 @@ def length: Int = content.foldLeft(0)(_ + _.source.length) - def compact_source: (String, Span) = - { + def compact_source: (String, Span) = { val source = Token.implode(content) val content1 = new mutable.ListBuffer[Token] var i = 0 @@ -110,8 +105,7 @@ (source, Span(kind, content1.toList)) } - def clean_arguments: List[(Token, Int)] = - { + def clean_arguments: List[(Token, Int)] = { if (name.nonEmpty) { def clean(toks: List[(Token, Int)]): List[(Token, Int)] = toks match { @@ -143,8 +137,7 @@ val empty: Span = Span(Ignored_Span, Nil) - def unparsed(source: String, theory: Boolean): Span = - { + def unparsed(source: String, theory: Boolean): Span = { val kind = if (theory) Theory_Span else Malformed_Span Span(kind, List(Token(Token.Kind.UNPARSED, source))) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/PIDE/document.scala --- a/src/Pure/PIDE/document.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/PIDE/document.scala Fri Apr 01 23:26:19 2022 +0200 @@ -11,24 +11,20 @@ import scala.collection.mutable -object Document -{ +object Document { /** document structure **/ /* overlays -- print functions with arguments */ - object Overlays - { + object Overlays { val empty = new Overlays(Map.empty) } - final class Overlays private(rep: Map[Node.Name, Node.Overlays]) - { + final class Overlays private(rep: Map[Node.Name, Node.Overlays]) { def apply(name: Node.Name): Node.Overlays = rep.getOrElse(name, Node.Overlays.empty) - private def update(name: Node.Name, f: Node.Overlays => Node.Overlays): Overlays = - { + private def update(name: Node.Name, f: Node.Overlays => Node.Overlays): Overlays = { val node_overlays = f(apply(name)) new Overlays(if (node_overlays.is_empty) rep - name else rep + (name -> node_overlays)) } @@ -45,19 +41,16 @@ /* document blobs: auxiliary files */ - sealed case class Blob(bytes: Bytes, source: String, chunk: Symbol.Text_Chunk, changed: Boolean) - { + sealed case class Blob(bytes: Bytes, source: String, chunk: Symbol.Text_Chunk, changed: Boolean) { def unchanged: Blob = copy(changed = false) } - object Blobs - { + object Blobs { def apply(blobs: Map[Node.Name, Blob]): Blobs = new Blobs(blobs) val empty: Blobs = apply(Map.empty) } - final class Blobs private(blobs: Map[Node.Name, Blob]) - { + final class Blobs private(blobs: Map[Node.Name, Blob]) { def get(name: Node.Name): Option[Blob] = blobs.get(name) def changed(name: Node.Name): Boolean = @@ -76,16 +69,15 @@ type Edit_Text = Edit[Text.Edit, Text.Perspective] type Edit_Command = Edit[Command.Edit, Command.Perspective] - object Node - { + object Node { /* header and name */ sealed case class Header( imports_pos: List[(Name, Position.T)] = Nil, keywords: Thy_Header.Keywords = Nil, abbrevs: Thy_Header.Abbrevs = Nil, - errors: List[String] = Nil) - { + errors: List[String] = Nil + ) { def imports_offset: Map[Int, Name] = (for { (name, Position.Offset(i)) <- imports_pos } yield i -> name).toMap @@ -101,12 +93,10 @@ val no_header: Header = Header() def bad_header(msg: String): Header = Header(errors = List(msg)) - object Name - { + object Name { val empty: Name = Name("") - object Ordering extends scala.math.Ordering[Name] - { + object Ordering extends scala.math.Ordering[Name] { def compare(name1: Name, name2: Name): Int = name1.node compare name2.node } @@ -116,8 +106,7 @@ Graph.make(entries, symmetric = true)(Ordering) } - sealed case class Name(node: String, master_dir: String = "", theory: String = "") - { + sealed case class Name(node: String, master_dir: String = "", theory: String = "") { override def hashCode: Int = node.hashCode override def equals(that: Any): Boolean = that match { @@ -146,8 +135,7 @@ JSON.Object("node_name" -> node, "theory_name" -> theory) } - sealed case class Entry(name: Node.Name, header: Node.Header) - { + sealed case class Entry(name: Node.Name, header: Node.Header) { def map(f: String => String): Entry = copy(name = name.map(f)) override def toString: String = name.toString @@ -156,13 +144,11 @@ /* node overlays */ - object Overlays - { + object Overlays { val empty = new Overlays(Multi_Map.empty) } - final class Overlays private(rep: Multi_Map[Command, (String, List[String])]) - { + final class Overlays private(rep: Multi_Map[Command, (String, List[String])]) { def commands: Set[Command] = rep.keySet def is_empty: Boolean = rep.isEmpty def dest: List[(Command, (String, List[String]))] = rep.iterator.toList @@ -177,10 +163,8 @@ /* edits */ - sealed abstract class Edit[A, B] - { - def foreach(f: A => Unit): Unit = - { + sealed abstract class Edit[A, B] { + def foreach(f: A => Unit): Unit = { this match { case Edits(es) => es.foreach(f) case _ => @@ -224,14 +208,14 @@ /* commands */ - object Commands - { + object Commands { def apply(commands: Linear_Set[Command]): Commands = new Commands(commands) val empty: Commands = apply(Linear_Set.empty) - def starts(commands: Iterator[Command], offset: Text.Offset = 0) - : Iterator[(Command, Text.Offset)] = - { + def starts( + commands: Iterator[Command], + offset: Text.Offset = 0 + ) : Iterator[(Command, Text.Offset)] = { var i = offset for (command <- commands) yield { val start = i @@ -240,9 +224,10 @@ } } - def starts_pos(commands: Iterator[Command], pos: Token.Pos = Token.Pos.start) - : Iterator[(Command, Token.Pos)] = - { + def starts_pos( + commands: Iterator[Command], + pos: Token.Pos = Token.Pos.start + ) : Iterator[(Command, Token.Pos)] = { var p = pos for (command <- commands) yield { val start = p @@ -254,13 +239,11 @@ private val block_size = 256 } - final class Commands private(val commands: Linear_Set[Command]) - { + final class Commands private(val commands: Linear_Set[Command]) { lazy val load_commands: List[Command] = commands.iterator.filter(cmd => cmd.blobs.nonEmpty).toList - private lazy val full_index: (Array[(Command, Text.Offset)], Text.Range) = - { + private lazy val full_index: (Array[(Command, Text.Offset)], Text.Range) = { val blocks = new mutable.ListBuffer[(Command, Text.Offset)] var next_block = 0 var last_stop = 0 @@ -276,8 +259,7 @@ private def full_range: Text.Range = full_index._2 - def iterator(i: Text.Offset = 0): Iterator[(Command, Text.Offset)] = - { + def iterator(i: Text.Offset = 0): Iterator[(Command, Text.Offset)] = { if (commands.nonEmpty && full_range.contains(i)) { val (cmd0, start0) = full_index._1(i / Commands.block_size) Node.Commands.starts(commands.iterator(cmd0), start0) dropWhile { @@ -296,8 +278,8 @@ val syntax: Option[Outer_Syntax] = None, val text_perspective: Text.Perspective = Text.Perspective.empty, val perspective: Node.Perspective_Command = Node.no_perspective_command, - _commands: Node.Commands = Node.Commands.empty) - { + _commands: Node.Commands = Node.Commands.empty + ) { def is_empty: Boolean = get_blob.isEmpty && header == Node.no_header && @@ -366,18 +348,15 @@ /* development graph */ - object Nodes - { + object Nodes { val empty: Nodes = new Nodes(Graph.empty(Node.Name.Ordering)) } - final class Nodes private(graph: Graph[Node.Name, Node]) - { + final class Nodes private(graph: Graph[Node.Name, Node]) { def apply(name: Node.Name): Node = graph.default_node(name, Node.empty).get_node(name) - def is_suppressed(name: Node.Name): Boolean = - { + def is_suppressed(name: Node.Name): Boolean = { val graph1 = graph.default_node(name, Node.empty) graph1.is_maximal(name) && graph1.get_node(name).is_empty } @@ -388,8 +367,7 @@ case del => Some(new Nodes(del.foldLeft(graph)(_.del_node(_)))) } - def + (entry: (Node.Name, Node)): Nodes = - { + def + (entry: (Node.Name, Node)): Nodes = { val (name, node) = entry val imports = node.header.imports val graph1 = @@ -431,14 +409,14 @@ /* particular document versions */ - object Version - { + object Version { val init: Version = new Version() def make(nodes: Nodes): Version = new Version(Document_ID.make(), nodes) - def purge_future(versions: Map[Document_ID.Version, Version], future: Future[Version]) - : Future[Version] = - { + def purge_future( + versions: Map[Document_ID.Version, Version], + future: Future[Version] + ) : Future[Version] = { if (future.is_finished) { val version = future.join versions.get(version.id) match { @@ -450,8 +428,8 @@ } def purge_suppressed( - versions: Map[Document_ID.Version, Version]): Map[Document_ID.Version, Version] = - { + versions: Map[Document_ID.Version, Version] + ): Map[Document_ID.Version, Version] = { (for ((id, v) <- versions.iterator; v1 <- v.purge_suppressed) yield (id, v1)). foldLeft(versions)(_ + _) } @@ -459,8 +437,8 @@ final class Version private( val id: Document_ID.Version = Document_ID.none, - val nodes: Nodes = Nodes.empty) - { + val nodes: Nodes = Nodes.empty + ) { override def toString: String = "Version(" + id + ")" def purge_suppressed: Option[Version] = @@ -470,8 +448,7 @@ /* changes of plain text, eventually resulting in document edits */ - object Change - { + object Change { val init: Change = new Change() def make(previous: Future[Version], edits: List[Edit_Text], version: Future[Version]): Change = @@ -481,16 +458,15 @@ final class Change private( val previous: Option[Future[Version]] = Some(Future.value(Version.init)), val rev_edits: List[Edit_Text] = Nil, - val version: Future[Version] = Future.value(Version.init)) - { + val version: Future[Version] = Future.value(Version.init) + ) { def is_finished: Boolean = (previous match { case None => true case Some(future) => future.is_finished }) && version.is_finished def truncate: Change = new Change(None, Nil, version) - def purge(versions: Map[Document_ID.Version, Version]): Option[Change] = - { + def purge(versions: Map[Document_ID.Version, Version]): Option[Change] = { val previous1 = previous.map(Version.purge_future(versions, _)) val version1 = Version.purge_future(versions, version) if ((previous eq previous1) && (version eq version1)) None @@ -501,21 +477,19 @@ /* history navigation */ - object History - { + object History { val init: History = new History() } final class History private( - val undo_list: List[Change] = List(Change.init)) // non-empty list - { + val undo_list: List[Change] = List(Change.init) // non-empty list + ) { override def toString: String = "History(" + undo_list.length + ")" def tip: Change = undo_list.head def + (change: Change): History = new History(change :: undo_list) - def prune(check: Change => Boolean, retain: Int): Option[(List[Change], History)] = - { + def prune(check: Change => Boolean, retain: Int): Option[(List[Change], History)] = { val n = undo_list.iterator.zipWithIndex.find(p => check(p._1)).get._2 + 1 val (retained, dropped) = undo_list.splitAt(n max retain) @@ -525,8 +499,7 @@ } } - def purge(versions: Map[Document_ID.Version, Version]): History = - { + def purge(versions: Map[Document_ID.Version, Version]): History = { val undo_list1 = undo_list.map(_.purge(versions)) if (undo_list1.forall(_.isEmpty)) this else new History(for ((a, b) <- undo_list1 zip undo_list) yield a.getOrElse(b)) @@ -536,8 +509,7 @@ /* snapshot: persistent user-view of document state */ - object Snapshot - { + object Snapshot { val init: Snapshot = State.init.snapshot() } @@ -546,8 +518,8 @@ val version: Version, val node_name: Node.Name, edits: List[Text.Edit], - val snippet_command: Option[Command]) - { + val snippet_command: Option[Command] + ) { override def toString: String = "Snapshot(node = " + node_name.node + ", version = " + version.id + (if (is_outdated) ", outdated" else "") + ")" @@ -595,8 +567,7 @@ /* command as add-on snippet */ - def snippet(command: Command): Snapshot = - { + def snippet(command: Command): Snapshot = { val node_name = command.node_name val nodes0 = version.nodes @@ -623,9 +594,9 @@ elements: Markup.Elements = Markup.Elements.full): XML.Body = state.xml_markup(version, node_name, range = range, elements = elements) - def xml_markup_blobs(elements: Markup.Elements = Markup.Elements.full) - : List[(Path, XML.Body)] = - { + def xml_markup_blobs( + elements: Markup.Elements = Markup.Elements.full + ) : List[(Path, XML.Body)] = { snippet_command match { case None => Nil case Some(command) => @@ -707,8 +678,8 @@ def command_results(range: Text.Range): Command.Results = Command.State.merge_results( - select[List[Command.State]](range, Markup.Elements.full, command_states => - { case _ => Some(command_states) }).flatMap(_.info)) + select[List[Command.State]](range, Markup.Elements.full, + command_states => _ => Some(command_states)).flatMap(_.info)) def command_results(command: Command): Command.Results = state.command_results(version, command) @@ -727,8 +698,8 @@ info: A, elements: Markup.Elements, result: List[Command.State] => (A, Text.Markup) => Option[A], - status: Boolean = false): List[Text.Info[A]] = - { + status: Boolean = false + ): List[Text.Info[A]] = { val former_range = revert(range).inflate_singularity val (chunk_name, command_iterator) = commands_loading.headOption match { @@ -755,10 +726,9 @@ range: Text.Range, elements: Markup.Elements, result: List[Command.State] => Text.Markup => Option[A], - status: Boolean = false): List[Text.Info[A]] = - { - def result1(states: List[Command.State]): (Option[A], Text.Markup) => Option[Option[A]] = - { + status: Boolean = false + ): List[Text.Info[A]] = { + def result1(states: List[Command.State]): (Option[A], Text.Markup) => Option[Option[A]] = { val res = result(states) (_: Option[A], x: Text.Markup) => res(x) match { @@ -774,13 +744,11 @@ /* model */ - trait Session - { + trait Session { def resources: Resources } - trait Model - { + trait Model { def session: Session def is_stable: Boolean def snapshot(): Snapshot @@ -798,8 +766,8 @@ def node_edits( node_header: Node.Header, text_edits: List[Text.Edit], - perspective: Node.Perspective_Text): List[Edit_Text] = - { + perspective: Node.Perspective_Text + ): List[Edit_Text] = { val edits: List[Node.Edit[Text.Edit, Text.Perspective]] = get_blob match { case None => @@ -823,26 +791,23 @@ type Assign_Update = List[(Document_ID.Command, List[Document_ID.Exec])] // update of exec state assignment - object State - { + object State { class Fail(state: State) extends Exception - object Assignment - { + object Assignment { val init: Assignment = new Assignment() } final class Assignment private( val command_execs: Map[Document_ID.Command, List[Document_ID.Exec]] = Map.empty, - val is_finished: Boolean = false) - { + val is_finished: Boolean = false + ) { override def toString: String = "Assignment(" + command_execs.size + "," + is_finished + ")" def check_finished: Assignment = { require(is_finished, "assignment not finished"); this } def unfinished: Assignment = new Assignment(command_execs, false) - def assign(update: Assign_Update): Assignment = - { + def assign(update: Assign_Update): Assignment = { require(!is_finished, "assignment already finished") val command_execs1 = update.foldLeft(command_execs) { @@ -876,8 +841,8 @@ /*explicit (linear) history*/ history: History = History.init, /*intermediate state between remove_versions/removed_versions*/ - removing_versions: Boolean = false) - { + removing_versions: Boolean = false + ) { override def toString: String = "State(versions = " + versions.size + ", blobs = " + blobs.size + @@ -890,8 +855,7 @@ private def fail[A]: A = throw new State.Fail(this) - def define_version(version: Version, assignment: State.Assignment): State = - { + def define_version(version: Version, assignment: State.Assignment): State = { val id = version.id copy(versions = versions + (id -> version), assignments = assignments + (id -> assignment.unfinished)) @@ -900,8 +864,7 @@ def define_blob(digest: SHA1.Digest): State = copy(blobs = blobs + digest) def defined_blob(digest: SHA1.Digest): Boolean = blobs.contains(digest) - def define_command(command: Command): State = - { + def define_command(command: Command): State = { val id = command.id if (commands.isDefinedAt(id)) fail else copy(commands = commands + (id -> command.init_state)) @@ -921,9 +884,10 @@ id == st.command.id || (execs.get(id) match { case Some(st1) => st1.command.id == st.command.id case None => false }) - private def other_id(node_name: Node.Name, id: Document_ID.Generic) - : Option[(Symbol.Text_Chunk.Id, Symbol.Text_Chunk)] = - { + private def other_id( + node_name: Node.Name, + id: Document_ID.Generic + ) : Option[(Symbol.Text_Chunk.Id, Symbol.Text_Chunk)] = { for { st <- lookup_id(id) if st.command.node_name == node_name @@ -936,11 +900,12 @@ graph.default_node(id, ()).default_node(st.command.id, ()).add_edge(id, st.command.id) } - def accumulate(id: Document_ID.Generic, message: XML.Elem, cache: XML.Cache) - : (Command.State, State) = - { - def update(st: Command.State): (Command.State, State) = - { + def accumulate( + id: Document_ID.Generic, + message: XML.Elem, + cache: XML.Cache + ) : (Command.State, State) = { + def update(st: Command.State): (Command.State, State) = { val st1 = st.accumulate(self_id(st), other_id, message, cache) (st1, copy(commands_redirection = redirection(st1))) } @@ -958,8 +923,10 @@ } } - def add_export(id: Document_ID.Generic, entry: Command.Exports.Entry): (Command.State, State) = - { + def add_export( + id: Document_ID.Generic, + entry: Command.Exports.Entry + ): (Command.State, State) = { execs.get(id) match { case Some(st) => st.add_export(entry) match { @@ -989,8 +956,8 @@ node_name: Node.Name, id: Document_ID.Exec, source: String, - blobs_info: Command.Blobs_Info): State = - { + blobs_info: Command.Blobs_Info + ): State = { if (theories.isDefinedAt(id)) fail else { val command = @@ -1013,9 +980,11 @@ (state1.snippet(command1), state1) } - def assign(id: Document_ID.Version, edited: List[String], update: Assign_Update) - : ((List[Node.Name], List[Command]), State) = - { + def assign( + id: Document_ID.Version, + edited: List[String], + update: Assign_Update + ) : ((List[Node.Name], List[Command]), State) = { val version = the_version(id) val edited_set = edited.toSet @@ -1064,14 +1033,13 @@ def continue_history( previous: Future[Version], edits: List[Edit_Text], - version: Future[Version]): State = - { + version: Future[Version] + ): State = { val change = Change.make(previous, edits, version) copy(history = history + change) } - def remove_versions(retain: Int = 0): (List[Version], State) = - { + def remove_versions(retain: Int = 0): (List[Version], State) = { history.prune(is_stable, retain) match { case Some((dropped, history1)) => val old_versions = dropped.map(change => change.version.get_finished) @@ -1082,8 +1050,7 @@ } } - def removed_versions(removed: List[Document_ID.Version]): State = - { + def removed_versions(removed: List[Document_ID.Version]): State = { val versions1 = Version.purge_suppressed(versions -- removed) val assignments1 = assignments -- removed @@ -1122,9 +1089,10 @@ removing_versions = false) } - def command_id_map(version: Version, commands: Iterable[Command]) - : Map[Document_ID.Generic, Command] = - { + def command_id_map( + version: Version, + commands: Iterable[Command] + ) : Map[Document_ID.Generic, Command] = { require(is_assigned(version), "version not assigned (command_id_map)") val assignment = the_assignment(version).check_finished (for { @@ -1133,8 +1101,7 @@ } yield (id -> command)).toMap } - def command_maybe_consolidated(version: Version, command: Command): Boolean = - { + def command_maybe_consolidated(version: Version, command: Command): Boolean = { require(is_assigned(version), "version not assigned (command_maybe_consolidated)") try { the_assignment(version).check_finished.command_execs.getOrElse(command.id, Nil) match { @@ -1147,9 +1114,10 @@ catch { case _: State.Fail => false } } - private def command_states_self(version: Version, command: Command) - : List[(Document_ID.Generic, Command.State)] = - { + private def command_states_self( + version: Version, + command: Command + ) : List[(Document_ID.Generic, Command.State)] = { require(is_assigned(version), "version not assigned (command_states_self)") try { the_assignment(version).check_finished.command_execs.getOrElse(command.id, Nil) @@ -1165,8 +1133,7 @@ } } - def command_states(version: Version, command: Command): List[Command.State] = - { + def command_states(version: Version, command: Command): List[Command.State] = { val self = command_states_self(version, command) val others = if (commands_redirection.defined(command.id)) { @@ -1191,8 +1158,8 @@ version: Version, node_name: Node.Name, range: Text.Range = Text.Range.full, - elements: Markup.Elements = Markup.Elements.full): XML.Body = - { + elements: Markup.Elements = Markup.Elements.full + ): XML.Body = { val node = version.nodes(node_name) if (node_name.is_theory) { val markup_index = Command.Markup_Index.markup @@ -1233,8 +1200,7 @@ version.nodes(name).commands.reverse.iterator.forall(command_maybe_consolidated(version, _)) def node_consolidated(version: Version, name: Node.Name): Boolean = - !name.is_theory || - { + !name.is_theory || { val it = version.nodes(name).commands.reverse.iterator it.hasNext && command_states(version, it.next()).exists(_.consolidated) } @@ -1242,8 +1208,8 @@ def snapshot( node_name: Node.Name = Node.Name.empty, pending_edits: List[Text.Edit] = Nil, - snippet_command: Option[Command] = None): Snapshot = - { + snippet_command: Option[Command] = None + ): Snapshot = { val stable = recent_stable val version = stable.version.get_finished diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/PIDE/document_id.scala --- a/src/Pure/PIDE/document_id.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/PIDE/document_id.scala Fri Apr 01 23:26:19 2022 +0200 @@ -9,8 +9,7 @@ package isabelle -object Document_ID -{ +object Document_ID { type Generic = Long type Version = Generic type Command = Generic diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/PIDE/document_status.scala --- a/src/Pure/PIDE/document_status.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/PIDE/document_status.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,12 +7,10 @@ package isabelle -object Document_Status -{ +object Document_Status { /* command status */ - object Command_Status - { + object Command_Status { val proper_elements: Markup.Elements = Markup.Elements(Markup.ACCEPTED, Markup.FORKED, Markup.JOINED, Markup.RUNNING, Markup.FINISHED, Markup.FAILED, Markup.CANCELED) @@ -20,8 +18,7 @@ val liberal_elements: Markup.Elements = proper_elements + Markup.WARNING + Markup.LEGACY + Markup.ERROR - def make(markup_iterator: Iterator[Markup]): Command_Status = - { + def make(markup_iterator: Iterator[Markup]): Command_Status = { var touched = false var accepted = false var warned = false @@ -73,8 +70,8 @@ private val canceled: Boolean, private val finalized: Boolean, forks: Int, - runs: Int) - { + runs: Int + ) { def + (that: Command_Status): Command_Status = Command_Status( touched = touched || that.touched, @@ -99,13 +96,11 @@ /* node status */ - object Node_Status - { + object Node_Status { def make( state: Document.State, version: Document.Version, - name: Document.Node.Name): Node_Status = - { + name: Document.Node.Name): Node_Status = { val node = version.nodes(name) var unprocessed = 0 @@ -159,8 +154,8 @@ terminated: Boolean, initialized: Boolean, finalized: Boolean, - consolidated: Boolean) - { + consolidated: Boolean + ) { def ok: Boolean = failed == 0 def total: Int = unprocessed + running + warned + failed + finished @@ -181,16 +176,15 @@ /* overall timing */ - object Overall_Timing - { + object Overall_Timing { val empty: Overall_Timing = Overall_Timing(0.0, Map.empty) def make( state: Document.State, version: Document.Version, commands: Iterable[Command], - threshold: Double = 0.0): Overall_Timing = - { + threshold: Double = 0.0 + ): Overall_Timing = { var total = 0.0 var command_timings = Map.empty[Command, Double] for { @@ -211,8 +205,7 @@ } } - sealed case class Overall_Timing(total: Double, command_timings: Map[Command, Double]) - { + sealed case class Overall_Timing(total: Double, command_timings: Map[Command, Double]) { def command_timing(command: Command): Double = command_timings.getOrElse(command, 0.0) } @@ -220,20 +213,18 @@ /* nodes status */ - object Overall_Node_Status extends Enumeration - { + object Overall_Node_Status extends Enumeration { val ok, failed, pending = Value } - object Nodes_Status - { + object Nodes_Status { val empty: Nodes_Status = new Nodes_Status(Map.empty, Document.Nodes.empty) } final class Nodes_Status private( private val rep: Map[Document.Node.Name, Node_Status], - nodes: Document.Nodes) - { + nodes: Document.Nodes + ) { def is_empty: Boolean = rep.isEmpty def apply(name: Document.Node.Name): Node_Status = rep(name) def get(name: Document.Node.Name): Option[Node_Status] = rep.get(name) @@ -262,8 +253,8 @@ state: Document.State, version: Document.Version, domain: Option[Set[Document.Node.Name]] = None, - trim: Boolean = false): (Boolean, Nodes_Status) = - { + trim: Boolean = false + ): (Boolean, Nodes_Status) = { val nodes1 = version.nodes val update_iterator = for { @@ -285,8 +276,7 @@ case _ => false } - override def toString: String = - { + override def toString: String = { var ok = 0 var failed = 0 var pending = 0 diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/PIDE/editor.scala --- a/src/Pure/PIDE/editor.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/PIDE/editor.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -abstract class Editor[Context] -{ +abstract class Editor[Context] { /* session */ def session: Session @@ -33,8 +32,7 @@ /* hyperlinks */ - abstract class Hyperlink - { + abstract class Hyperlink { def external: Boolean = false def follow(context: Context): Unit } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/PIDE/headless.scala --- a/src/Pure/PIDE/headless.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/PIDE/headless.scala Fri Apr 01 23:26:19 2022 +0200 @@ -13,13 +13,14 @@ import scala.collection.mutable -object Headless -{ +object Headless { /** session **/ private def stable_snapshot( - state: Document.State, version: Document.Version, name: Document.Node.Name): Document.Snapshot = - { + state: Document.State, + version: Document.Version, + name: Document.Node.Name + ): Document.Snapshot = { val snapshot = state.snapshot(name) assert(version.id == snapshot.version.id) snapshot @@ -29,10 +30,9 @@ val state: Document.State, val version: Document.Version, val nodes: List[(Document.Node.Name, Document_Status.Node_Status)], - val nodes_committed: List[(Document.Node.Name, Document_Status.Node_Status)]) - { - def nodes_pending: List[(Document.Node.Name, Document_Status.Node_Status)] = - { + val nodes_committed: List[(Document.Node.Name, Document_Status.Node_Status)] + ) { + def nodes_pending: List[(Document.Node.Name, Document_Status.Node_Status)] = { val committed = nodes_committed.iterator.map(_._1).toSet nodes.filter(p => !committed(p._1)) } @@ -47,8 +47,8 @@ class Session private[Headless]( session_name: String, _session_options: => Options, - override val resources: Resources) extends isabelle.Session(_session_options, resources) - { + override val resources: Resources) + extends isabelle.Session(_session_options, resources) { session => @@ -78,8 +78,7 @@ override def toString: String = session_name - override def stop(): Process_Result = - { + override def stop(): Process_Result = { try { super.stop() } finally { Isabelle_System.rm_tree(tmp_dir) } } @@ -87,8 +86,7 @@ /* theories */ - private object Load_State - { + private object Load_State { def finished: Load_State = Load_State(Nil, Nil, 0) def count_file(name: Document.Node.Name): Long = @@ -96,15 +94,18 @@ } private case class Load_State( - pending: List[Document.Node.Name], rest: List[Document.Node.Name], load_limit: Long) - { + pending: List[Document.Node.Name], + rest: List[Document.Node.Name], + load_limit: Long + ) { def next( dep_graph: Document.Node.Name.Graph[Unit], - finished: Document.Node.Name => Boolean): (List[Document.Node.Name], Load_State) = - { - def load_requirements(pending1: List[Document.Node.Name], rest1: List[Document.Node.Name]) - : (List[Document.Node.Name], Load_State) = - { + finished: Document.Node.Name => Boolean + ): (List[Document.Node.Name], Load_State) = { + def load_requirements( + pending1: List[Document.Node.Name], + rest1: List[Document.Node.Name] + ) : (List[Document.Node.Name], Load_State) = { val load_theories = dep_graph.all_preds_rev(pending1).filterNot(finished) (load_theories, Load_State(pending1, rest1, load_limit)) } @@ -129,8 +130,8 @@ last_update: Time = Time.now(), nodes_status: Document_Status.Nodes_Status = Document_Status.Nodes_Status.empty, already_committed: Map[Document.Node.Name, Document_Status.Node_Status] = Map.empty, - result: Option[Exn.Result[Use_Theories_Result]] = None) - { + result: Option[Exn.Result[Use_Theories_Result]] = None + ) { def update(new_nodes_status: Document_Status.Nodes_Status): Use_Theories_State = copy(last_update = Time.now(), nodes_status = new_nodes_status) @@ -145,11 +146,11 @@ def cancel_result: Use_Theories_State = if (finished_result) this else copy(result = Some(Exn.Exn(Exn.Interrupt()))) - def clean_theories: (List[Document.Node.Name], Use_Theories_State) = - { - @tailrec def frontier(base: List[Document.Node.Name], front: Set[Document.Node.Name]) - : Set[Document.Node.Name] = - { + def clean_theories: (List[Document.Node.Name], Use_Theories_State) = { + @tailrec def frontier( + base: List[Document.Node.Name], + front: Set[Document.Node.Name] + ) : Set[Document.Node.Name] = { val add = base.filter(name => dep_graph.imm_succs(name).forall(front)) if (add.isEmpty) front else { @@ -176,9 +177,11 @@ } } - def check(state: Document.State, version: Document.Version, beyond_limit: Boolean) - : (List[Document.Node.Name], Use_Theories_State) = - { + def check( + state: Document.State, + version: Document.Version, + beyond_limit: Boolean + ) : (List[Document.Node.Name], Use_Theories_State) = { val already_committed1 = commit match { case None => already_committed @@ -189,8 +192,7 @@ version.nodes(name).header.imports.forall(parent => loaded_theory(parent) || committed.isDefinedAt(parent)) if (!committed.isDefinedAt(name) && parents_committed && - state.node_consolidated(version, name)) - { + state.node_consolidated(version, name)) { val snapshot = stable_snapshot(state, version, name) val status = Document_Status.Node_Status.make(state, version, name) commit_fn(snapshot, status) @@ -209,8 +211,7 @@ if (!finished_result && (beyond_limit || watchdog || dep_graph.keys_iterator.forall(name => - finished_theory(name) || nodes_status.quasi_consolidated(name)))) - { + finished_theory(name) || nodes_status.quasi_consolidated(name)))) { val nodes = (for { name <- dep_graph.keys_iterator @@ -245,10 +246,9 @@ // commit: must not block, must not fail commit: Option[(Document.Snapshot, Document_Status.Node_Status) => Unit] = None, commit_cleanup_delay: Time = default_commit_cleanup_delay, - progress: Progress = new Progress): Use_Theories_Result = - { - val dependencies = - { + progress: Progress = new Progress + ): Use_Theories_Result = { + val dependencies = { val import_names = theories.map(thy => resources.import_name(qualifier, master_directory(master_dir), thy) -> Position.none) @@ -260,8 +260,7 @@ for (path <- dependencies.loaded_files) yield Document.Node.Name(resources.append("", path)) - val use_theories_state = - { + val use_theories_state = { val dep_graph = dependencies.theory_graph val maximals = dep_graph.maximals @@ -279,8 +278,7 @@ Synchronized(Use_Theories_State(dep_graph, load_state, watchdog_timeout, commit)) } - def check_state(beyond_limit: Boolean = false): Unit = - { + def check_state(beyond_limit: Boolean = false): Unit = { val state = session.get_state() for { version <- state.stable_tip_version @@ -289,21 +287,18 @@ } resources.load_theories(session, id, load_theories, dep_files, unicode_symbols, progress) } - val check_progress = - { + val check_progress = { var check_count = 0 - Event_Timer.request(Time.now(), repeat = Some(check_delay)) - { - if (progress.stopped) use_theories_state.change(_.cancel_result) - else { - check_count += 1 - check_state(check_limit > 0 && check_count > check_limit) - } + Event_Timer.request(Time.now(), repeat = Some(check_delay)) { + if (progress.stopped) use_theories_state.change(_.cancel_result) + else { + check_count += 1 + check_state(check_limit > 0 && check_count > check_limit) } + } } - val consumer = - { + val consumer = { val delay_nodes_status = Delay.first(nodes_status_delay max Time.zero) { progress.nodes_status(use_theories_state.value.nodes_status) @@ -326,30 +321,29 @@ val version = snapshot.version val theory_progress = - use_theories_state.change_result(st => - { - val domain = - if (st.nodes_status.is_empty) dep_theories_set - else changed.nodes.iterator.filter(dep_theories_set).toSet + use_theories_state.change_result { st => + val domain = + if (st.nodes_status.is_empty) dep_theories_set + else changed.nodes.iterator.filter(dep_theories_set).toSet - val (nodes_status_changed, nodes_status1) = - st.nodes_status.update(resources, state, version, - domain = Some(domain), trim = changed.assignment) + val (nodes_status_changed, nodes_status1) = + st.nodes_status.update(resources, state, version, + domain = Some(domain), trim = changed.assignment) - if (nodes_status_delay >= Time.zero && nodes_status_changed) { - delay_nodes_status.invoke() - } + if (nodes_status_delay >= Time.zero && nodes_status_changed) { + delay_nodes_status.invoke() + } - val theory_progress = - (for { - (name, node_status) <- nodes_status1.present.iterator - if changed.nodes.contains(name) && !st.already_committed.isDefinedAt(name) - p1 = node_status.percentage - if p1 > 0 && Some(p1) != st.nodes_status.get(name).map(_.percentage) - } yield Progress.Theory(name.theory, percentage = Some(p1))).toList + val theory_progress = + (for { + (name, node_status) <- nodes_status1.present.iterator + if changed.nodes.contains(name) && !st.already_committed.isDefinedAt(name) + p1 = node_status.percentage + if p1 > 0 && Some(p1) != st.nodes_status.get(name).map(_.percentage) + } yield Progress.Theory(name.theory, percentage = Some(p1))).toList - (theory_progress, st.update(nodes_status1)) - }) + (theory_progress, st.update(nodes_status1)) + } theory_progress.foreach(progress.theory) @@ -382,8 +376,8 @@ theories: List[String], qualifier: String = Sessions.DRAFT, master_dir: String = "", - all: Boolean = false): (List[Document.Node.Name], List[Document.Node.Name]) = - { + all: Boolean = false + ): (List[Document.Node.Name], List[Document.Node.Name]) = { val nodes = if (all) None else Some(theories.map(resources.import_name(qualifier, master_directory(master_dir), _))) @@ -395,8 +389,7 @@ /** resources **/ - object Resources - { + object Resources { def apply(options: Options, base_info: Sessions.Base_Info, log: Logger = No_Logger): Resources = new Resources(options, base_info, log = log) @@ -406,8 +399,8 @@ session_dirs: List[Path] = Nil, include_sessions: List[String] = Nil, progress: Progress = new Progress, - log: Logger = No_Logger): Resources = - { + log: Logger = No_Logger + ): Resources = { val base_info = Sessions.base_info(options, session_name, dirs = session_dirs, include_sessions = include_sessions, progress = progress) @@ -418,8 +411,8 @@ val node_name: Document.Node.Name, val node_header: Document.Node.Header, val text: String, - val node_required: Boolean) - { + val node_required: Boolean + ) { override def toString: String = node_name.toString def node_perspective: Document.Node.Perspective_Text = @@ -430,8 +423,7 @@ node_name -> Document.Node.Edits(text_edits), node_name -> node_perspective) - def node_edits(old: Option[Theory]): List[Document.Edit_Text] = - { + def node_edits(old: Option[Theory]): List[Document.Edit_Text] = { val (text_edits, old_required) = if (old.isEmpty) (Text.Edit.inserts(0, text), false) else (Text.Edit.replace(0, old.get.text, text), old.get.node_required) @@ -451,20 +443,17 @@ sealed case class State( blobs: Map[Document.Node.Name, Document.Blob] = Map.empty, theories: Map[Document.Node.Name, Theory] = Map.empty, - required: Multi_Map[Document.Node.Name, UUID.T] = Multi_Map.empty) - { + required: Multi_Map[Document.Node.Name, UUID.T] = Multi_Map.empty + ) { /* blobs */ def doc_blobs: Document.Blobs = Document.Blobs(blobs) - def update_blobs(names: List[Document.Node.Name]): (Document.Blobs, State) = - { + def update_blobs(names: List[Document.Node.Name]): (Document.Blobs, State) = { val new_blobs = - names.flatMap(name => - { + names.flatMap { name => val bytes = Bytes.read(name.path) - def new_blob: Document.Blob = - { + def new_blob: Document.Blob = { val text = bytes.text Document.Blob(bytes, text, Symbol.Text_Chunk(text), changed = true) } @@ -472,15 +461,16 @@ case Some(blob) => if (blob.bytes == bytes) None else Some(name -> new_blob) case None => Some(name -> new_blob) } - }) + } val blobs1 = new_blobs.foldLeft(blobs)(_ + _) val blobs2 = new_blobs.foldLeft(blobs) { case (map, (a, b)) => map + (a -> b.unchanged) } (Document.Blobs(blobs1), copy(blobs = blobs2)) } - def blob_edits(name: Document.Node.Name, old_blob: Option[Document.Blob]) - : List[Document.Edit_Text] = - { + def blob_edits( + name: Document.Node.Name, + old_blob: Option[Document.Blob] + ) : List[Document.Edit_Text] = { val blob = blobs.getOrElse(name, error("Missing blob " + quote(name.toString))) val text_edits = old_blob match { @@ -517,15 +507,16 @@ } }) - def remove_theories(remove: List[Document.Node.Name]): State = - { + def remove_theories(remove: List[Document.Node.Name]): State = { require(remove.forall(name => !is_required(name)), "attempt to remove required nodes") copy(theories = theories -- remove) } - def unload_theories(session: Session, id: UUID.T, theories: List[Document.Node.Name]) - : (List[Document.Edit_Text], State) = - { + def unload_theories( + session: Session, + id: UUID.T, + theories: List[Document.Node.Name] + ) : (List[Document.Edit_Text], State) = { val st1 = remove_required(id, theories) val theory_edits = for { @@ -540,9 +531,10 @@ (theory_edits.flatMap(_._1), st1.update_theories(theory_edits.map(_._2))) } - def purge_theories(session: Session, nodes: Option[List[Document.Node.Name]]) - : ((List[Document.Node.Name], List[Document.Node.Name], List[Document.Edit_Text]), State) = - { + def purge_theories( + session: Session, + nodes: Option[List[Document.Node.Name]] + ) : ((List[Document.Node.Name], List[Document.Node.Name], List[Document.Edit_Text]), State) = { val all_nodes = theory_graph.topological_order val purge = nodes.getOrElse(all_nodes).filterNot(is_required).toSet @@ -559,9 +551,7 @@ val options: Options, val session_base_info: Sessions.Base_Info, log: Logger = No_Logger) - extends isabelle.Resources( - session_base_info.sessions_structure, session_base_info.check.base, log = log) - { + extends isabelle.Resources(session_base_info.sessions_structure, session_base_info.check.base, log = log) { resources => val store: Sessions.Store = Sessions.store(options) @@ -569,8 +559,10 @@ /* session */ - def start_session(print_mode: List[String] = Nil, progress: Progress = new Progress): Session = - { + def start_session( + print_mode: List[String] = Nil, + progress: Progress = new Progress + ): Session = { val session = new Session(session_base_info.session, options, resources) progress.echo("Starting session " + session_base_info.session + " ...") @@ -591,8 +583,8 @@ theories: List[Document.Node.Name], files: List[Document.Node.Name], unicode_symbols: Boolean, - progress: Progress): Unit = - { + progress: Progress + ): Unit = { val loaded_theories = for (node_name <- theories) yield { @@ -609,55 +601,50 @@ val loaded = loaded_theories.length if (loaded > 1) progress.echo("Loading " + loaded + " theories ...") - state.change(st => - { - val (doc_blobs1, st1) = st.insert_required(id, theories).update_blobs(files) - val theory_edits = - for (theory <- loaded_theories) - yield { - val node_name = theory.node_name - val theory1 = theory.required(st1.is_required(node_name)) - val edits = theory1.node_edits(st1.theories.get(node_name)) - (edits, (node_name, theory1)) - } - val file_edits = - for { node_name <- files if doc_blobs1.changed(node_name) } - yield st1.blob_edits(node_name, st.blobs.get(node_name)) + state.change { st => + val (doc_blobs1, st1) = st.insert_required(id, theories).update_blobs(files) + val theory_edits = + for (theory <- loaded_theories) + yield { + val node_name = theory.node_name + val theory1 = theory.required(st1.is_required(node_name)) + val edits = theory1.node_edits(st1.theories.get(node_name)) + (edits, (node_name, theory1)) + } + val file_edits = + for { node_name <- files if doc_blobs1.changed(node_name) } + yield st1.blob_edits(node_name, st.blobs.get(node_name)) - session.update(doc_blobs1, theory_edits.flatMap(_._1) ::: file_edits.flatten) - st1.update_theories(theory_edits.map(_._2)) - }) + session.update(doc_blobs1, theory_edits.flatMap(_._1) ::: file_edits.flatten) + st1.update_theories(theory_edits.map(_._2)) + } } - def unload_theories(session: Session, id: UUID.T, theories: List[Document.Node.Name]): Unit = - { - state.change(st => - { + def unload_theories(session: Session, id: UUID.T, theories: List[Document.Node.Name]): Unit = { + state.change { st => val (edits, st1) = st.unload_theories(session, id, theories) session.update(st.doc_blobs, edits) st1 - }) + } } - def clean_theories(session: Session, id: UUID.T, theories: List[Document.Node.Name]): Unit = - { - state.change(st => - { + def clean_theories(session: Session, id: UUID.T, theories: List[Document.Node.Name]): Unit = { + state.change { st => val (edits1, st1) = st.unload_theories(session, id, theories) val ((_, _, edits2), st2) = st1.purge_theories(session, None) session.update(st.doc_blobs, edits1 ::: edits2) st2 - }) + } } - def purge_theories(session: Session, nodes: Option[List[Document.Node.Name]]) - : (List[Document.Node.Name], List[Document.Node.Name]) = - { - state.change_result(st => - { + def purge_theories( + session: Session, + nodes: Option[List[Document.Node.Name]] + ) : (List[Document.Node.Name], List[Document.Node.Name]) = { + state.change_result { st => val ((purged, retained, _), st1) = st.purge_theories(session, nodes) ((purged, retained), st1) - }) + } } } } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/PIDE/line.scala --- a/src/Pure/PIDE/line.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/PIDE/line.scala Fri Apr 01 23:26:19 2022 +0200 @@ -10,8 +10,7 @@ import scala.annotation.tailrec -object Line -{ +object Line { /* logical lines */ def normalize(text: String): String = @@ -23,19 +22,16 @@ /* position */ - object Position - { + object Position { val zero: Position = Position() val offside: Position = Position(line = -1) - object Ordering extends scala.math.Ordering[Position] - { + object Ordering extends scala.math.Ordering[Position] { def compare(p1: Position, p2: Position): Int = p1 compare p2 } } - sealed case class Position(line: Int = 0, column: Int = 0) - { + sealed case class Position(line: Int = 0, column: Int = 0) { def line1: Int = line + 1 def column1: Int = column + 1 def print: String = line1.toString + ":" + column1.toString @@ -59,14 +55,12 @@ /* range (right-open interval) */ - object Range - { + object Range { def apply(start: Position): Range = Range(start, start) val zero: Range = Range(Position.zero) } - sealed case class Range(start: Position, stop: Position) - { + sealed case class Range(start: Position, stop: Position) { if (start.compare(stop) > 0) error("Bad line range: " + start.print + ".." + stop.print) @@ -78,21 +72,18 @@ /* positions within document node */ - object Node_Position - { + object Node_Position { def offside(name: String): Node_Position = Node_Position(name, Position.offside) } - sealed case class Node_Position(name: String, pos: Position = Position.zero) - { + sealed case class Node_Position(name: String, pos: Position = Position.zero) { def line: Int = pos.line def line1: Int = pos.line1 def column: Int = pos.column def column1: Int = pos.column1 } - sealed case class Node_Range(name: String, range: Range = Range.zero) - { + sealed case class Node_Range(name: String, range: Range = Range.zero) { def start: Position = range.start def stop: Position = range.stop } @@ -100,8 +91,7 @@ /* document with newline as separator (not terminator) */ - object Document - { + object Document { def apply(text: String): Document = Document(logical_lines(text).map(s => Line(Library.isolate_substring(s)))) @@ -123,8 +113,7 @@ def text(lines: List[Line]): String = lines.mkString("", "\n", "") } - sealed case class Document(lines: List[Line]) - { + sealed case class Document(lines: List[Line]) { lazy val text_length: Text.Offset = Document.length(lines) def text_range: Text.Range = Text.Range(0, text_length) @@ -142,10 +131,8 @@ } override def hashCode(): Int = lines.hashCode - def position(text_offset: Text.Offset): Position = - { - @tailrec def move(i: Text.Offset, lines_count: Int, lines_rest: List[Line]): Position = - { + def position(text_offset: Text.Offset): Position = { + @tailrec def move(i: Text.Offset, lines_count: Int, lines_rest: List[Line]): Position = { lines_rest match { case Nil => require(i == 0, "bad Line.position.move"); Position(lines_count) case line :: ls => @@ -161,8 +148,7 @@ def range(text_range: Text.Range): Range = Range(position(text_range.start), position(text_range.stop)) - def offset(pos: Position): Option[Text.Offset] = - { + def offset(pos: Position): Option[Text.Offset] = { val l = pos.line val c = pos.column val n = lines.length @@ -178,16 +164,14 @@ else None } - def change(remove: Range, insert: String): Option[(List[Text.Edit], Document)] = - { + def change(remove: Range, insert: String): Option[(List[Text.Edit], Document)] = { for { edit_start <- offset(remove.start) if remove.stop == remove.start || offset(remove.stop).isDefined l1 = remove.start.line l2 = remove.stop.line if l1 <= l2 - (removed_text, new_lines) <- - { + (removed_text, new_lines) <- { val c1 = remove.start.column val c2 = remove.stop.column @@ -239,8 +223,7 @@ def apply(text: String): Line = if (text == "") empty else new Line(text) } -final class Line private(val text: String) -{ +final class Line private(val text: String) { require(text.forall(c => c != '\r' && c != '\n'), "bad line text") override def equals(that: Any): Boolean = diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/PIDE/markup.scala --- a/src/Pure/PIDE/markup.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/PIDE/markup.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,25 +7,21 @@ package isabelle -object Markup -{ +object Markup { /* elements */ - object Elements - { + object Elements { def apply(elems: Set[String]): Elements = new Elements(elems) def apply(elems: String*): Elements = apply(Set(elems: _*)) val empty: Elements = apply() val full: Elements = - new Elements(Set.empty) - { + new Elements(Set.empty) { override def apply(elem: String): Boolean = true override def toString: String = "Elements.full" } } - sealed class Elements private[Markup](private val rep: Set[String]) - { + sealed class Elements private[Markup](private val rep: Set[String]) { def apply(elem: String): Boolean = rep.contains(elem) def + (elem: String): Elements = new Elements(rep + elem) def ++ (elems: Elements): Elements = new Elements(rep ++ elems.rep) @@ -61,15 +57,13 @@ val Empty: Markup = Markup("", Nil) val Broken: Markup = Markup("broken", Nil) - class Markup_Elem(val name: String) - { + class Markup_Elem(val name: String) { def apply(props: Properties.T = Nil): Markup = Markup(name, props) def unapply(markup: Markup): Option[Properties.T] = if (markup.name == name) Some(markup.properties) else None } - class Markup_String(val name: String, prop: String) - { + class Markup_String(val name: String, prop: String) { val Prop: Properties.String = new Properties.String(prop) def apply(s: String): Markup = Markup(name, Prop(s)) @@ -78,8 +72,7 @@ def get(markup: Markup): String = unapply(markup).getOrElse("") } - class Markup_Int(val name: String, prop: String) - { + class Markup_Int(val name: String, prop: String) { val Prop: Properties.Int = new Properties.Int(prop) def apply(i: Int): Markup = Markup(name, Prop(i)) @@ -88,8 +81,7 @@ def get(markup: Markup): Int = unapply(markup).getOrElse(0) } - class Markup_Long(val name: String, prop: String) - { + class Markup_Long(val name: String, prop: String) { val Prop: Properties.Long = new Properties.Long(prop) def apply(i: Long): Markup = Markup(name, Prop(i)) @@ -114,13 +106,11 @@ val BINDING = "binding" val ENTITY = "entity" - object Entity - { + object Entity { val Def = new Markup_Long(ENTITY, "def") val Ref = new Markup_Long(ENTITY, "ref") - object Occ - { + object Occ { def unapply(markup: Markup): Option[Long] = Def.unapply(markup) orElse Ref.unapply(markup) } @@ -176,8 +166,7 @@ /* expression */ val EXPRESSION = "expression" - object Expression - { + object Expression { def unapply(markup: Markup): Option[String] = markup match { case Markup(EXPRESSION, props) => Some(Kind.get(props)) @@ -199,8 +188,7 @@ val Delimited = new Properties.Boolean("delimited") val LANGUAGE = "language" - object Language - { + object Language { val DOCUMENT = "document" val ML = "ML" val SML = "SML" @@ -218,8 +206,7 @@ case _ => None } - object Path - { + object Path { def unapply(markup: Markup): Option[Boolean] = markup match { case Language(PATH, _, _, delimited) => Some(delimited) @@ -250,8 +237,7 @@ val Indent = new Properties.Int("indent") val Width = new Properties.Int("width") - object Block - { + object Block { val name = "block" def apply(c: Boolean, i: Int): Markup = Markup(name, @@ -266,8 +252,7 @@ else None } - object Break - { + object Break { val name = "break" def apply(w: Int, i: Int): Markup = Markup(name, @@ -360,8 +345,7 @@ val PARAGRAPH = "paragraph" val TEXT_FOLD = "text_fold" - object Document_Tag extends Markup_String("document_tag", NAME) - { + object Document_Tag extends Markup_String("document_tag", NAME) { val IMPORTANT = "important" val UNIMPORTANT = "unimportant" } @@ -452,8 +436,7 @@ val CPU = new Properties.Double("cpu") val GC = new Properties.Double("gc") - object Timing_Properties - { + object Timing_Properties { def apply(timing: isabelle.Timing): Properties.T = Elapsed(timing.elapsed.seconds) ::: CPU(timing.cpu.seconds) ::: GC(timing.gc.seconds) @@ -470,8 +453,7 @@ val TIMING = "timing" - object Timing - { + object Timing { def apply(timing: isabelle.Timing): Markup = Markup(TIMING, Timing_Properties(timing)) def unapply(markup: Markup): Option[isabelle.Timing] = @@ -486,8 +468,7 @@ val Return_Code = new Properties.Int("return_code") - object Process_Result - { + object Process_Result { def apply(result: Process_Result): Properties.T = Return_Code(result.rc) ::: (if (result.timing.is_zero) Nil else Timing_Properties(result.timing)) @@ -594,8 +575,7 @@ val ML_PROFILING_ENTRY = "ML_profiling_entry" val ML_PROFILING = "ML_profiling" - object ML_Profiling - { + object ML_Profiling { def unapply_entry(tree: XML.Tree): Option[isabelle.ML_Profiling.Entry] = tree match { case XML.Elem(Markup(ML_PROFILING_ENTRY, List((NAME, name), (COUNT, Value.Long(count)))), _) => @@ -633,13 +613,11 @@ val FUNCTION = "function" - class Function(val name: String) - { + class Function(val name: String) { val PROPERTY: Properties.Entry = (FUNCTION, name) } - class Properties_Function(name: String) extends Function(name) - { + class Properties_Function(name: String) extends Function(name) { def unapply(props: Properties.T): Option[Properties.T] = props match { case PROPERTY :: args => Some(args) @@ -647,8 +625,7 @@ } } - class Name_Function(name: String) extends Function(name) - { + class Name_Function(name: String) extends Function(name) { def unapply(props: Properties.T): Option[String] = props match { case List(PROPERTY, (NAME, a)) => Some(a) @@ -656,8 +633,7 @@ } } - object ML_Statistics extends Function("ML_statistics") - { + object ML_Statistics extends Function("ML_statistics") { def unapply(props: Properties.T): Option[(Long, String)] = props match { case List(PROPERTY, ("pid", Value.Long(pid)), ("stats_dir", stats_dir)) => @@ -671,8 +647,7 @@ object Command_Timing extends Properties_Function("command_timing") object Theory_Timing extends Properties_Function("theory_timing") - object Session_Timing extends Properties_Function("session_timing") - { + object Session_Timing extends Properties_Function("session_timing") { val Threads = new Properties.Int("threads") } object Task_Statistics extends Properties_Function("task_statistics") @@ -684,8 +659,7 @@ object Assign_Update extends Function("assign_update") object Removed_Versions extends Function("removed_versions") - object Invoke_Scala extends Function("invoke_scala") - { + object Invoke_Scala extends Function("invoke_scala") { def unapply(props: Properties.T): Option[(String, String)] = props match { case List(PROPERTY, (NAME, name), (ID, id)) => Some((name, id)) @@ -693,8 +667,7 @@ } } - object Cancel_Scala extends Function("cancel_scala") - { + object Cancel_Scala extends Function("cancel_scala") { def unapply(props: Properties.T): Option[String] = props match { case List(PROPERTY, (ID, id)) => Some(id) @@ -717,8 +690,7 @@ /* debugger output */ val DEBUGGER_STATE = "debugger_state" - object Debugger_State - { + object Debugger_State { def unapply(props: Properties.T): Option[String] = props match { case List((FUNCTION, DEBUGGER_STATE), (NAME, name)) => Some(name) @@ -727,8 +699,7 @@ } val DEBUGGER_OUTPUT = "debugger_output" - object Debugger_Output - { + object Debugger_Output { def unapply(props: Properties.T): Option[String] = props match { case List((FUNCTION, DEBUGGER_OUTPUT), (NAME, name)) => Some(name) @@ -748,8 +719,7 @@ val SIMP_TRACE_IGNORE = "simp_trace_ignore" val SIMP_TRACE_CANCEL = "simp_trace_cancel" - object Simp_Trace_Cancel - { + object Simp_Trace_Cancel { def unapply(props: Properties.T): Option[Long] = props match { case (FUNCTION, SIMP_TRACE_CANCEL) :: Serial(i) => Some(i) @@ -760,14 +730,12 @@ /* XML data representation */ - def encode: XML.Encode.T[Markup] = (markup: Markup) => - { + def encode: XML.Encode.T[Markup] = { (markup: Markup) => import XML.Encode._ pair(string, properties)((markup.name, markup.properties)) } - def decode: XML.Decode.T[Markup] = (body: XML.Body) => - { + def decode: XML.Decode.T[Markup] = { (body: XML.Body) => import XML.Decode._ val (name, props) = pair(string, properties)(body) Markup(name, props) @@ -775,8 +743,7 @@ } -sealed case class Markup(name: String, properties: Properties.T) -{ +sealed case class Markup(name: String, properties: Properties.T) { def is_empty: Boolean = name.isEmpty def position_properties: Position.T = properties.filter(Markup.position_property) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/PIDE/markup_tree.scala --- a/src/Pure/PIDE/markup_tree.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/PIDE/markup_tree.scala Fri Apr 01 23:26:19 2022 +0200 @@ -13,8 +13,7 @@ import scala.annotation.tailrec -object Markup_Tree -{ +object Markup_Tree { /* construct trees */ val empty: Markup_Tree = new Markup_Tree(Branches.empty) @@ -40,8 +39,7 @@ /* tree building blocks */ - object Entry - { + object Entry { def apply(markup: Text.Markup, subtree: Markup_Tree): Entry = Entry(markup.range, List(markup.info), subtree) } @@ -49,12 +47,11 @@ sealed case class Entry( range: Text.Range, rev_markup: List[XML.Elem], - subtree: Markup_Tree) - { + subtree: Markup_Tree + ) { def markup: List[XML.Elem] = rev_markup.reverse - def filter_markup(elements: Markup.Elements): List[XML.Elem] = - { + def filter_markup(elements: Markup.Elements): List[XML.Elem] = { var result: List[XML.Elem] = Nil for (elem <- rev_markup if elements(elem.name)) result ::= elem @@ -65,8 +62,7 @@ def \ (markup: Text.Markup): Entry = copy(subtree = subtree + markup) } - object Branches - { + object Branches { type T = SortedMap[Text.Range, Entry] val empty: T = SortedMap.empty(Text.Range.Ordering) } @@ -84,33 +80,34 @@ case _ => (elems, body) } - private def make_trees(acc: (Int, List[Markup_Tree]), tree: XML.Tree): (Int, List[Markup_Tree]) = - { - val (offset, markup_trees) = acc + private def make_trees( + acc: (Int, List[Markup_Tree]), + tree: XML.Tree + ): (Int, List[Markup_Tree]) = { + val (offset, markup_trees) = acc - strip_elems(Nil, List(tree)) match { - case (Nil, body) => - (offset + XML.text_length(body), markup_trees) + strip_elems(Nil, List(tree)) match { + case (Nil, body) => + (offset + XML.text_length(body), markup_trees) - case (elems, body) => - val (end_offset, subtrees) = - body.foldLeft((offset, List.empty[Markup_Tree]))(make_trees) - if (offset == end_offset) acc - else { - val range = Text.Range(offset, end_offset) - val entry = Entry(range, elems, merge_disjoint(subtrees)) - (end_offset, new Markup_Tree(Branches.empty, entry) :: markup_trees) - } - } + case (elems, body) => + val (end_offset, subtrees) = + body.foldLeft((offset, List.empty[Markup_Tree]))(make_trees) + if (offset == end_offset) acc + else { + val range = Text.Range(offset, end_offset) + val entry = Entry(range, elems, merge_disjoint(subtrees)) + (end_offset, new Markup_Tree(Branches.empty, entry) :: markup_trees) + } } + } def from_XML(body: XML.Body): Markup_Tree = merge_disjoint(body.foldLeft((0, List.empty[Markup_Tree]))(make_trees)._2) } -final class Markup_Tree private(val branches: Markup_Tree.Branches.T) -{ +final class Markup_Tree private(val branches: Markup_Tree.Branches.T) { import Markup_Tree._ private def this(branches: Markup_Tree.Branches.T, entry: Markup_Tree.Entry) = @@ -135,8 +132,7 @@ def is_empty: Boolean = branches.isEmpty - def + (new_markup: Text.Markup): Markup_Tree = - { + def + (new_markup: Text.Markup): Markup_Tree = { val new_range = new_markup.range branches.get(new_range) match { @@ -161,8 +157,7 @@ } } - def merge(other: Markup_Tree, root_range: Text.Range, elements: Markup.Elements): Markup_Tree = - { + def merge(other: Markup_Tree, root_range: Text.Range, elements: Markup.Elements): Markup_Tree = { def merge_trees(tree1: Markup_Tree, tree2: Markup_Tree): Markup_Tree = tree2.branches.foldLeft(tree1) { case (tree, (range, entry)) => @@ -183,11 +178,13 @@ } } - def cumulate[A](root_range: Text.Range, root_info: A, elements: Markup.Elements, - result: (A, Text.Markup) => Option[A]): List[Text.Info[A]] = - { - def results(x: A, entry: Entry): Option[A] = - { + def cumulate[A]( + root_range: Text.Range, + root_info: A, + elements: Markup.Elements, + result: (A, Text.Markup) => Option[A] + ): List[Text.Info[A]] = { + def results(x: A, entry: Entry): Option[A] = { var y = x var changed = false for { @@ -199,8 +196,8 @@ def traverse( last: Text.Offset, - stack: List[(Text.Info[A], List[(Text.Range, Entry)])]): List[Text.Info[A]] = - { + stack: List[(Text.Info[A], List[(Text.Range, Entry)])] + ): List[Text.Info[A]] = { stack match { case (parent, (range, entry) :: more) :: rest => val subrange = range.restrict(root_range) @@ -232,8 +229,7 @@ List((Text.Info(root_range, root_info), overlapping(root_range).toList))) } - def to_XML(root_range: Text.Range, text: CharSequence, elements: Markup.Elements): XML.Body = - { + def to_XML(root_range: Text.Range, text: CharSequence, elements: Markup.Elements): XML.Body = { def make_text(start: Text.Offset, stop: Text.Offset): XML.Body = if (start == stop) Nil else List(XML.Text(text.subSequence(start, stop).toString)) @@ -246,9 +242,11 @@ else List(XML.Wrapped_Elem(elem.markup, elem.body, b)) } - def make_body(elem_range: Text.Range, elem_markup: List[XML.Elem], entries: Branches.T) - : XML.Body = - { + def make_body( + elem_range: Text.Range, + elem_markup: List[XML.Elem], + entries: Branches.T + ) : XML.Body = { val body = new mutable.ListBuffer[XML.Tree] var last = elem_range.start for ((range, entry) <- entries) { diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/PIDE/protocol.scala --- a/src/Pure/PIDE/protocol.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/PIDE/protocol.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Protocol -{ +object Protocol { /* markers for inlined messages */ val Loading_Theory_Marker = Protocol_Message.Marker("loading_theory") @@ -23,8 +22,7 @@ /* batch build */ - object Loading_Theory - { + object Loading_Theory { def unapply(props: Properties.T): Option[(Document.Node.Name, Document_ID.Exec)] = (props, props, props) match { case (Markup.Name(name), Position.File(file), Position.Id(id)) @@ -38,8 +36,7 @@ /* document editing */ - object Commands_Accepted - { + object Commands_Accepted { def unapply(text: String): Option[List[Document_ID.Command]] = try { Some(space_explode(',', text).map(Value.Long.parse)) } catch { case ERROR(_) => None } @@ -47,11 +44,10 @@ val message: XML.Elem = XML.elem(Markup.STATUS, List(XML.elem(Markup.ACCEPTED))) } - object Assign_Update - { - def unapply(text: String) - : Option[(Document_ID.Version, List[String], Document.Assign_Update)] = - { + object Assign_Update { + def unapply( + text: String + ) : Option[(Document_ID.Version, List[String], Document.Assign_Update)] = { try { import XML.Decode._ def decode_upd(body: XML.Body): (Long, List[Long]) = @@ -68,8 +64,7 @@ } } - object Removed - { + object Removed { def unapply(text: String): Option[List[Document_ID.Version]] = try { import XML.Decode._ @@ -84,8 +79,7 @@ /* command timing */ - object Command_Timing - { + object Command_Timing { def unapply(props: Properties.T): Option[(Properties.T, Document_ID.Generic, isabelle.Timing)] = props match { case Markup.Command_Timing(args) => @@ -100,8 +94,7 @@ /* theory timing */ - object Theory_Timing - { + object Theory_Timing { def unapply(props: Properties.T): Option[(String, isabelle.Timing)] = props match { case Markup.Theory_Timing(args) => @@ -182,8 +175,8 @@ pos: Position.T = Position.none, margin: Double = Pretty.default_margin, breakgain: Double = Pretty.default_breakgain, - metric: Pretty.Metric = Pretty.Default_Metric): String = - { + metric: Pretty.Metric = Pretty.Default_Metric + ): String = { val text1 = if (heading) { val h = @@ -212,8 +205,7 @@ /* ML profiling */ - object ML_Profiling - { + object ML_Profiling { def unapply(msg: XML.Tree): Option[isabelle.ML_Profiling.Report] = msg match { case XML.Elem(_, List(tree)) if is_tracing(msg) => @@ -225,8 +217,7 @@ /* export */ - object Export - { + object Export { sealed case class Args( id: Option[String] = None, serial: Long = 0L, @@ -234,8 +225,8 @@ name: String, executable: Boolean = false, compress: Boolean = true, - strict: Boolean = true) - { + strict: Boolean = true + ) { def compound_name: String = isabelle.Export.compound_name(theory_name, name) } @@ -259,8 +250,7 @@ /* breakpoints */ - object ML_Breakpoint - { + object ML_Breakpoint { def unapply(tree: XML.Tree): Option[Long] = tree match { case XML.Elem(Markup(Markup.ML_BREAKPOINT, Markup.Serial(breakpoint)), _) => Some(breakpoint) @@ -271,8 +261,7 @@ /* dialogs */ - object Dialog_Args - { + object Dialog_Args { def unapply(props: Properties.T): Option[(Document_ID.Generic, Long, String)] = (props, props, props) match { case (Position.Id(id), Markup.Serial(serial), Markup.Result(result)) => @@ -281,8 +270,7 @@ } } - object Dialog - { + object Dialog { def unapply(tree: XML.Tree): Option[(Document_ID.Generic, Long, String)] = tree match { case XML.Elem(Markup(Markup.DIALOG, Dialog_Args(id, serial, result)), _) => @@ -291,10 +279,8 @@ } } - object Dialog_Result - { - def apply(id: Document_ID.Generic, serial: Long, result: String): XML.Elem = - { + object Dialog_Result { + def apply(id: Document_ID.Generic, serial: Long, result: String): XML.Elem = { val props = Position.Id(id) ::: Markup.Serial(serial) XML.Elem(Markup(Markup.RESULT, props), List(XML.Text(result))) } @@ -308,8 +294,7 @@ } -trait Protocol -{ +trait Protocol { /* protocol commands */ def protocol_command_raw(name: String, args: List[Bytes]): Unit @@ -334,16 +319,16 @@ def define_blob(digest: SHA1.Digest, bytes: Bytes): Unit = protocol_command_raw("Document.define_blob", List(Bytes(digest.toString), bytes)) - private def encode_command(resources: Resources, command: Command) - : (String, String, String, String, String, List[String]) = - { + private def encode_command( + resources: Resources, + command: Command + ) : (String, String, String, String, String, List[String]) = { import XML.Encode._ val parents = command.theory_parents(resources).map(name => File.standard_url(name.node)) val parents_yxml = Symbol.encode_yxml(list(string)(parents)) - val blobs_yxml = - { + val blobs_yxml = { val encode_blob: T[Exn.Result[Command.Blob]] = variant(List( { case Exn.Res(Command.Blob(a, b, c)) => @@ -354,8 +339,7 @@ Symbol.encode_yxml(pair(list(encode_blob), int)(command.blobs, command.blobs_index)) } - val toks_yxml = - { + val toks_yxml = { val encode_tok: T[Token] = (tok => pair(int, int)((tok.kind.id, Symbol.length(tok.source)))) Symbol.encode_yxml(list(encode_tok)(command.span.content)) } @@ -365,8 +349,7 @@ blobs_yxml, toks_yxml, toks_sources) } - def define_command(resources: Resources, command: Command): Unit = - { + def define_command(resources: Resources, command: Command): Unit = { val (command_id, name, parents_yxml, blobs_yxml, toks_yxml, toks_sources) = encode_command(resources, command) protocol_command_args( @@ -375,10 +358,8 @@ } def define_commands(resources: Resources, commands: List[Command]): Unit = - { protocol_command_args("Document.define_commands", - commands.map(command => - { + commands.map { command => import XML.Encode._ val (command_id, name, parents_yxml, blobs_yxml, toks_yxml, toks_sources) = encode_command(resources, command) @@ -386,11 +367,9 @@ pair(string, pair(string, pair(string, pair(string, pair(string, list(string))))))( command_id, (name, (parents_yxml, (blobs_yxml, (toks_yxml, toks_sources))))) YXML.string_of_body(body) - })) - } + }) - def define_commands_bulk(resources: Resources, commands: List[Command]): Unit = - { + def define_commands_bulk(resources: Resources, commands: List[Command]): Unit = { val (irregular, regular) = commands.partition(command => YXML.detect(command.source)) irregular.foreach(define_command(resources, _)) regular match { @@ -412,16 +391,17 @@ /* document versions */ - def update(old_id: Document_ID.Version, new_id: Document_ID.Version, - edits: List[Document.Edit_Command], consolidate: List[Document.Node.Name]): Unit = - { - val consolidate_yxml = - { + def update( + old_id: Document_ID.Version, + new_id: Document_ID.Version, + edits: List[Document.Edit_Command], + consolidate: List[Document.Node.Name] + ): Unit = { + val consolidate_yxml = { import XML.Encode._ Symbol.encode_yxml(list(string)(consolidate.map(_.node))) } - val edits_yxml = - { + val edits_yxml = { import XML.Encode._ def id: T[Command] = (cmd => long(cmd.id)) def encode_edit(name: Document.Node.Name) @@ -446,8 +426,7 @@ Document_ID(old_id) :: Document_ID(new_id) :: consolidate_yxml :: edits_yxml) } - def remove_versions(versions: List[Document.Version]): Unit = - { + def remove_versions(versions: List[Document.Version]): Unit = { val versions_yxml = { import XML.Encode._ Symbol.encode_yxml(list(long)(versions.map(_.id))) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/PIDE/protocol_handlers.scala --- a/src/Pure/PIDE/protocol_handlers.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/PIDE/protocol_handlers.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,20 +7,18 @@ package isabelle -object Protocol_Handlers -{ +object Protocol_Handlers { private def err_handler(exn: Throwable, name: String): Nothing = error("Failed to initialize protocol handler: " + quote(name) + "\n" + Exn.message(exn)) sealed case class State( session: Session, handlers: Map[String, Session.Protocol_Handler] = Map.empty, - functions: Map[String, Session.Protocol_Function] = Map.empty) - { + functions: Map[String, Session.Protocol_Function] = Map.empty + ) { def get(name: String): Option[Session.Protocol_Handler] = handlers.get(name) - def init(handler: Session.Protocol_Handler): State = - { + def init(handler: Session.Protocol_Handler): State = { val name = handler.getClass.getName try { if (handlers.isDefinedAt(name)) error("Duplicate protocol handler: " + name) @@ -34,8 +32,7 @@ catch { case exn: Throwable => err_handler(exn, name) } } - def init(name: String): State = - { + def init(name: String): State = { val handler = try { Class.forName(name).getDeclaredConstructor().newInstance() @@ -58,8 +55,7 @@ case _ => false } - def exit(): State = - { + def exit(): State = { for ((_, handler) <- handlers) handler.exit() copy(handlers = Map.empty, functions = Map.empty) } @@ -69,8 +65,7 @@ new Protocol_Handlers(session) } -class Protocol_Handlers private(session: Session) -{ +class Protocol_Handlers private(session: Session) { private val state = Synchronized(Protocol_Handlers.State(session)) def prover_options(options: Options): Options = diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/PIDE/protocol_message.scala --- a/src/Pure/PIDE/protocol_message.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/PIDE/protocol_message.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,20 +7,17 @@ package isabelle -object Protocol_Message -{ +object Protocol_Message { /* message markers */ - object Marker - { + object Marker { def apply(a: String): Marker = new Marker { override def name: String = a } def test(line: String): Boolean = line.startsWith("\f") } - abstract class Marker private - { + abstract class Marker private { def name: String val prefix: String = "\f" + name + " = " diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/PIDE/prover.scala --- a/src/Pure/PIDE/prover.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/PIDE/prover.scala Fri Apr 01 23:26:19 2022 +0200 @@ -11,23 +11,20 @@ import java.io.{InputStream, OutputStream, BufferedOutputStream, IOException} -object Prover -{ +object Prover { /* messages */ sealed abstract class Message type Receiver = Message => Unit - class Input(val name: String, val args: List[String]) extends Message - { + class Input(val name: String, val args: List[String]) extends Message { override def toString: String = XML.Elem(Markup(Markup.PROVER_COMMAND, List((Markup.NAME, name))), args.flatMap(s => List(XML.newline, XML.elem(Markup.PROVER_ARG, YXML.parse_body(s))))).toString } - class Output(val message: XML.Elem) extends Message - { + class Output(val message: XML.Elem) extends Message { def kind: String = message.markup.name def properties: Properties.T = message.markup.properties def body: XML.Body = message.body @@ -41,8 +38,7 @@ def is_report: Boolean = kind == Markup.REPORT def is_syslog: Boolean = is_init || is_exit || is_system || is_stderr - override def toString: String = - { + override def toString: String = { val res = if (is_status || is_report) message.body.map(_.toString).mkString else Pretty.string_of(message.body, metric = Symbol.Metric) @@ -65,8 +61,7 @@ } class Protocol_Output(props: Properties.T, val chunks: List[Bytes]) - extends Output(XML.Elem(Markup(Markup.PROTOCOL, props), Nil)) - { + extends Output(XML.Elem(Markup(Markup.PROTOCOL, props), Nil)) { def chunk: Bytes = the_chunk(chunks, toString) lazy val text: String = chunk.text } @@ -77,29 +72,25 @@ receiver: Prover.Receiver, cache: XML.Cache, channel: System_Channel, - process: Bash.Process) extends Protocol -{ + process: Bash.Process +) extends Protocol { /** receiver output **/ - private def system_output(text: String): Unit = - { + private def system_output(text: String): Unit = { receiver(new Prover.Output(XML.Elem(Markup(Markup.SYSTEM, Nil), List(XML.Text(text))))) } - private def protocol_output(props: Properties.T, chunks: List[Bytes]): Unit = - { + private def protocol_output(props: Properties.T, chunks: List[Bytes]): Unit = { receiver(new Prover.Protocol_Output(props, chunks)) } - private def output(kind: String, props: Properties.T, body: XML.Body): Unit = - { + private def output(kind: String, props: Properties.T, body: XML.Body): Unit = { val main = XML.Elem(Markup(kind, props), Protocol_Message.clean_reports(body)) val reports = Protocol_Message.reports(props, body) for (msg <- main :: reports) receiver(new Prover.Output(cache.elem(msg))) } - private def exit_message(result: Process_Result): Unit = - { + private def exit_message(result: Process_Result): Unit = { output(Markup.EXIT, Markup.Process_Result(result), List(XML.Text(result.print_return_code))) } @@ -115,20 +106,17 @@ Process_Result(rc, timing = timing) } - private def terminate_process(): Unit = - { + private def terminate_process(): Unit = { try { process.terminate() } catch { case exn @ ERROR(_) => system_output("Failed to terminate prover process: " + exn.getMessage) } } - private val process_manager = Isabelle_Thread.fork(name = "process_manager") - { + private val process_manager = Isabelle_Thread.fork(name = "process_manager") { val stdout = physical_output(false) - val (startup_failed, startup_errors) = - { + val (startup_failed, startup_errors) = { var finished: Option[Boolean] = None val result = new StringBuilder(100) while (finished.isEmpty && (process.stderr.ready || !process_result.is_finished)) { @@ -174,8 +162,7 @@ def join(): Unit = process_manager.join() - def terminate(): Unit = - { + def terminate(): Unit = { system_output("Terminating prover process") command_input_close() @@ -197,8 +184,7 @@ private def command_input_close(): Unit = command_input.foreach(_.shutdown()) - private def command_input_init(raw_stream: OutputStream): Unit = - { + private def command_input_init(raw_stream: OutputStream): Unit = { val name = "command_input" val stream = new BufferedOutputStream(raw_stream) command_input = @@ -223,8 +209,7 @@ /* physical output */ - private def physical_output(err: Boolean): Thread = - { + private def physical_output(err: Boolean): Thread = { val (name, reader, markup) = if (err) ("standard_error", process.stderr, Markup.STDERR) else ("standard_output", process.stdout, Markup.STDOUT) @@ -261,8 +246,7 @@ /* message output */ - private def message_output(stream: InputStream): Thread = - { + private def message_output(stream: InputStream): Thread = { def decode_chunk(chunk: Bytes): XML.Body = Symbol.decode_yxml_failsafe(chunk.text, cache = cache) @@ -312,8 +296,7 @@ case _ => error("Inactive prover input thread for command " + quote(name)) } - def protocol_command_args(name: String, args: List[String]): Unit = - { + def protocol_command_args(name: String, args: List[String]): Unit = { receiver(new Prover.Input(name, args)) protocol_command_raw(name, args.map(Bytes(_))) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/PIDE/query_operation.scala --- a/src/Pure/PIDE/query_operation.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/PIDE/query_operation.scala Fri Apr 01 23:26:19 2022 +0200 @@ -8,17 +8,14 @@ package isabelle -object Query_Operation -{ - object Status extends Enumeration - { +object Query_Operation { + object Status extends Enumeration { val WAITING = Value("waiting") val RUNNING = Value("running") val FINISHED = Value("finished") } - object State - { + object State { val empty: State = State() def make(command: Command, query: List[String]): State = @@ -43,8 +40,8 @@ editor_context: Editor_Context, operation_name: String, consume_status: Query_Operation.Status.Value => Unit, - consume_output: (Document.Snapshot, Command.Results, XML.Body) => Unit) -{ + consume_output: (Document.Snapshot, Command.Results, XML.Body) => Unit +) { private val print_function = operation_name + "_query" @@ -54,8 +51,7 @@ def get_location: Option[Command] = current_state.value.location - private def remove_overlay(): Unit = - { + private def remove_overlay(): Unit = { val state = current_state.value for (command <- state.location) { editor.remove_overlay(command, print_function, state.instance :: state.query) @@ -65,8 +61,7 @@ /* content update */ - private def content_update(): Unit = - { + private def content_update(): Unit = { editor.require_dispatcher {} @@ -94,8 +89,7 @@ /* resolve sendback: static command id */ - def resolve_sendback(body: XML.Body): XML.Body = - { + def resolve_sendback(body: XML.Body): XML.Body = { state0.location match { case None => body case Some(command) => @@ -176,8 +170,7 @@ def cancel_query(): Unit = editor.require_dispatcher { editor.session.cancel_exec(current_state.value.exec_id) } - def apply_query(query: List[String]): Unit = - { + def apply_query(query: List[String]): Unit = { editor.require_dispatcher {} editor.current_node_snapshot(editor_context) match { @@ -200,8 +193,7 @@ } } - def locate_query(): Unit = - { + def locate_query(): Unit = { editor.require_dispatcher {} val state = current_state.value @@ -229,13 +221,11 @@ } } - def activate(): Unit = - { + def activate(): Unit = { editor.session.commands_changed += main } - def deactivate(): Unit = - { + def deactivate(): Unit = { editor.session.commands_changed -= main remove_overlay() current_state.change(_ => Query_Operation.State.empty) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/PIDE/rendering.scala --- a/src/Pure/PIDE/rendering.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/PIDE/rendering.scala Fri Apr 01 23:26:19 2022 +0200 @@ -15,12 +15,10 @@ -object Rendering -{ +object Rendering { /* color */ - object Color extends Enumeration - { + object Color extends Enumeration { // background val unprocessed1, running1, canceled, bad, intensify, entity, active, active_result, markdown_bullet1, markdown_bullet2, markdown_bullet3, markdown_bullet4 = Value @@ -97,8 +95,7 @@ legacy_pri -> Color.legacy_message, error_pri -> Color.error_message) - def output_messages(results: Command.Results): List[XML.Elem] = - { + def output_messages(results: Command.Results): List[XML.Elem] = { val (states, other) = results.iterator.map(_._2).filterNot(Protocol.is_result).toList .partition(Protocol.is_state) @@ -170,23 +167,20 @@ /* entity focus */ - object Focus - { + object Focus { def apply(ids: Set[Long]): Focus = new Focus(ids) val empty: Focus = apply(Set.empty) def make(args: List[Text.Info[Focus]]): Focus = args.foldLeft(empty) { case (focus1, Text.Info(_, focus2)) => focus1 ++ focus2 } val full: Focus = - new Focus(Set.empty) - { + new Focus(Set.empty) { override def apply(id: Long): Boolean = true override def toString: String = "Focus.full" } } - sealed class Focus private[Rendering](protected val rep: Set[Long]) - { + sealed class Focus private[Rendering](protected val rep: Set[Long]) { def defined: Boolean = rep.nonEmpty def apply(id: Long): Boolean = rep.contains(id) def + (id: Long): Focus = if (rep.contains(id)) this else new Focus(rep + id) @@ -265,8 +259,8 @@ class Rendering( val snapshot: Document.Snapshot, val options: Options, - val session: Session) -{ + val session: Session +) { override def toString: String = "Rendering(" + snapshot.toString + ")" def get_text(range: Text.Range): Option[String] = None @@ -274,8 +268,7 @@ /* caret */ - def before_caret_range(caret: Text.Offset): Text.Range = - { + def before_caret_range(caret: Text.Offset): Text.Range = { val former_caret = snapshot.revert(caret) snapshot.convert(Text.Range(former_caret - 1, former_caret)) } @@ -302,8 +295,8 @@ history: Completion.History, unicode: Boolean, completed_range: Option[Text.Range], - caret_range: Text.Range): (Boolean, Option[Completion.Result]) = - { + caret_range: Text.Range + ): (Boolean, Option[Completion.Result]) = { semantic_completion(completed_range, caret_range) match { case Some(Text.Info(_, Completion.No_Completion)) => (true, None) case Some(Text.Info(range, names: Completion.Names)) => @@ -347,10 +340,8 @@ case _ => None }).headOption.map({ case Text.Info(_, (delimited, range)) => Text.Info(range, delimited) }) - def path_completion(caret: Text.Offset): Option[Completion.Result] = - { - def complete(text: String): List[(String, List[String])] = - { + def path_completion(caret: Text.Offset): Option[Completion.Result] = { + def complete(text: String): List[(String, List[String])] = { try { val path = Path.explode(text) val (dir, base_name) = @@ -414,8 +405,7 @@ spell_checker_include ++ Markup.Elements(space_explode(',', options.string("spell_checker_exclude")): _*) - def spell_checker(range: Text.Range): List[Text.Info[Text.Range]] = - { + def spell_checker(range: Text.Range): List[Text.Info[Text.Range]] = { val result = snapshot.select(range, spell_checker_elements, _ => { @@ -435,9 +425,11 @@ /* text background */ - def background(elements: Markup.Elements, range: Text.Range, focus: Rendering.Focus) - : List[Text.Info[Rendering.Color.Value]] = - { + def background( + elements: Markup.Elements, + range: Text.Range, + focus: Rendering.Focus + ) : List[Text.Info[Rendering.Color.Value]] = { for { Text.Info(r, result) <- snapshot.cumulate[(List[Markup], Option[Rendering.Color.Value])]( @@ -521,8 +513,7 @@ /* caret focus */ - def caret_focus(caret_range: Text.Range, defs_range: Text.Range): Rendering.Focus = - { + def caret_focus(caret_range: Text.Range, defs_range: Text.Range): Rendering.Focus = { val focus = entity_focus_defs(caret_range) if (focus.defined) focus else if (defs_range == Text.Range.offside) Rendering.Focus.empty @@ -534,8 +525,7 @@ } } - def caret_focus_ranges(caret_range: Text.Range, defs_range: Text.Range): List[Text.Range] = - { + def caret_focus_ranges(caret_range: Text.Range, defs_range: Text.Range): List[Text.Range] = { val focus = caret_focus(caret_range, defs_range) if (focus.defined) { snapshot.cumulate[Boolean](defs_range, false, Rendering.entity_elements, _ => @@ -550,9 +540,10 @@ /* messages */ - def message_underline_color(elements: Markup.Elements, range: Text.Range) - : List[Text.Info[Rendering.Color.Value]] = - { + def message_underline_color( + elements: Markup.Elements, + range: Text.Range + ) : List[Text.Info[Rendering.Color.Value]] = { val results = snapshot.cumulate[Int](range, 0, elements, _ => { @@ -564,8 +555,7 @@ } yield Text.Info(r, color) } - def text_messages(range: Text.Range): List[Text.Info[XML.Elem]] = - { + def text_messages(range: Text.Range): List[Text.Info[XML.Elem]] = { val results = snapshot.cumulate[Vector[Command.Results.Entry]]( range, Vector.empty, Rendering.message_elements, command_states => @@ -577,8 +567,7 @@ }) var seen_serials = Set.empty[Long] - def seen(i: Long): Boolean = - { + def seen(i: Long): Boolean = { val b = seen_serials(i) seen_serials += i b @@ -598,17 +587,15 @@ range: Text.Range, timing: Timing = Timing.zero, messages: List[(Long, XML.Tree)] = Nil, - rev_infos: List[(Boolean, XML.Tree)] = Nil) - { + rev_infos: List[(Boolean, XML.Tree)] = Nil + ) { def + (t: Timing): Tooltip_Info = copy(timing = timing + t) - def + (r0: Text.Range, serial: Long, tree: XML.Tree): Tooltip_Info = - { + def + (r0: Text.Range, serial: Long, tree: XML.Tree): Tooltip_Info = { val r = snapshot.convert(r0) if (range == r) copy(messages = (serial -> tree) :: messages) else copy(range = r, messages = List(serial -> tree)) } - def + (r0: Text.Range, important: Boolean, tree: XML.Tree): Tooltip_Info = - { + def + (r0: Text.Range, important: Boolean, tree: XML.Tree): Tooltip_Info = { val r = snapshot.convert(r0) if (range == r) copy(rev_infos = (important -> tree) :: rev_infos) else copy (range = r, rev_infos = List(important -> tree)) @@ -630,8 +617,7 @@ def perhaps_append_file(node_name: Document.Node.Name, name: String): String = if (Path.is_valid(name)) session.resources.append(node_name, Path.explode(name)) else name - def tooltips(elements: Markup.Elements, range: Text.Range): Option[Text.Info[List[XML.Tree]]] = - { + def tooltips(elements: Markup.Elements, range: Text.Range): Option[Text.Info[List[XML.Tree]]] = { val results = snapshot.cumulate[Tooltip_Info](range, Tooltip_Info(range), elements, command_states => { @@ -727,8 +713,7 @@ /* command status overview */ - def overview_color(range: Text.Range): Option[Rendering.Color.Value] = - { + def overview_color(range: Text.Range): Option[Rendering.Color.Value] = { if (snapshot.is_outdated) None else { val results = @@ -762,8 +747,7 @@ /* meta data */ - def meta_data(range: Text.Range): Properties.T = - { + def meta_data(range: Text.Range): Properties.T = { val results = snapshot.cumulate[Properties.T](range, Nil, Rendering.meta_data_elements, _ => { @@ -777,8 +761,7 @@ /* document tags */ - def document_tags(range: Text.Range): List[String] = - { + def document_tags(range: Text.Range): List[String] = { val results = snapshot.cumulate[List[String]](range, Nil, Rendering.document_tag_elements, _ => { diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/PIDE/resources.scala --- a/src/Pure/PIDE/resources.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/PIDE/resources.scala Fri Apr 01 23:26:19 2022 +0200 @@ -12,8 +12,7 @@ import java.io.{File => JFile} -object Resources -{ +object Resources { def empty: Resources = new Resources(Sessions.Structure.empty, Sessions.Structure.empty.bootstrap) @@ -31,15 +30,14 @@ val sessions_structure: Sessions.Structure, val session_base: Sessions.Base, val log: Logger = No_Logger, - command_timings: List[Properties.T] = Nil) -{ + command_timings: List[Properties.T] = Nil +) { resources => /* init session */ - def init_session_yxml: String = - { + def init_session_yxml: String = { import XML.Encode._ YXML.string_of_body( @@ -78,8 +76,7 @@ def append(node_name: Document.Node.Name, source_path: Path): String = append(node_name.master_dir, source_path) - def file_node(file: Path, dir: String = "", theory: String = ""): Document.Node.Name = - { + def file_node(file: Path, dir: String = "", theory: String = ""): Document.Node.Name = { val node = append(dir, file) val master_dir = append(dir, file.dir) Document.Node.Name(node, master_dir, theory) @@ -91,8 +88,7 @@ /* source files of Isabelle/ML bootstrap */ - def source_file(raw_name: String): Option[String] = - { + def source_file(raw_name: String): Option[String] = { if (Path.is_wellformed(raw_name)) { if (Path.is_valid(raw_name)) { def check(p: Path): Option[Path] = if (p.is_file) Some(p) else None @@ -115,9 +111,10 @@ /* theory files */ - def load_commands(syntax: Outer_Syntax, name: Document.Node.Name) - : () => List[Command_Span.Span] = - { + def load_commands( + syntax: Outer_Syntax, + name: Document.Node.Name + ) : () => List[Command_Span.Span] = { val (is_utf8, raw_text) = with_thy_reader(name, reader => (Scan.reader_is_utf8(reader), reader.source.toString)) () => @@ -130,16 +127,17 @@ } } - def loaded_files(syntax: Outer_Syntax, name: Document.Node.Name, spans: List[Command_Span.Span]) - : List[Path] = - { + def loaded_files( + syntax: Outer_Syntax, + name: Document.Node.Name, + spans: List[Command_Span.Span] + ) : List[Path] = { val dir = name.master_dir_path for { span <- spans; file <- span.loaded_files(syntax).files } yield (dir + Path.explode(file)).expand } - def pure_files(syntax: Outer_Syntax): List[Path] = - { + def pure_files(syntax: Outer_Syntax): List[Path] = { val pure_dir = Path.explode("~~/src/Pure") for { (name, theory) <- Thy_Header.ml_roots @@ -154,8 +152,7 @@ theory else Long_Name.qualify(qualifier, theory) - def find_theory_node(theory: String): Option[Document.Node.Name] = - { + def find_theory_node(theory: String): Option[Document.Node.Name] = { val thy_file = Path.basic(Long_Name.base_name(theory)).thy val session = session_base.theory_qualifier(theory) val dirs = @@ -167,8 +164,7 @@ case dir if (dir + thy_file).is_file => file_node(dir + thy_file, theory = theory) }) } - def import_name(qualifier: String, dir: String, s: String): Document.Node.Name = - { + def import_name(qualifier: String, dir: String, s: String): Document.Node.Name = { val theory = theory_name(qualifier, Thy_Header.import_name(s)) def theory_node = file_node(Path.explode(s).thy, dir = dir, theory = theory) @@ -188,8 +184,7 @@ def import_name(info: Sessions.Info, s: String): Document.Node.Name = import_name(info.name, info.dir.implode, s) - def find_theory(file: JFile): Option[Document.Node.Name] = - { + def find_theory(file: JFile): Option[Document.Node.Name] = { for { qualifier <- session_base.session_directories.get(File.canonical(file).getParentFile) theory_base <- proper_string(Thy_Header.theory_name(file.getName)) @@ -199,8 +194,7 @@ } yield theory_node } - def complete_import_name(context_name: Document.Node.Name, s: String): List[String] = - { + def complete_import_name(context_name: Document.Node.Name, s: String): List[String] = { val context_session = session_base.theory_qualifier(context_name) val context_dir = try { Some(context_name.master_dir_path) } @@ -216,8 +210,7 @@ }).toList.sorted } - def with_thy_reader[A](name: Document.Node.Name, f: Reader[Char] => A): A = - { + def with_thy_reader[A](name: Document.Node.Name, f: Reader[Char] => A): A = { val path = name.path if (path.is_file) using(Scan.byte_reader(path.file))(f) else if (name.node == name.theory) @@ -225,9 +218,12 @@ else error ("Cannot load theory file " + path) } - def check_thy(node_name: Document.Node.Name, reader: Reader[Char], - command: Boolean = true, strict: Boolean = true): Document.Node.Header = - { + def check_thy( + node_name: Document.Node.Name, + reader: Reader[Char], + command: Boolean = true, + strict: Boolean = true + ): Document.Node.Header = { if (node_name.is_theory && reader.source.length > 0) { try { val header = Thy_Header.read(node_name, reader, command = command, strict = strict) @@ -248,8 +244,7 @@ /* special header */ - def special_header(name: Document.Node.Name): Option[Document.Node.Header] = - { + def special_header(name: Document.Node.Name): Option[Document.Node.Header] = { val imports = if (name.theory == Sessions.root_name) List(import_name(name, Sessions.theory_name)) else if (Thy_Header.is_ml_root(name.theory)) List(import_name(name, Thy_Header.ML_BOOTSTRAP)) @@ -291,9 +286,10 @@ progress: Progress = new Progress): Dependencies[Unit] = Dependencies.empty[Unit].require_thys((), thys, progress = progress) - def session_dependencies(info: Sessions.Info, progress: Progress = new Progress) - : Dependencies[Options] = - { + def session_dependencies( + info: Sessions.Info, + progress: Progress = new Progress + ) : Dependencies[Options] = { info.theories.foldLeft(Dependencies.empty[Options]) { case (dependencies, (options, thys)) => dependencies.require_thys(options, @@ -302,8 +298,7 @@ } } - object Dependencies - { + object Dependencies { def empty[A]: Dependencies[A] = new Dependencies[A](Nil, Map.empty) private def show_path(names: List[Document.Node.Name]): String = @@ -319,16 +314,15 @@ final class Dependencies[A] private( rev_entries: List[Document.Node.Entry], - seen: Map[Document.Node.Name, A]) - { + seen: Map[Document.Node.Name, A] + ) { private def cons(entry: Document.Node.Entry): Dependencies[A] = new Dependencies[A](entry :: rev_entries, seen) def require_thy(adjunct: A, thy: (Document.Node.Name, Position.T), initiators: List[Document.Node.Name] = Nil, - progress: Progress = new Progress): Dependencies[A] = - { + progress: Progress = new Progress): Dependencies[A] = { val (name, pos) = thy def message: String = @@ -380,8 +374,7 @@ case errs => error(cat_lines(errs)) } - lazy val theory_graph: Document.Node.Name.Graph[Unit] = - { + lazy val theory_graph: Document.Node.Name.Graph[Unit] = { val regular = theories.toSet val irregular = (for { @@ -420,9 +413,10 @@ theories.map(name => resources.load_commands(get_syntax(name), name)))) .filter(p => p._2.nonEmpty) - def loaded_files(name: Document.Node.Name, spans: List[Command_Span.Span]) - : (String, List[Path]) = - { + def loaded_files( + name: Document.Node.Name, + spans: List[Command_Span.Span] + ) : (String, List[Path]) = { val theory = name.theory val syntax = get_syntax(name) val files1 = resources.loaded_files(syntax, name, spans) @@ -436,8 +430,7 @@ file <- loaded_files(name, spans)._2 } yield file - def imported_files: List[Path] = - { + def imported_files: List[Path] = { val base_theories = loaded_theories.all_preds(theories.map(_.theory)). filter(session_base.loaded_theories.defined) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/PIDE/session.scala --- a/src/Pure/PIDE/session.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/PIDE/session.scala Fri Apr 01 23:26:19 2022 +0200 @@ -13,26 +13,22 @@ import scala.annotation.tailrec -object Session -{ +object Session { /* outlets */ - object Consumer - { + object Consumer { def apply[A](name: String)(consume: A => Unit): Consumer[A] = new Consumer[A](name, consume) } final class Consumer[-A] private(val name: String, val consume: A => Unit) - class Outlet[A](dispatcher: Consumer_Thread[() => Unit]) - { + class Outlet[A](dispatcher: Consumer_Thread[() => Unit]) { private val consumers = Synchronized[List[Consumer[A]]](Nil) def += (c: Consumer[A]): Unit = consumers.change(Library.update(c)) def -= (c: Consumer[A]): Unit = consumers.change(Library.remove(c)) - def post(a: A): Unit = - { + def post(a: A): Unit = { for (c <- consumers.value.iterator) { dispatcher.send(() => try { c.consume(a) } @@ -73,8 +69,7 @@ case class Commands_Changed( assignment: Boolean, nodes: Set[Document.Node.Name], commands: Set[Command]) - sealed abstract class Phase - { + sealed abstract class Phase { def print: String = this match { case Terminated(result) => if (result.ok) "finished" else "failed" @@ -91,8 +86,7 @@ /* syslog */ - private[Session] class Syslog(limit: Int) - { + private[Session] class Syslog(limit: Int) { private var queue = Queue.empty[XML.Elem] private var length = 0 @@ -113,8 +107,7 @@ type Protocol_Function = Prover.Protocol_Output => Boolean - abstract class Protocol_Handler extends Isabelle_System.Service - { + abstract class Protocol_Handler extends Isabelle_System.Service { def init(session: Session): Unit = {} def exit(): Unit = {} def functions: List[(String, Protocol_Function)] = Nil @@ -123,8 +116,7 @@ } -class Session(_session_options: => Options, val resources: Resources) extends Document.Session -{ +class Session(_session_options: => Options, val resources: Resources) extends Document.Session { session => val cache: Term.Cache = Term.Cache.make() @@ -156,26 +148,22 @@ private val dispatcher = Consumer_Thread.fork[() => Unit]("Session.dispatcher", daemon = true) { case e => e(); true } - def assert_dispatcher[A](body: => A): A = - { + def assert_dispatcher[A](body: => A): A = { assert(dispatcher.check_thread()) body } - def require_dispatcher[A](body: => A): A = - { + def require_dispatcher[A](body: => A): A = { require(dispatcher.check_thread(), "not on dispatcher thread") body } - def send_dispatcher(body: => Unit): Unit = - { + def send_dispatcher(body: => Unit): Unit = { if (dispatcher.check_thread()) body else dispatcher.send(() => body) } - def send_wait_dispatcher(body: => Unit): Unit = - { + def send_wait_dispatcher(body: => Unit): Unit = { if (dispatcher.check_thread()) body else dispatcher.send_wait(() => body) } @@ -218,8 +206,7 @@ /* phase */ - private def post_phase(new_phase: Session.Phase): Session.Phase = - { + private def post_phase(new_phase: Session.Phase): Session.Phase = { phase_changed.post(new_phase) new_phase } @@ -245,8 +232,7 @@ consolidate: List[Document.Node.Name], version_result: Promise[Document.Version]) - private val change_parser = Consumer_Thread.fork[Text_Edits]("change_parser", daemon = true) - { + private val change_parser = Consumer_Thread.fork[Text_Edits]("change_parser", daemon = true) { case Text_Edits(previous, doc_blobs, text_edits, consolidate, version_result) => val prev = previous.get_finished val change = @@ -261,8 +247,7 @@ /* buffered changes */ - private object change_buffer - { + private object change_buffer { private var assignment: Boolean = false private var nodes: Set[Document.Node.Name] = Set.empty private var commands: Set[Command] = Set.empty @@ -291,8 +276,7 @@ delay_flush.invoke() } - def shutdown(): Unit = - { + def shutdown(): Unit = { delay_flush.revoke() flush() } @@ -301,8 +285,7 @@ /* postponed changes */ - private object postponed_changes - { + private object postponed_changes { private var postponed: List[Session.Change] = Nil def store(change: Session.Change): Unit = synchronized { postponed ::= change } @@ -317,22 +300,19 @@ /* node consolidation */ - private object consolidation - { + private object consolidation { private val delay = Delay.first(consolidate_delay) { manager.send(Consolidate_Execution) } private val init_state: Option[Set[Document.Node.Name]] = Some(Set.empty) private val state = Synchronized(init_state) - def exit(): Unit = - { + def exit(): Unit = { delay.revoke() state.change(_ => None) } - def update(new_nodes: Set[Document.Node.Name] = Set.empty): Unit = - { + def update(new_nodes: Set[Document.Node.Name] = Set.empty): Unit = { val active = state.change_result(st => (st.isDefined, st.map(nodes => if (nodes.isEmpty) new_nodes else nodes ++ new_nodes))) @@ -346,8 +326,7 @@ /* prover process */ - private object prover - { + private object prover { private val variable = Synchronized[Option[Prover]](None) def defined: Boolean = variable.value.isDefined @@ -391,8 +370,7 @@ private val delay_prune = Delay.first(prune_delay) { manager.send(Prune_History) } - private val manager: Consumer_Thread[Any] = - { + private val manager: Consumer_Thread[Any] = { /* global state */ val global_state = Synchronized(Document.State.init) @@ -402,9 +380,8 @@ def handle_raw_edits( doc_blobs: Document.Blobs = Document.Blobs.empty, edits: List[Document.Edit_Text] = Nil, - consolidate: List[Document.Node.Name] = Nil): Unit = + consolidate: List[Document.Node.Name] = Nil): Unit = { //{{{ - { require(prover.defined, "prover process not defined (handle_raw_edits)") if (edits.nonEmpty) prover.get.discontinue_execution() @@ -415,22 +392,20 @@ raw_edits.post(Session.Raw_Edits(doc_blobs, edits)) change_parser.send(Text_Edits(previous, doc_blobs, edits, consolidate, version)) + //}}} } - //}}} /* resulting changes */ - def handle_change(change: Session.Change): Unit = + def handle_change(change: Session.Change): Unit = { //{{{ - { require(prover.defined, "prover process not defined (handle_change)") // define commands { val id_commands = new mutable.ListBuffer[Command] - def id_command(command: Command): Unit = - { + def id_command(command: Command): Unit = { for { (name, digest) <- command.blobs_defined if !global_state.value.defined_blob(digest) @@ -462,23 +437,20 @@ prover.get.update(change.previous.id, change.version.id, change.doc_edits, change.consolidate) resources.commit(change) + //}}} } - //}}} /* prover output */ - def handle_output(output: Prover.Output): Unit = + def handle_output(output: Prover.Output): Unit = { //{{{ - { - def bad_output(): Unit = - { + def bad_output(): Unit = { if (verbose) Output.warning("Ignoring bad prover output: " + output.message.toString) } - def change_command(f: Document.State => (Command.State, Document.State)): Unit = - { + def change_command(f: Document.State => (Command.State, Document.State)): Unit = { try { val st = global_state.change_result(f) if (!st.command.span.is_theory) { @@ -591,14 +563,13 @@ raw_output_messages.post(output) } } + //}}} } - //}}} /* main thread */ - Consumer_Thread.fork[Any]("Session.manager", daemon = true) - { + Consumer_Thread.fork[Any]("Session.manager", daemon = true) { case arg: Any => //{{{ arg match { @@ -697,8 +668,7 @@ /* main operations */ - def get_state(): Document.State = - { + def get_state(): Document.State = { if (manager.is_active()) { val promise = Future.promise[Document.State] manager.send_wait(Get_State(promise)) @@ -715,8 +685,7 @@ get_state().recent_finished.version.get_finished.nodes(name).syntax getOrElse resources.session_base.overall_syntax - @tailrec final def await_stable_snapshot(): Document.Snapshot = - { + @tailrec final def await_stable_snapshot(): Document.Snapshot = { val snapshot = this.snapshot() if (snapshot.is_outdated) { output_delay.sleep() @@ -725,8 +694,7 @@ else snapshot } - def start(start_prover: Prover.Receiver => Prover): Unit = - { + def start(start_prover: Prover.Receiver => Prover): Unit = { file_formats _phase.change( { @@ -737,8 +705,7 @@ }) } - def stop(): Process_Result = - { + def stop(): Process_Result = { val was_ready = _phase.guarded_access( { diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/PIDE/text.scala --- a/src/Pure/PIDE/text.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/PIDE/text.scala Fri Apr 01 23:26:19 2022 +0200 @@ -12,8 +12,7 @@ import scala.util.Sorting -object Text -{ +object Text { /* offset */ type Offset = Int @@ -21,21 +20,18 @@ /* range -- with total quasi-ordering */ - object Range - { + object Range { def apply(start: Offset): Range = Range(start, start) val full: Range = apply(0, Integer.MAX_VALUE / 2) val offside: Range = apply(-1) - object Ordering extends scala.math.Ordering[Range] - { + object Ordering extends scala.math.Ordering[Range] { def compare(r1: Range, r2: Range): Int = r1 compare r2 } } - sealed case class Range(start: Offset, stop: Offset) - { + sealed case class Range(start: Offset, stop: Offset) { // denotation: {start} Un {i. start < i & i < stop} if (start > stop) error("Bad range: [" + start.toString + ".." + stop.toString + "]") @@ -82,20 +78,17 @@ /* perspective */ - object Perspective - { + object Perspective { val empty: Perspective = Perspective(Nil) def full: Perspective = Perspective(List(Range.full)) - def apply(ranges: List[Range]): Perspective = - { + def apply(ranges: List[Range]): Perspective = { val result = new mutable.ListBuffer[Range] var last: Option[Range] = None def ship(next: Option[Range]): Unit = { result ++= last; last = next } - for (range <- ranges.sortBy(_.start)) - { + for (range <- ranges.sortBy(_.start)) { last match { case None => ship(Some(range)) case Some(last_range) => @@ -111,8 +104,8 @@ } final class Perspective private( - val ranges: List[Range]) // visible text partitioning in canonical order - { + val ranges: List[Range] // visible text partitioning in canonical order + ) { def is_empty: Boolean = ranges.isEmpty def range: Range = if (is_empty) Range(0) @@ -130,8 +123,7 @@ /* information associated with text range */ - sealed case class Info[A](range: Range, info: A) - { + sealed case class Info[A](range: Range, info: A) { def restrict(r: Range): Info[A] = Info(range.restrict(r), info) def try_restrict(r: Range): Option[Info[A]] = range.try_restrict(r).map(Info(_, info)) @@ -143,8 +135,7 @@ /* editing */ - object Edit - { + object Edit { def insert(start: Offset, text: String): Edit = new Edit(true, start, text) def remove(start: Offset, text: String): Edit = new Edit(false, start, text) def inserts(start: Offset, text: String): List[Edit] = @@ -156,8 +147,7 @@ else removes(start, old_text) ::: inserts(start, new_text) } - final class Edit private(val is_insert: Boolean, val start: Offset, val text: String) - { + final class Edit private(val is_insert: Boolean, val start: Offset, val text: String) { override def toString: String = (if (is_insert) "Insert(" else "Remove(") + (start, text).toString + ")" diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/PIDE/xml.scala --- a/src/Pure/PIDE/xml.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/PIDE/xml.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object XML -{ +object XML { /** XML trees **/ /* datatype representation */ @@ -18,8 +17,7 @@ sealed abstract class Tree { override def toString: String = string_of_tree(this) } type Body = List[Tree] - case class Elem(markup: Markup, body: Body) extends Tree - { + case class Elem(markup: Markup, body: Body) extends Tree { private lazy val hash: Int = (markup, body).hashCode() override def hashCode(): Int = hash @@ -31,8 +29,7 @@ def + (att: Attribute): Elem = Elem(markup + att, body) } - case class Text(content: String) extends Tree - { + case class Text(content: String) extends Tree { private lazy val hash: Int = content.hashCode() override def hashCode(): Int = hash } @@ -52,14 +49,12 @@ /* name space */ - object Namespace - { + object Namespace { def apply(prefix: String, target: String): Namespace = new Namespace(prefix, target) } - final class Namespace private(prefix: String, target: String) - { + final class Namespace private(prefix: String, target: String) { def apply(name: String): String = prefix + ":" + name val attribute: XML.Attribute = ("xmlns:" + prefix, target) @@ -73,8 +68,7 @@ val XML_NAME = "xml_name" val XML_BODY = "xml_body" - object Wrapped_Elem - { + object Wrapped_Elem { def apply(markup: Markup, body1: Body, body2: Body): XML.Elem = XML.Elem(Markup(XML_ELEM, (XML_NAME, markup.name) :: markup.properties), XML.Elem(Markup(XML_BODY, Nil), body1) :: body2) @@ -89,8 +83,7 @@ } } - object Root_Elem - { + object Root_Elem { def apply(body: Body): XML.Elem = XML.Elem(Markup(XML_ELEM, Nil), body) def unapply(tree: Tree): Option[Body] = tree match { @@ -102,8 +95,7 @@ /* traverse text */ - def traverse_text[A](body: Body)(a: A)(op: (A, String) => A): A = - { + def traverse_text[A](body: Body)(a: A)(op: (A, String) => A): A = { def traverse(x: A, t: Tree): A = t match { case XML.Wrapped_Elem(_, _, ts) => ts.foldLeft(x)(traverse) @@ -119,8 +111,7 @@ /* text content */ - def content(body: Body): String = - { + def content(body: Body): String = { val text = new StringBuilder(text_length(body)) traverse_text(body)(()) { case (_, s) => text.append(s) } text.toString @@ -134,8 +125,7 @@ val header: String = "\n" - def output_char(s: StringBuilder, c: Char, permissive: Boolean = false): Unit = - { + def output_char(s: StringBuilder, c: Char, permissive: Boolean = false): Unit = { c match { case '<' => s ++= "<" case '>' => s ++= ">" @@ -146,14 +136,12 @@ } } - def output_string(s: StringBuilder, str: String, permissive: Boolean = false): Unit = - { + def output_string(s: StringBuilder, str: String, permissive: Boolean = false): Unit = { if (str == null) s ++= str else str.iterator.foreach(output_char(s, _, permissive = permissive)) } - def output_elem(s: StringBuilder, markup: Markup, end: Boolean = false): Unit = - { + def output_elem(s: StringBuilder, markup: Markup, end: Boolean = false): Unit = { s += '<' s ++= markup.name for ((a, b) <- markup.properties) { @@ -168,16 +156,14 @@ s += '>' } - def output_elem_end(s: StringBuilder, name: String): Unit = - { + def output_elem_end(s: StringBuilder, name: String): Unit = { s += '<' s += '/' s ++= name s += '>' } - def string_of_body(body: Body): String = - { + def string_of_body(body: Body): String = { val s = new StringBuilder def tree(t: Tree): Unit = @@ -201,8 +187,7 @@ /** cache **/ - object Cache - { + object Cache { def make( xz: XZ.Cache = XZ.Cache.make(), max_string: Int = isabelle.Cache.default_max_string, @@ -213,10 +198,8 @@ } class Cache(val xz: XZ.Cache, max_string: Int, initial_size: Int) - extends isabelle.Cache(max_string, initial_size) - { - protected def cache_props(x: Properties.T): Properties.T = - { + extends isabelle.Cache(max_string, initial_size) { + protected def cache_props(x: Properties.T): Properties.T = { if (x.isEmpty) x else lookup(x) match { @@ -225,8 +208,7 @@ } } - protected def cache_markup(x: Markup): Markup = - { + protected def cache_markup(x: Markup): Markup = { lookup(x) match { case Some(y) => y case None => @@ -237,8 +219,7 @@ } } - protected def cache_tree(x: XML.Tree): XML.Tree = - { + protected def cache_tree(x: XML.Tree): XML.Tree = { lookup(x) match { case Some(y) => y case None => @@ -250,8 +231,7 @@ } } - protected def cache_body(x: XML.Body): XML.Body = - { + protected def cache_body(x: XML.Body): XML.Body = { if (x.isEmpty) x else lookup(x) match { @@ -285,8 +265,7 @@ class XML_Atom(s: String) extends Error(s) class XML_Body(body: XML.Body) extends Error("") - object Encode - { + object Encode { type T[A] = A => XML.Body type V[A] = PartialFunction[A, (List[String], XML.Body)] type P[A] = PartialFunction[A, List[String]] @@ -340,22 +319,19 @@ def list[A](f: T[A]): T[List[A]] = (xs => xs.map((x: A) => node(f(x)))) - def option[A](f: T[A]): T[Option[A]] = - { + def option[A](f: T[A]): T[Option[A]] = { case None => Nil case Some(x) => List(node(f(x))) } - def variant[A](fs: List[V[A]]): T[A] = - { + def variant[A](fs: List[V[A]]): T[A] = { case x => val (f, tag) = fs.iterator.zipWithIndex.find(p => p._1.isDefinedAt(x)).get List(tagged(tag, f(x))) } } - object Decode - { + object Decode { type T[A] = XML.Body => A type V[A] = (List[String], XML.Body) => A type P[A] = PartialFunction[List[String], A] @@ -401,20 +377,17 @@ /* representation of standard types */ - val tree: T[XML.Tree] = - { + val tree: T[XML.Tree] = { case List(t) => t case ts => throw new XML_Body(ts) } - val properties: T[Properties.T] = - { + val properties: T[Properties.T] = { case List(XML.Elem(Markup(":", props), Nil)) => props case ts => throw new XML_Body(ts) } - val string: T[String] = - { + val string: T[String] = { case Nil => "" case List(XML.Text(s)) => s case ts => throw new XML_Body(ts) @@ -428,14 +401,12 @@ val unit: T[Unit] = (x => unit_atom(string(x))) - def pair[A, B](f: T[A], g: T[B]): T[(A, B)] = - { + def pair[A, B](f: T[A], g: T[B]): T[(A, B)] = { case List(t1, t2) => (f(node(t1)), g(node(t2))) case ts => throw new XML_Body(ts) } - def triple[A, B, C](f: T[A], g: T[B], h: T[C]): T[(A, B, C)] = - { + def triple[A, B, C](f: T[A], g: T[B], h: T[C]): T[(A, B, C)] = { case List(t1, t2, t3) => (f(node(t1)), g(node(t2)), h(node(t3))) case ts => throw new XML_Body(ts) } @@ -443,15 +414,13 @@ def list[A](f: T[A]): T[List[A]] = (ts => ts.map(t => f(node(t)))) - def option[A](f: T[A]): T[Option[A]] = - { + def option[A](f: T[A]): T[Option[A]] = { case Nil => None case List(t) => Some(f(node(t))) case ts => throw new XML_Body(ts) } - def variant[A](fs: List[V[A]]): T[A] = - { + def variant[A](fs: List[V[A]]): T[A] = { case List(t) => val (tag, (xs, ts)) = tagged(t) val f = diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/PIDE/yxml.scala --- a/src/Pure/PIDE/yxml.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/PIDE/yxml.scala Fri Apr 01 23:26:19 2022 +0200 @@ -11,8 +11,7 @@ import scala.collection.mutable -object YXML -{ +object YXML { /* chunk markers */ val X = '\u0005' @@ -32,8 +31,7 @@ /* string representation */ - def traversal(string: String => Unit, body: XML.Body): Unit = - { + def traversal(string: String => Unit, body: XML.Body): Unit = { def tree(t: XML.Tree): Unit = t match { case XML.Elem(markup @ Markup(name, atts), ts) => @@ -51,8 +49,7 @@ body.foreach(tree) } - def string_of_body(body: XML.Body): String = - { + def string_of_body(body: XML.Body): String = { val s = new StringBuilder traversal(str => s ++= str, body) s.toString @@ -74,8 +71,7 @@ Properties.Eq.unapply(source.toString) getOrElse err_attribute() - def parse_body(source: CharSequence, cache: XML.Cache = XML.Cache.none): XML.Body = - { + def parse_body(source: CharSequence, cache: XML.Cache = XML.Cache.none): XML.Body = { /* stack operations */ def buffer(): mutable.ListBuffer[XML.Tree] = new mutable.ListBuffer[XML.Tree] @@ -134,14 +130,12 @@ private def markup_broken(source: CharSequence) = XML.Elem(Markup.Broken, List(XML.Text(source.toString))) - def parse_body_failsafe(source: CharSequence, cache: XML.Cache = XML.Cache.none): XML.Body = - { + def parse_body_failsafe(source: CharSequence, cache: XML.Cache = XML.Cache.none): XML.Body = { try { parse_body(source, cache = cache) } catch { case ERROR(_) => List(markup_broken(source)) } } - def parse_failsafe(source: CharSequence, cache: XML.Cache = XML.Cache.none): XML.Tree = - { + def parse_failsafe(source: CharSequence, cache: XML.Cache = XML.Cache.none): XML.Tree = { try { parse(source, cache = cache) } catch { case ERROR(_) => markup_broken(source) } } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/ROOT.scala --- a/src/Pure/ROOT.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/ROOT.scala Fri Apr 01 23:26:19 2022 +0200 @@ -4,8 +4,7 @@ Root of isabelle package. */ -package object isabelle -{ +package object isabelle { val ERROR = Exn.ERROR val error = Exn.error _ def cat_error(msgs: String*): Nothing = Exn.cat_error(msgs:_*) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/System/bash.scala --- a/src/Pure/System/bash.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/System/bash.scala Fri Apr 01 23:26:19 2022 +0200 @@ -15,12 +15,10 @@ import scala.jdk.OptionConverters._ -object Bash -{ +object Bash { /* concrete syntax */ - private def bash_chr(c: Byte): String = - { + private def bash_chr(c: Byte): String = { val ch = c.toChar ch match { case '\t' => "$'\\t'" @@ -61,13 +59,13 @@ new Process(script, description, cwd, env, redirect, cleanup) class Process private[Bash]( - script: String, - description: String, - cwd: JFile, - env: JMap[String, String], - redirect: Boolean, - cleanup: () => Unit) - { + script: String, + description: String, + cwd: JFile, + env: JMap[String, String], + redirect: Boolean, + cleanup: () => Unit + ) { override def toString: String = make_description(description) private val timing_file = Isabelle_System.tmp_file("bash_timing") @@ -123,10 +121,8 @@ file.exists() && process_alive(Library.trim_line(File.read(file))) } - @tailrec private def signal(s: String, count: Int = 1): Boolean = - { - count <= 0 || - { + @tailrec private def signal(s: String, count: Int = 1): Boolean = { + count <= 0 || { isabelle.setup.Environment.kill_process(group_pid, s) val running = root_process_alive() || @@ -139,15 +135,13 @@ } } - def terminate(): Unit = Isabelle_Thread.try_uninterruptible - { + def terminate(): Unit = Isabelle_Thread.try_uninterruptible { signal("INT", count = 7) && signal("TERM", count = 3) && signal("KILL") proc.destroy() do_cleanup() } - def interrupt(): Unit = Isabelle_Thread.try_uninterruptible - { + def interrupt(): Unit = Isabelle_Thread.try_uninterruptible { isabelle.setup.Environment.kill_process(group_pid, "INT") } @@ -162,8 +156,7 @@ // cleanup - private def do_cleanup(): Unit = - { + private def do_cleanup(): Unit = { try { Runtime.getRuntime.removeShutdownHook(shutdown_hook) } catch { case _: IllegalStateException => } @@ -192,8 +185,7 @@ // join - def join(): Int = - { + def join(): Int = { val rc = proc.waitFor() do_cleanup() rc @@ -207,8 +199,8 @@ progress_stdout: String => Unit = (_: String) => (), progress_stderr: String => Unit = (_: String) => (), watchdog: Option[Watchdog] = None, - strict: Boolean = true): Process_Result = - { + strict: Boolean = true + ): Process_Result = { val in = if (input.isEmpty) Future.value(stdin.close()) else Future.thread("bash_stdin") { stdin.write(input); stdin.flush(); stdin.close(); } @@ -247,8 +239,7 @@ /* server */ - object Server - { + object Server { // input messages private val RUN = "run" private val KILL = "kill" @@ -259,8 +250,7 @@ private val FAILURE = "failure" private val RESULT = "result" - def start(port: Int = 0, debugging: => Boolean = false): Server = - { + def start(port: Int = 0, debugging: => Boolean = false): Server = { val server = new Server(port, debugging) server.start() server @@ -268,20 +258,17 @@ } class Server private(port: Int, debugging: => Boolean) - extends isabelle.Server.Handler(port) - { + extends isabelle.Server.Handler(port) { server => private val _processes = Synchronized(Map.empty[UUID.T, Bash.Process]) - override def stop(): Unit = - { + override def stop(): Unit = { for ((_, process) <- _processes.value) process.terminate() super.stop() } - override def handle(connection: isabelle.Server.Connection): Unit = - { + override def handle(connection: isabelle.Server.Connection): Unit = { def reply(chunks: List[String]): Unit = try { connection.write_byte_message(chunks.map(Bytes.apply)) } catch { case _: IOException => } @@ -366,26 +353,22 @@ } } - class Handler extends Session.Protocol_Handler - { + class Handler extends Session.Protocol_Handler { private var server: Server = null - override def init(session: Session): Unit = - { + override def init(session: Session): Unit = { exit() server = Server.start(debugging = session.session_options.bool("bash_process_debugging")) } - override def exit(): Unit = - { + override def exit(): Unit = { if (server != null) { server.stop() server = null } } - override def prover_options(options: Options): Options = - { + override def prover_options(options: Options): Options = { val address = if (server == null) "" else server.address val password = if (server == null) "" else server.password options + diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/System/command_line.scala --- a/src/Pure/System/command_line.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/System/command_line.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,10 +7,8 @@ package isabelle -object Command_Line -{ - object Chunks - { +object Command_Line { + object Chunks { private def chunks(list: List[String]): List[List[String]] = list.indexWhere(_ == "\n") match { case -1 => List(list) @@ -21,8 +19,7 @@ def unapplySeq(list: List[String]): Option[List[List[String]]] = Some(chunks(list)) } - def tool(body: => Unit): Unit = - { + def tool(body: => Unit): Unit = { val thread = Isabelle_Thread.fork(name = "command_line", inherit_locals = true) { val rc = diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/System/components.scala --- a/src/Pure/System/components.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/System/components.scala Fri Apr 01 23:26:19 2022 +0200 @@ -10,20 +10,17 @@ import java.io.{File => JFile} -object Components -{ +object Components { /* archive name */ - object Archive - { + object Archive { val suffix: String = ".tar.gz" def apply(name: String): String = if (name == "") error("Bad component name: " + quote(name)) else name + suffix - def unapply(archive: String): Option[String] = - { + def unapply(archive: String): Option[String] = { for { name0 <- Library.try_unsuffix(suffix, archive) name <- proper_string(name0) @@ -48,8 +45,7 @@ def contrib(dir: Path = Path.current, name: String = ""): Path = dir + Path.explode("contrib") + Path.explode(name) - def unpack(dir: Path, archive: Path, progress: Progress = new Progress): String = - { + def unpack(dir: Path, archive: Path, progress: Progress = new Progress): String = { val name = Archive.get_name(archive.file_name) progress.echo("Unpacking " + name) Isabelle_System.gnutar("-xzf " + File.bash_path(archive), dir = dir).check @@ -59,8 +55,8 @@ def resolve(base_dir: Path, names: List[String], target_dir: Option[Path] = None, copy_dir: Option[Path] = None, - progress: Progress = new Progress): Unit = - { + progress: Progress = new Progress + ): Unit = { Isabelle_System.make_directory(base_dir) for (name <- names) { val archive_name = Archive(name) @@ -89,8 +85,7 @@ private val platforms_all: Set[String] = Set("x86-linux", "x86-cygwin") ++ platforms_family.iterator.flatMap(_._2) - def purge(dir: Path, platform: Platform.Family.Value): Unit = - { + def purge(dir: Path, platform: Platform.Family.Value): Unit = { val purge_set = platforms_all -- platforms_family(platform) File.find_files(dir.file, @@ -124,8 +119,7 @@ val components_sha1: Path = Path.explode("~~/Admin/components/components.sha1") - sealed case class SHA1_Digest(digest: SHA1.Digest, name: String) - { + sealed case class SHA1_Digest(digest: SHA1.Digest, name: String) { override def toString: String = digest.shasum(name) } @@ -149,14 +143,12 @@ if (components_path.is_file) Library.trim_split_lines(File.read(components_path)) else Nil - def write_components(lines: List[String]): Unit = - { + def write_components(lines: List[String]): Unit = { Isabelle_System.make_directory(components_path.dir) File.write(components_path, Library.terminate_lines(lines)) } - def update_components(add: Boolean, path0: Path, progress: Progress = new Progress): Unit = - { + def update_components(add: Boolean, path0: Path, progress: Progress = new Progress): Unit = { val path = path0.expand.absolute if (!check_dir(path) && !Sessions.is_session_dir(path)) error("Bad component directory: " + path) @@ -175,8 +167,7 @@ /* main entry point */ - def main(args: Array[String]): Unit = - { + def main(args: Array[String]): Unit = { Command_Line.tool { for (arg <- args) { val add = @@ -198,8 +189,8 @@ progress: Progress = new Progress, publish: Boolean = false, force: Boolean = false, - update_components_sha1: Boolean = false): Unit = - { + update_components_sha1: Boolean = false + ): Unit = { val archives: List[Path] = for (path <- components) yield { path.file_name match { @@ -232,8 +223,7 @@ if ((publish && archives.nonEmpty) || update_components_sha1) { options.string("isabelle_components_server") match { case SSH.Target(user, host) => - using(SSH.open_session(options, host = host, user = user))(ssh => - { + using(SSH.open_session(options, host = host, user = user)) { ssh => val components_dir = Path.explode(options.string("isabelle_components_dir")) val contrib_dir = Path.explode(options.string("isabelle_components_contrib_dir")) @@ -257,11 +247,10 @@ // contrib directory val is_standard_component = - Isabelle_System.with_tmp_dir("component")(tmp_dir => - { + Isabelle_System.with_tmp_dir("component") { tmp_dir => Isabelle_System.gnutar("-xzf " + File.bash_path(archive), dir = tmp_dir).check check_dir(tmp_dir + Path.explode(name)) - }) + } if (is_standard_component) { if (ssh.is_dir(remote_contrib)) { if (force) ssh.rm_tree(remote_contrib) @@ -291,7 +280,7 @@ } write_components_sha1(read_components_sha1(lines)) } - }) + } case s => error("Bad isabelle_components_server: " + quote(s)) } } @@ -321,17 +310,17 @@ val isabelle_tool = Isabelle_Tool("build_components", "build and publish Isabelle components", - Scala_Project.here, args => - { - var publish = false - var update_components_sha1 = false - var force = false - var options = Options.init() + Scala_Project.here, + { args => + var publish = false + var update_components_sha1 = false + var force = false + var options = Options.init() - def show_options: String = - cat_lines(relevant_options.map(name => options.options(name).print)) + def show_options: String = + cat_lines(relevant_options.map(name => options.options(name).print)) - val getopts = Getopts(""" + val getopts = Getopts(""" Usage: isabelle build_components [OPTIONS] ARCHIVES... DIRS... Options are: @@ -344,17 +333,17 @@ depending on system options: """ + Library.indent_lines(2, show_options) + "\n", - "P" -> (_ => publish = true), - "f" -> (_ => force = true), - "o:" -> (arg => options = options + arg), - "u" -> (_ => update_components_sha1 = true)) + "P" -> (_ => publish = true), + "f" -> (_ => force = true), + "o:" -> (arg => options = options + arg), + "u" -> (_ => update_components_sha1 = true)) - val more_args = getopts(args) - if (more_args.isEmpty && !update_components_sha1) getopts.usage() + val more_args = getopts(args) + if (more_args.isEmpty && !update_components_sha1) getopts.usage() - val progress = new Console_Progress + val progress = new Console_Progress - build_components(options, more_args.map(Path.explode), progress = progress, - publish = publish, force = force, update_components_sha1 = update_components_sha1) - }) + build_components(options, more_args.map(Path.explode), progress = progress, + publish = publish, force = force, update_components_sha1 = update_components_sha1) + }) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/System/executable.scala --- a/src/Pure/System/executable.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/System/executable.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,19 +7,17 @@ package isabelle -object Executable -{ +object Executable { def libraries_closure(path: Path, mingw: MinGW = MinGW.none, filter: String => Boolean = _ => true, - patchelf: Boolean = false): List[String] = - { + patchelf: Boolean = false + ): List[String] = { val exe_path = path.expand val exe_dir = exe_path.dir val exe = exe_path.base - val ldd_lines = - { + val ldd_lines = { val ldd = if (Platform.is_macos) "otool -L" else "ldd" val script = mingw.bash_script(ldd + " " + File.bash_path(exe)) Library.split_lines(Isabelle_System.bash(script, cwd = exe_dir.file).check.out) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/System/getopts.scala --- a/src/Pure/System/getopts.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/System/getopts.scala Fri Apr 01 23:26:19 2022 +0200 @@ -10,10 +10,8 @@ import scala.annotation.tailrec -object Getopts -{ - def apply(usage_text: String, option_specs: (String, String => Unit)*): Getopts = - { +object Getopts { + def apply(usage_text: String, option_specs: (String, String => Unit)*): Getopts = { val options = option_specs.foldLeft(Map.empty[Char, (Boolean, String => Unit)]) { case (m, (s, f)) => @@ -28,10 +26,8 @@ } } -class Getopts private(usage_text: String, options: Map[Char, (Boolean, String => Unit)]) -{ - def usage(): Nothing = - { +class Getopts private(usage_text: String, options: Map[Char, (Boolean, String => Unit)]) { + def usage(): Nothing = { Output.writeln(usage_text, stdout = true) sys.exit(Process_Result.RC.error) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/System/isabelle_charset.scala --- a/src/Pure/System/isabelle_charset.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/System/isabelle_charset.scala Fri Apr 01 23:26:19 2022 +0200 @@ -14,15 +14,13 @@ import java.nio.charset.spi.CharsetProvider -object Isabelle_Charset -{ +object Isabelle_Charset { val name: String = "UTF-8-Isabelle-test" // FIXME lazy val charset: Charset = new Isabelle_Charset } -class Isabelle_Charset extends Charset(Isabelle_Charset.name, null) -{ +class Isabelle_Charset extends Charset(Isabelle_Charset.name, null) { override def contains(cs: Charset): Boolean = cs.name.equalsIgnoreCase(UTF8.charset_name) || UTF8.charset.contains(cs) @@ -32,18 +30,15 @@ } -class Isabelle_Charset_Provider extends CharsetProvider -{ - override def charsetForName(name: String): Charset = - { +class Isabelle_Charset_Provider extends CharsetProvider { + override def charsetForName(name: String): Charset = { // FIXME inactive // if (name.equalsIgnoreCase(Isabelle_Charset.name)) Isabelle_Charset.charset // else null null } - override def charsets(): java.util.Iterator[Charset] = - { + override def charsets(): java.util.Iterator[Charset] = { // FIXME inactive // Iterator(Isabelle_Charset.charset) JList.of[Charset]().listIterator() diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/System/isabelle_fonts.scala --- a/src/Pure/System/isabelle_fonts.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/System/isabelle_fonts.scala Fri Apr 01 23:26:19 2022 +0200 @@ -11,8 +11,7 @@ import java.awt.{Font, GraphicsEnvironment} -object Isabelle_Fonts -{ +object Isabelle_Fonts { /* standard names */ val mono: String = "Isabelle DejaVu Sans Mono" @@ -22,10 +21,8 @@ /* environment entries */ - object Entry - { - object Ordering extends scala.math.Ordering[Entry] - { + object Entry { + object Ordering extends scala.math.Ordering[Entry] { def compare(entry1: Entry, entry2: Entry): Int = entry1.family compare entry2.family match { case 0 => entry1.style compare entry2.style @@ -34,8 +31,7 @@ } } - sealed case class Entry(path: Path, hidden: Boolean = false) - { + sealed case class Entry(path: Path, hidden: Boolean = false) { lazy val bytes: Bytes = Bytes.read(path) lazy val font: Font = Font.createFont(Font.TRUETYPE_FONT, path.file) @@ -56,8 +52,8 @@ def make_entries( getenv: String => String = Isabelle_System.getenv_strict(_), - hidden: Boolean = false): List[Entry] = - { + hidden: Boolean = false + ): List[Entry] = { Path.split(getenv("ISABELLE_FONTS")).map(Entry(_)) ::: (if (hidden) Path.split(getenv("ISABELLE_FONTS_HIDDEN")).map(Entry(_, hidden = true)) else Nil) } @@ -71,8 +67,7 @@ /* system init */ - def init(): Unit = - { + def init(): Unit = { val ge = GraphicsEnvironment.getLocalGraphicsEnvironment() for (entry <- fonts()) ge.registerFont(entry.font) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/System/isabelle_platform.scala --- a/src/Pure/System/isabelle_platform.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/System/isabelle_platform.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Isabelle_Platform -{ +object Isabelle_Platform { val settings: List[String] = List( "ISABELLE_PLATFORM_FAMILY", @@ -17,8 +16,7 @@ "ISABELLE_WINDOWS_PLATFORM64", "ISABELLE_APPLE_PLATFORM64") - def apply(ssh: Option[SSH.Session] = None): Isabelle_Platform = - { + def apply(ssh: Option[SSH.Session] = None): Isabelle_Platform = { ssh match { case None => new Isabelle_Platform(settings.map(a => (a, Isabelle_System.getenv(a)))) @@ -36,8 +34,7 @@ lazy val self: Isabelle_Platform = apply() } -class Isabelle_Platform private(val settings: List[(String, String)]) -{ +class Isabelle_Platform private(val settings: List[(String, String)]) { private def get(name: String): String = settings.collectFirst({ case (a, b) if a == name => b }). getOrElse(error("Bad platform settings variable: " + quote(name))) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/System/isabelle_process.scala --- a/src/Pure/System/isabelle_process.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/System/isabelle_process.scala Fri Apr 01 23:26:19 2022 +0200 @@ -11,8 +11,7 @@ import java.io.{File => JFile} -object Isabelle_Process -{ +object Isabelle_Process { def start( session: Session, options: Options, @@ -24,8 +23,8 @@ eval_main: String = "", modes: List[String] = Nil, cwd: JFile = null, - env: JMap[String, String] = Isabelle_System.settings()): Isabelle_Process = - { + env: JMap[String, String] = Isabelle_System.settings() + ): Isabelle_Process = { val channel = System_Channel() val process = try { @@ -47,8 +46,7 @@ } } -class Isabelle_Process private(session: Session, process: Bash.Process) -{ +class Isabelle_Process private(session: Session, process: Bash.Process) { private val startup = Future.promise[String] private val terminated = Future.promise[Process_Result] @@ -72,8 +70,7 @@ case err => session.stop(); error(err) } - def await_shutdown(): Process_Result = - { + def await_shutdown(): Process_Result = { val result = terminated.join session.stop() result diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/System/isabelle_system.scala --- a/src/Pure/System/isabelle_system.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/System/isabelle_system.scala Fri Apr 01 23:26:19 2022 +0200 @@ -16,12 +16,10 @@ import scala.jdk.CollectionConverters._ -object Isabelle_System -{ +object Isabelle_System { /* settings environment */ - def settings(putenv: List[(String, String)] = Nil): JMap[String, String] = - { + def settings(putenv: List[(String, String)] = Nil): JMap[String, String] = { val env0 = isabelle.setup.Environment.settings() if (putenv.isEmpty) env0 else { @@ -45,8 +43,7 @@ @volatile private var _services: Option[List[Class[Service]]] = None - def services(): List[Class[Service]] = - { + def services(): List[Class[Service]] = { if (_services.isEmpty) init() // unsynchronized check _services.get } @@ -58,10 +55,8 @@ /* init settings + services */ - def make_services(): List[Class[Service]] = - { - def make(where: String, names: List[String]): List[Class[Service]] = - { + def make_services(): List[Class[Service]] = { + def make(where: String, names: List[String]): List[Class[Service]] = { for (name <- names) yield { def err(msg: String): Nothing = error("Bad Isabelle/Scala service " + quote(name) + " in " + where + "\n" + msg) @@ -83,8 +78,7 @@ from_env("ISABELLE_SCALA_SERVICES") ::: Scala.class_path().flatMap(from_jar) } - def init(isabelle_root: String = "", cygwin_root: String = ""): Unit = - { + def init(isabelle_root: String = "", cygwin_root: String = ""): Unit = { isabelle.setup.Environment.init(isabelle_root, cygwin_root) synchronized { if (_services.isEmpty) { _services = Some(make_services()) } } } @@ -92,8 +86,7 @@ /* getetc -- static distribution parameters */ - def getetc(name: String, root: Path = Path.ISABELLE_HOME): Option[String] = - { + def getetc(name: String, root: Path = Path.ISABELLE_HOME): Option[String] = { val path = root + Path.basic("etc") + Path.basic(name) if (path.is_file) { Library.trim_split_lines(File.read(path)) match { @@ -114,8 +107,7 @@ else error("Failed to identify Isabelle distribution " + root.expand) } - object Isabelle_Id extends Scala.Fun_String("isabelle_id") - { + object Isabelle_Id extends Scala.Fun_String("isabelle_id") { val here = Scala_Project.here def apply(arg: String): String = isabelle_id() } @@ -151,8 +143,10 @@ /* scala functions */ - private def apply_paths(args: List[String], fun: List[Path] => Unit): List[String] = - { fun(args.map(Path.explode)); Nil } + private def apply_paths(args: List[String], fun: List[Path] => Unit): List[String] = { + fun(args.map(Path.explode)) + Nil + } private def apply_paths1(args: List[String], fun: Path => Unit): List[String] = apply_paths(args, { case List(path) => fun(path) }) @@ -175,8 +169,7 @@ /* directories */ - def make_directory(path: Path): Path = - { + def make_directory(path: Path): Path = { if (!path.is_dir) { try { Files.createDirectories(path.java_path) } catch { case ERROR(_) => error("Failed to create directory: " + path.absolute) } @@ -188,16 +181,14 @@ if (path.is_dir) error("Directory already exists: " + path.absolute) else make_directory(path) - def copy_dir(dir1: Path, dir2: Path): Unit = - { + def copy_dir(dir1: Path, dir2: Path): Unit = { val res = bash("cp -a " + File.bash_path(dir1) + " " + File.bash_path(dir2)) if (!res.ok) { cat_error("Failed to copy directory " + dir1.absolute + " to " + dir2.absolute, res.err) } } - def with_copy_dir[A](dir1: Path, dir2: Path)(body: => A): A = - { + def with_copy_dir[A](dir1: Path, dir2: Path)(body: => A): A = { if (dir2.is_file || dir2.is_dir) error("Directory already exists: " + dir2.absolute) else { try { copy_dir(dir1, dir2); body } @@ -206,14 +197,12 @@ } - object Make_Directory extends Scala.Fun_Strings("make_directory") - { + object Make_Directory extends Scala.Fun_Strings("make_directory") { val here = Scala_Project.here def apply(args: List[String]): List[String] = apply_paths1(args, make_directory) } - object Copy_Dir extends Scala.Fun_Strings("copy_dir") - { + object Copy_Dir extends Scala.Fun_Strings("copy_dir") { val here = Scala_Project.here def apply(args: List[String]): List[String] = apply_paths2(args, copy_dir) } @@ -221,8 +210,7 @@ /* copy files */ - def copy_file(src: JFile, dst: JFile): Unit = - { + def copy_file(src: JFile, dst: JFile): Unit = { val target = if (dst.isDirectory) new JFile(dst, src.getName) else dst if (!File.eq(src, target)) { try { @@ -240,8 +228,7 @@ def copy_file(src: Path, dst: Path): Unit = copy_file(src.file, dst.file) - def copy_file_base(base_dir: Path, src: Path, target_dir: Path): Unit = - { + def copy_file_base(base_dir: Path, src: Path, target_dir: Path): Unit = { val src1 = src.expand val src1_dir = src1.dir if (!src1.starts_basic) error("Illegal path specification " + src1 + " beyond base directory") @@ -249,14 +236,12 @@ } - object Copy_File extends Scala.Fun_Strings("copy_file") - { + object Copy_File extends Scala.Fun_Strings("copy_file") { val here = Scala_Project.here def apply(args: List[String]): List[String] = apply_paths2(args, copy_file) } - object Copy_File_Base extends Scala.Fun_Strings("copy_file_base") - { + object Copy_File_Base extends Scala.Fun_Strings("copy_file_base") { val here = Scala_Project.here def apply(args: List[String]): List[String] = apply_paths3(args, copy_file_base) } @@ -264,8 +249,7 @@ /* move files */ - def move_file(src: JFile, dst: JFile): Unit = - { + def move_file(src: JFile, dst: JFile): Unit = { val target = if (dst.isDirectory) new JFile(dst, src.getName) else dst if (!File.eq(src, target)) Files.move(src.toPath, target.toPath, StandardCopyOption.REPLACE_EXISTING) @@ -276,16 +260,14 @@ /* symbolic link */ - def symlink(src: Path, dst: Path, force: Boolean = false, native: Boolean = false): Unit = - { + def symlink(src: Path, dst: Path, force: Boolean = false, native: Boolean = false): Unit = { val src_file = src.file val dst_file = dst.file val target = if (dst_file.isDirectory) new JFile(dst_file, src_file.getName) else dst_file if (force) target.delete - def cygwin_link(): Unit = - { + def cygwin_link(): Unit = { if (native) { error("Failed to create native symlink on Windows: " + quote(src_file.toString) + "\n(but it could work as Administrator)") @@ -304,23 +286,20 @@ /* tmp files */ - def isabelle_tmp_prefix(): JFile = - { + def isabelle_tmp_prefix(): JFile = { val path = Path.explode("$ISABELLE_TMP_PREFIX") path.file.mkdirs // low-level mkdirs to avoid recursion via Isabelle environment File.platform_file(path) } - def tmp_file(name: String, ext: String = "", base_dir: JFile = isabelle_tmp_prefix()): JFile = - { + def tmp_file(name: String, ext: String = "", base_dir: JFile = isabelle_tmp_prefix()): JFile = { val suffix = if (ext == "") "" else "." + ext val file = Files.createTempFile(base_dir.toPath, name, suffix).toFile file.deleteOnExit() file } - def with_tmp_file[A](name: String, ext: String = "")(body: Path => A): A = - { + def with_tmp_file[A](name: String, ext: String = "")(body: Path => A): A = { val file = tmp_file(name, ext) try { body(File.path(file)) } finally { file.delete } } @@ -328,21 +307,18 @@ /* tmp dirs */ - def rm_tree(root: JFile): Unit = - { + def rm_tree(root: JFile): Unit = { root.delete if (root.isDirectory) { Files.walkFileTree(root.toPath, new SimpleFileVisitor[JPath] { - override def visitFile(file: JPath, attrs: BasicFileAttributes): FileVisitResult = - { + override def visitFile(file: JPath, attrs: BasicFileAttributes): FileVisitResult = { try { Files.deleteIfExists(file) } catch { case _: IOException => } FileVisitResult.CONTINUE } - override def postVisitDirectory(dir: JPath, e: IOException): FileVisitResult = - { + override def postVisitDirectory(dir: JPath, e: IOException): FileVisitResult = { if (e == null) { try { Files.deleteIfExists(dir) } catch { case _: IOException => } @@ -357,21 +333,18 @@ def rm_tree(root: Path): Unit = rm_tree(root.file) - object Rm_Tree extends Scala.Fun_Strings("rm_tree") - { + object Rm_Tree extends Scala.Fun_Strings("rm_tree") { val here = Scala_Project.here def apply(args: List[String]): List[String] = apply_paths1(args, rm_tree) } - def tmp_dir(name: String, base_dir: JFile = isabelle_tmp_prefix()): JFile = - { + def tmp_dir(name: String, base_dir: JFile = isabelle_tmp_prefix()): JFile = { val dir = Files.createTempDirectory(base_dir.toPath, name).toFile dir.deleteOnExit() dir } - def with_tmp_dir[A](name: String)(body: Path => A): A = - { + def with_tmp_dir[A](name: String)(body: Path => A): A = { val dir = tmp_dir(name) try { body(File.path(dir)) } finally { rm_tree(dir) } } @@ -379,8 +352,7 @@ /* quasi-atomic update of directory */ - def update_directory(dir: Path, f: Path => Unit): Unit = - { + def update_directory(dir: Path, f: Path => Unit): Unit = { val new_dir = dir.ext("new") val old_dir = dir.ext("old") @@ -410,8 +382,8 @@ progress_stderr: String => Unit = (_: String) => (), watchdog: Option[Bash.Watchdog] = None, strict: Boolean = true, - cleanup: () => Unit = () => ()): Process_Result = - { + cleanup: () => Unit = () => () + ): Process_Result = { Bash.process(script, description = description, cwd = cwd, env = env, redirect = redirect, cleanup = cleanup). result(input = input, progress_stdout = progress_stdout, progress_stderr = progress_stderr, @@ -421,8 +393,7 @@ /* command-line tools */ - def require_command(cmd: String, test: String = "--version"): Unit = - { + def require_command(cmd: String, test: String = "--version"): Unit = { if (!bash(Bash.string(cmd) + " " + test).ok) error("Missing system command: " + quote(cmd)) } @@ -435,8 +406,8 @@ dir: Path = Path.current, original_owner: Boolean = false, strip: Int = 0, - redirect: Boolean = false): Process_Result = - { + redirect: Boolean = false + ): Process_Result = { val options = (if (dir.is_current) "" else "-C " + File.bash_path(dir) + " ") + (if (original_owner) "" else "--owner=root --group=staff ") + @@ -446,16 +417,14 @@ else error("Expected to find GNU tar executable") } - def make_patch(base_dir: Path, src: Path, dst: Path, diff_options: String = ""): String = - { - with_tmp_file("patch")(patch => - { + def make_patch(base_dir: Path, src: Path, dst: Path, diff_options: String = ""): String = { + with_tmp_file("patch") { patch => Isabelle_System.bash( "diff -ru " + diff_options + " -- " + File.bash_path(src) + " " + File.bash_path(dst) + " > " + File.bash_path(patch), cwd = base_dir.file).check_rc(_ <= 1) File.read(patch) - }) + } } def hostname(): String = bash("hostname -s").check.out @@ -466,8 +435,7 @@ def pdf_viewer(arg: Path): Unit = bash("exec \"$PDF_VIEWER\" " + File.bash_path(arg) + " >/dev/null 2>/dev/null &") - def open_external_file(name: String): Boolean = - { + def open_external_file(name: String): Boolean = { val ext = Library.take_suffix((c: Char) => c != '.', name.toList)._2.mkString val external = ext.nonEmpty && @@ -490,8 +458,7 @@ /* default logic */ - def default_logic(args: String*): String = - { + def default_logic(args: String*): String = { args.find(_ != "") match { case Some(logic) => logic case None => getenv_strict("ISABELLE_LOGIC") @@ -501,8 +468,7 @@ /* download file */ - def download(url_name: String, progress: Progress = new Progress): HTTP.Content = - { + def download(url_name: String, progress: Progress = new Progress): HTTP.Content = { val url = Url(url_name) progress.echo("Getting " + quote(url_name)) try { HTTP.Client.get(url) } @@ -512,8 +478,7 @@ def download_file(url_name: String, file: Path, progress: Progress = new Progress): Unit = Bytes.write(file, download(url_name, progress = progress).bytes) - object Download extends Scala.Fun("download", thread = true) - { + object Download extends Scala.Fun("download", thread = true) { val here = Scala_Project.here override def invoke(args: List[Bytes]): List[Bytes] = args match { case List(url) => List(download(url.text).bytes) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/System/isabelle_tool.scala --- a/src/Pure/System/isabelle_tool.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/System/isabelle_tool.scala Fri Apr 01 23:26:19 2022 +0200 @@ -12,14 +12,12 @@ import scala.tools.reflect.{ToolBox, ToolBoxError} -object Isabelle_Tool -{ +object Isabelle_Tool { /* Scala source tools */ abstract class Body extends Function[List[String], Unit] - private def compile(path: Path): Body = - { + private def compile(path: Path): Body = { def err(msg: String): Nothing = cat_error(msg, "The error(s) above occurred in Isabelle/Scala tool " + path) @@ -58,8 +56,7 @@ private def dirs(): List[Path] = Path.split(Isabelle_System.getenv_strict("ISABELLE_TOOLS")) - private def is_external(dir: Path, file_name: String): Boolean = - { + private def is_external(dir: Path, file_name: String): Boolean = { val file = (dir + Path.explode(file_name)).file try { file.isFile && file.canRead && @@ -70,17 +67,17 @@ } private def find_external(name: String): Option[List[String] => Unit] = - dirs().collectFirst({ - case dir if is_external(dir, name + ".scala") => - compile(dir + Path.explode(name + ".scala")) - case dir if is_external(dir, name) => - (args: List[String]) => - { + dirs().collectFirst( + { + case dir if is_external(dir, name + ".scala") => + compile(dir + Path.explode(name + ".scala")) + case dir if is_external(dir, name) => + { (args: List[String]) => val tool = dir + Path.explode(name) val result = Isabelle_System.bash(File.bash_path(tool) + " " + Bash.strings(args)) sys.exit(result.print_stdout.rc) } - }) + }) /* internal tools */ @@ -97,8 +94,7 @@ /* list tools */ - abstract class Entry - { + abstract class Entry { def name: String def position: Properties.T def description: String @@ -109,18 +105,15 @@ } } - sealed case class External(name: String, path: Path) extends Entry - { + sealed case class External(name: String, path: Path) extends Entry { def position: Properties.T = Position.File(path.absolute.implode) - def description: String = - { + def description: String = { val Pattern = """.*\bDESCRIPTION: *(.*)""".r split_lines(File.read(path)).collectFirst({ case Pattern(s) => s }) getOrElse "" } } - def external_tools(): List[External] = - { + def external_tools(): List[External] = { for { dir <- dirs() if dir.is_dir file_name <- File.read_dir(dir) if is_external(dir, file_name) @@ -134,8 +127,7 @@ def isabelle_tools(): List[Entry] = (external_tools() ::: internal_tools).sortBy(_.name) - object Isabelle_Tools extends Scala.Fun_String("isabelle_tools") - { + object Isabelle_Tools extends Scala.Fun_String("isabelle_tools") { val here = Scala_Project.here def apply(arg: String): String = if (arg.nonEmpty) error("Bad argument: " + quote(arg)) @@ -149,8 +141,7 @@ /* command line entry point */ - def main(args: Array[String]): Unit = - { + def main(args: Array[String]): Unit = { Command_Line.tool { args.toList match { case Nil | List("-?") => @@ -175,8 +166,8 @@ name: String, description: String, here: Scala_Project.Here, - body: List[String] => Unit) extends Isabelle_Tool.Entry -{ + body: List[String] => Unit +) extends Isabelle_Tool.Entry { def position: Position.T = here.position } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/System/java_statistics.scala --- a/src/Pure/System/java_statistics.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/System/java_statistics.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,19 +7,16 @@ package isabelle -object Java_Statistics -{ +object Java_Statistics { /* memory status */ - sealed case class Memory_Status(heap_size: Long, heap_free: Long) - { + sealed case class Memory_Status(heap_size: Long, heap_free: Long) { def heap_used: Long = (heap_size - heap_free) max 0 def heap_used_fraction: Double = if (heap_size == 0) 0.0 else heap_used.toDouble / heap_size } - def memory_status(): Memory_Status = - { + def memory_status(): Memory_Status = { val heap_size = Runtime.getRuntime.totalMemory() val heap_free = Runtime.getRuntime.freeMemory() Memory_Status(heap_size, heap_free) @@ -28,8 +25,7 @@ /* JVM statistics */ - def jvm_statistics(): Properties.T = - { + def jvm_statistics(): Properties.T = { val status = memory_status() val threads = Thread.activeCount() val workers = Isabelle_Thread.pool.getPoolSize diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/System/linux.scala --- a/src/Pure/System/linux.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/System/linux.scala Fri Apr 01 23:26:19 2022 +0200 @@ -10,15 +10,13 @@ import scala.util.matching.Regex -object Linux -{ +object Linux { /* check system */ def check_system(): Unit = if (!Platform.is_linux) error("Not a Linux system") - def check_system_root(): Unit = - { + def check_system_root(): Unit = { check_system() if (Isabelle_System.bash("id -u").check.out != "0") error("Not running as superuser (root)") } @@ -26,22 +24,19 @@ /* release */ - object Release - { + object Release { private val ID = """^Distributor ID:\s*(\S.*)$""".r private val RELEASE = """^Release:\s*(\S.*)$""".r private val DESCRIPTION = """^Description:\s*(\S.*)$""".r - def apply(): Release = - { + def apply(): Release = { val lines = Isabelle_System.bash("lsb_release -a").check.out_lines def find(R: Regex): String = lines.collectFirst({ case R(a) => a }).getOrElse("Unknown") new Release(find(ID), find(RELEASE), find(DESCRIPTION)) } } - final class Release private(val id: String, val release: String, val description: String) - { + final class Release private(val id: String, val release: String, val description: String) { override def toString: String = description def is_ubuntu: Boolean = id == "Ubuntu" @@ -65,8 +60,7 @@ def package_install(packages: List[String], progress: Progress = new Progress): Unit = progress.bash("apt-get install -y -- " + Bash.strings(packages), echo = true).check - def package_installed(name: String): Boolean = - { + def package_installed(name: String): Boolean = { val result = Isabelle_System.bash("dpkg-query -s " + Bash.string(name)) val pattern = """^Status:.*installed.*$""".r.pattern result.ok && result.out_lines.exists(line => pattern.matcher(line).matches) @@ -78,8 +72,7 @@ def user_exists(name: String): Boolean = Isabelle_System.bash("id " + Bash.string(name)).ok - def user_entry(name: String, field: Int): String = - { + def user_entry(name: String, field: Int): String = { val result = Isabelle_System.bash("getent passwd " + Bash.string(name)).check val fields = space_explode(':', result.out) @@ -94,8 +87,8 @@ def user_add(name: String, description: String = "", system: Boolean = false, - ssh_setup: Boolean = false): Unit = - { + ssh_setup: Boolean = false + ): Unit = { require(!description.contains(','), "malformed description") if (user_exists(name)) error("User already exists: " + quote(name)) @@ -133,8 +126,7 @@ try { service_stop(name) } catch { case ERROR(_) => } - def service_install(name: String, spec: String): Unit = - { + def service_install(name: String, spec: String): Unit = { service_shutdown(name) val service_file = Path.explode("/lib/systemd/system") + Path.basic(name).ext("service") @@ -148,8 +140,7 @@ /* passwords */ - def generate_password(length: Int = 10): String = - { + def generate_password(length: Int = 10): String = { require(length >= 6, "password too short") Isabelle_System.bash("pwgen " + length + " 1").check.out } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/System/mingw.scala --- a/src/Pure/System/mingw.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/System/mingw.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object MinGW -{ +object MinGW { def environment: List[String] = List("PATH=/usr/bin:/bin:/mingw64/bin", "CONFIG_SITE=/mingw64/etc/config.site") @@ -19,8 +18,7 @@ def apply(path: Path) = new MinGW(Some(path)) } -class MinGW private(val root: Option[Path]) -{ +class MinGW private(val root: Option[Path]) { override def toString: String = root match { case None => "MinGW.none" @@ -40,8 +38,7 @@ else if (root.isEmpty) error("Windows platform requires msys/mingw root specification") else root.get - def check: Unit = - { + def check: Unit = { if (Platform.is_windows) { get_root try { require(Isabelle_System.bash(bash_script("uname -s")).check.out.startsWith("MSYS")) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/System/numa.scala --- a/src/Pure/System/numa.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/System/numa.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,12 +7,10 @@ package isabelle -object NUMA -{ +object NUMA { /* available nodes */ - def nodes(): List[Int] = - { + def nodes(): List[Int] = { val numa_nodes_linux = Path.explode("/sys/devices/system/node/online") val Single = """^(\d+)$""".r @@ -52,8 +50,7 @@ try { nodes().length >= 2 && numactl_available } catch { case ERROR(_) => false } - def enabled_warning(progress: Progress, enabled: Boolean): Boolean = - { + def enabled_warning(progress: Progress, enabled: Boolean): Boolean = { def warning = if (nodes().length < 2) Some("no NUMA nodes available") else if (!numactl_available) Some("bad numactl tool") @@ -66,8 +63,7 @@ }) } - class Nodes(enabled: Boolean = true) - { + class Nodes(enabled: Boolean = true) { private val available = nodes().zipWithIndex private var next_index = 0 diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/System/options.scala --- a/src/Pure/System/options.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/System/options.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Options -{ +object Options { type Spec = (String, Option[String]) val empty: Options = new Options() @@ -16,8 +15,7 @@ /* representation */ - sealed abstract class Type - { + sealed abstract class Type { def print: String = Word.lowercase(toString) } case object Bool extends Type @@ -35,8 +33,8 @@ default_value: String, standard_value: Option[String], description: String, - section: String) - { + section: String + ) { private def print_value(x: String): String = if (typ == Options.String) quote(x) else x private def print_standard: String = standard_value match { @@ -44,8 +42,7 @@ case Some(s) if s == default_value => " (standard)" case Some(s) => " (standard " + print_value(s) + ")" } - private def print(default: Boolean): String = - { + private def print(default: Boolean): String = { val x = if (default) default_value else value "option " + name + " : " + typ.print + " = " + print_value(x) + print_standard + (if (description == "") "" else "\n -- " + quote(description)) @@ -54,8 +51,7 @@ def print: String = print(false) def print_default: String = print(true) - def title(strip: String = ""): String = - { + def title(strip: String = ""): String = { val words = Word.explode('_', name) val words1 = words match { @@ -88,8 +84,7 @@ val prefs_syntax: Outer_Syntax = Outer_Syntax.empty + "=" - trait Parser extends Parse.Parser - { + trait Parser extends Parse.Parser { val option_name: Parser[String] = atom("option name", _.is_name) val option_type: Parser[String] = atom("option type", _.is_name) val option_value: Parser[String] = @@ -100,13 +95,11 @@ $$$("(") ~! $$$(STANDARD) ~ opt(option_value) ~ $$$(")") ^^ { case _ ~ _ ~ a ~ _ => a } } - private object Parser extends Parser - { + private object Parser extends Parser { def comment_marker: Parser[String] = $$$("--") | $$$(Symbol.comment) | $$$(Symbol.comment_decoded) - val option_entry: Parser[Options => Options] = - { + val option_entry: Parser[Options => Options] = { command(SECTION) ~! text ^^ { case _ ~ a => (options: Options) => options.set_section(a) } | opt($$$(PUBLIC)) ~ command(OPTION) ~! (position(option_name) ~ $$$(":") ~ option_type ~ @@ -116,16 +109,18 @@ (options: Options) => options.declare(a.isDefined, pos, b, c, d, e, f) } } - val prefs_entry: Parser[Options => Options] = - { + val prefs_entry: Parser[Options => Options] = { option_name ~ ($$$("=") ~! option_value) ^^ { case a ~ (_ ~ b) => (options: Options) => options.add_permissive(a, b) } } - def parse_file(options: Options, file_name: String, content: String, + def parse_file( + options: Options, + file_name: String, + content: String, syntax: Outer_Syntax = options_syntax, - parser: Parser[Options => Options] = option_entry): Options = - { + parser: Parser[Options => Options] = option_entry + ): Options = { val toks = Token.explode(syntax.keywords, content) val ops = parse_all(rep(parser), Token.reader(toks, Token.Pos.file(file_name))) match { @@ -143,8 +138,7 @@ def read_prefs(file: Path = PREFS): String = if (file.is_file) File.read(file) else "" - def init(prefs: String = read_prefs(PREFS), opts: List[String] = Nil): Options = - { + def init(prefs: String = read_prefs(PREFS), opts: List[String] = Nil): Options = { var options = empty for { dir <- Components.directories() @@ -157,14 +151,14 @@ /* Isabelle tool wrapper */ val isabelle_tool = Isabelle_Tool("options", "print Isabelle system options", - Scala_Project.here, args => - { - var build_options = false - var get_option = "" - var list_options = false - var export_file = "" + Scala_Project.here, + { args => + var build_options = false + var get_option = "" + var list_options = false + var export_file = "" - val getopts = Getopts(""" + val getopts = Getopts(""" Usage: isabelle options [OPTIONS] [MORE_OPTIONS ...] Options are: @@ -176,40 +170,39 @@ Report Isabelle system options, augmented by MORE_OPTIONS given as arguments NAME=VAL or NAME. """, - "b" -> (_ => build_options = true), - "g:" -> (arg => get_option = arg), - "l" -> (_ => list_options = true), - "x:" -> (arg => export_file = arg)) + "b" -> (_ => build_options = true), + "g:" -> (arg => get_option = arg), + "l" -> (_ => list_options = true), + "x:" -> (arg => export_file = arg)) - val more_options = getopts(args) - if (get_option == "" && !list_options && export_file == "") getopts.usage() + val more_options = getopts(args) + if (get_option == "" && !list_options && export_file == "") getopts.usage() - val options = - { - val options0 = Options.init() - val options1 = - if (build_options) - Word.explode(Isabelle_System.getenv("ISABELLE_BUILD_OPTIONS")).foldLeft(options0)(_ + _) - else options0 - more_options.foldLeft(options1)(_ + _) - } + val options = { + val options0 = Options.init() + val options1 = + if (build_options) + Word.explode(Isabelle_System.getenv("ISABELLE_BUILD_OPTIONS")).foldLeft(options0)(_ + _) + else options0 + more_options.foldLeft(options1)(_ + _) + } - if (get_option != "") - Output.writeln(options.check_name(get_option).value, stdout = true) + if (get_option != "") + Output.writeln(options.check_name(get_option).value, stdout = true) - if (export_file != "") - File.write(Path.explode(export_file), YXML.string_of_body(options.encode)) + if (export_file != "") + File.write(Path.explode(export_file), YXML.string_of_body(options.encode)) - if (get_option == "" && export_file == "") - Output.writeln(options.print, stdout = true) - }) + if (get_option == "" && export_file == "") + Output.writeln(options.print, stdout = true) + }) } final class Options private( val options: Map[String, Options.Opt] = Map.empty, - val section: String = "") -{ + val section: String = "" +) { override def toString: String = options.iterator.mkString("Options(", ",", ")") private def print_opt(opt: Options.Opt): String = @@ -228,8 +221,7 @@ case _ => error("Unknown option " + quote(name)) } - private def check_type(name: String, typ: Options.Type): Options.Opt = - { + 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) @@ -238,14 +230,12 @@ /* basic operations */ - private def put[A](name: String, typ: Options.Type, value: String): Options = - { + 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 = - { + 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 @@ -258,32 +248,28 @@ /* internal lookup and update */ - class Bool_Access - { + class Bool_Access { def apply(name: String): Boolean = get(name, Options.Bool, Value.Boolean.unapply) def update(name: String, x: Boolean): Options = put(name, Options.Bool, Value.Boolean(x)) } val bool = new Bool_Access - class Int_Access - { + class Int_Access { def apply(name: String): Int = get(name, Options.Int, Value.Int.unapply) def update(name: String, x: Int): Options = put(name, Options.Int, Value.Int(x)) } val int = new Int_Access - class Real_Access - { + class Real_Access { def apply(name: String): Double = get(name, Options.Real, Value.Double.unapply) def update(name: String, x: Double): Options = put(name, Options.Real, Value.Double(x)) } val real = new Real_Access - class String_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) } @@ -297,8 +283,7 @@ /* external updates */ - private def check_value(name: String): Options = - { + private def check_value(name: String): Options = { val opt = check_name(name) opt.typ match { case Options.Bool => bool(name); this @@ -316,8 +301,8 @@ typ_name: String, value: String, standard: Option[Option[String]], - description: String): Options = - { + description: String + ): Options = { options.get(name) match { case Some(other) => error("Duplicate declaration of option " + quote(name) + Position.here(pos) + @@ -347,8 +332,7 @@ } } - def add_permissive(name: String, value: String): Options = - { + def add_permissive(name: String, value: String): Options = { if (options.isDefinedAt(name)) this + (name, value) else { val opt = Options.Opt(false, Position.none, name, Options.Unknown, value, value, None, "", "") @@ -356,14 +340,12 @@ } } - def + (name: String, value: String): Options = - { + 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 = - { + def + (name: String, opt_value: Option[String]): Options = { val opt = check_name(name) opt_value orElse opt.standard_value match { case Some(value) => this + (name, value) @@ -393,8 +375,7 @@ /* encode */ - def encode: XML.Body = - { + def encode: XML.Body = { val opts = for ((_, opt) <- options.toList; if !opt.unknown) yield (opt.pos, (opt.name, (opt.typ.print, opt.value))) @@ -406,8 +387,7 @@ /* save preferences */ - def save_prefs(file: Path = Options.PREFS): Unit = - { + def save_prefs(file: Path = Options.PREFS): Unit = { val defaults: Options = Options.init(prefs = "") val changed = (for { @@ -426,8 +406,7 @@ } -class Options_Variable(init_options: Options) -{ +class Options_Variable(init_options: Options) { private var options = init_options def value: Options = synchronized { options } @@ -435,29 +414,25 @@ private def upd(f: Options => Options): Unit = synchronized { options = f(options) } def += (name: String, x: String): Unit = upd(opts => opts + (name, x)) - class Bool_Access - { + class Bool_Access { def apply(name: String): Boolean = value.bool(name) def update(name: String, x: Boolean): Unit = upd(opts => opts.bool.update(name, x)) } val bool = new Bool_Access - class Int_Access - { + class Int_Access { def apply(name: String): Int = value.int(name) def update(name: String, x: Int): Unit = upd(opts => opts.int.update(name, x)) } val int = new Int_Access - class Real_Access - { + class Real_Access { def apply(name: String): Double = value.real(name) def update(name: String, x: Double): Unit = upd(opts => opts.real.update(name, x)) } val real = new Real_Access - class String_Access - { + class String_Access { def apply(name: String): String = value.string(name) def update(name: String, x: String): Unit = upd(opts => opts.string.update(name, x)) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/System/platform.scala --- a/src/Pure/System/platform.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/System/platform.scala Fri Apr 01 23:26:19 2022 +0200 @@ -7,8 +7,7 @@ package isabelle -object Platform -{ +object Platform { /* platform family */ val is_windows: Boolean = isabelle.setup.Environment.is_windows() @@ -25,8 +24,7 @@ else if (is_windows) Family.windows else error("Failed to determine current platform family") - object Family extends Enumeration - { + object Family extends Enumeration { val linux_arm, linux, macos, windows = Value val list0: List[Value] = List(linux, windows, macos) val list: List[Value] = List(linux, linux_arm, windows, macos) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/System/posix_interrupt.scala --- a/src/Pure/System/posix_interrupt.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/System/posix_interrupt.scala Fri Apr 01 23:26:19 2022 +0200 @@ -10,18 +10,15 @@ import sun.misc.{Signal, SignalHandler} -object POSIX_Interrupt -{ - def handler[A](h: => Unit)(e: => A): A = - { +object POSIX_Interrupt { + def handler[A](h: => Unit)(e: => A): A = { val SIGINT = new Signal("INT") val new_handler = new SignalHandler { def handle(s: Signal): Unit = h } val old_handler = Signal.handle(SIGINT, new_handler) try { e } finally { Signal.handle(SIGINT, old_handler) } } - def exception[A](e: => A): A = - { + def exception[A](e: => A): A = { val thread = Thread.currentThread handler { thread.interrupt() } { e } } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/System/process_result.scala --- a/src/Pure/System/process_result.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/System/process_result.scala Fri Apr 01 23:26:19 2022 +0200 @@ -6,20 +6,17 @@ package isabelle -object Process_Result -{ +object Process_Result { /* return code */ - object RC - { + object RC { val ok = 0 val error = 1 val failure = 2 val interrupt = 130 val timeout = 142 - private def text(rc: Int): String = - { + private def text(rc: Int): String = { val txt = rc match { case `ok` => "OK" @@ -47,8 +44,8 @@ rc: Int, out_lines: List[String] = Nil, err_lines: List[String] = Nil, - timing: Timing = Timing.zero) -{ + timing: Timing = Timing.zero +) { def out: String = Library.trim_line(cat_lines(out_lines)) def err: String = Library.trim_line(cat_lines(err_lines)) @@ -81,15 +78,13 @@ def print_return_code: String = Process_Result.RC.print_long(rc) def print_rc: String = Process_Result.RC.print(rc) - def print: Process_Result = - { + def print: Process_Result = { Output.warning(err) Output.writeln(out) copy(out_lines = Nil, err_lines = Nil) } - def print_stdout: Process_Result = - { + def print_stdout: Process_Result = { Output.warning(err, stdout = true) Output.writeln(out, stdout = true) copy(out_lines = Nil, err_lines = Nil) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/System/progress.scala --- a/src/Pure/System/progress.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/System/progress.scala Fri Apr 01 23:26:19 2022 +0200 @@ -11,10 +11,8 @@ import java.io.{File => JFile} -object Progress -{ - sealed case class Theory(theory: String, session: String = "", percentage: Option[Int] = None) - { +object Progress { + sealed case class Theory(theory: String, session: String = "", percentage: Option[Int] = None) { def message: String = print_session + print_theory + print_percentage def print_session: String = if (session == "") "" else session + ": " @@ -24,8 +22,7 @@ } } -class Progress -{ +class Progress { def echo(msg: String): Unit = {} def echo_if(cond: Boolean, msg: String): Unit = { if (cond) echo(msg) } def theory(theory: Progress.Theory): Unit = {} @@ -39,8 +36,7 @@ @volatile protected var is_stopped = false def stop(): Unit = { is_stopped = true } - def stopped: Boolean = - { + def stopped: Boolean = { if (Thread.interrupted()) is_stopped = true is_stopped } @@ -55,8 +51,8 @@ redirect: Boolean = false, echo: Boolean = false, watchdog: Time = Time.zero, - strict: Boolean = true): Process_Result = - { + strict: Boolean = true + ): Process_Result = { val result = Isabelle_System.bash(script, cwd = cwd, env = env, redirect = redirect, progress_stdout = echo_if(echo, _), @@ -67,8 +63,7 @@ } } -class Console_Progress(verbose: Boolean = false, stderr: Boolean = false) extends Progress -{ +class Console_Progress(verbose: Boolean = false, stderr: Boolean = false) extends Progress { override def echo(msg: String): Unit = Output.writeln(msg, stdout = !stderr, include_empty = true) @@ -76,8 +71,7 @@ if (verbose) echo(theory.message) } -class File_Progress(path: Path, verbose: Boolean = false) extends Progress -{ +class File_Progress(path: Path, verbose: Boolean = false) extends Progress { override def echo(msg: String): Unit = File.append(path, msg + "\n") diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/System/scala.scala --- a/src/Pure/System/scala.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/System/scala.scala Fri Apr 01 23:26:19 2022 +0200 @@ -13,12 +13,10 @@ import scala.tools.nsc.interpreter.{IMain, Results} import scala.tools.nsc.interpreter.shell.ReplReporterImpl -object Scala -{ +object Scala { /** registered functions **/ - abstract class Fun(val name: String, val thread: Boolean = false) - { + abstract class Fun(val name: String, val thread: Boolean = false) { override def toString: String = name def multi: Boolean = true def position: Properties.T = here.position @@ -27,16 +25,14 @@ } abstract class Fun_Strings(name: String, thread: Boolean = false) - extends Fun(name, thread = thread) - { + extends Fun(name, thread = thread) { override def invoke(args: List[Bytes]): List[Bytes] = apply(args.map(_.text)).map(Bytes.apply) def apply(args: List[String]): List[String] } abstract class Fun_String(name: String, thread: Boolean = false) - extends Fun_Strings(name, thread = thread) - { + extends Fun_Strings(name, thread = thread) { override def multi: Boolean = false override def apply(args: List[String]): List[String] = List(apply(Library.the_single(args))) @@ -52,17 +48,14 @@ /** demo functions **/ - object Echo extends Fun_String("echo") - { + object Echo extends Fun_String("echo") { val here = Scala_Project.here def apply(arg: String): String = arg } - object Sleep extends Fun_String("sleep") - { + object Sleep extends Fun_String("sleep") { val here = Scala_Project.here - def apply(seconds: String): String = - { + def apply(seconds: String): String = { val t = seconds match { case Value.Double(s) => Time.seconds(s) @@ -86,12 +79,11 @@ elem <- space_explode(JFile.pathSeparatorChar, elems) if elem.nonEmpty } yield elem - object Compiler - { + object Compiler { def context( error: String => Unit = Exn.error, - jar_dirs: List[JFile] = Nil): Context = - { + jar_dirs: List[JFile] = Nil + ): Context = { def find_jars(dir: JFile): List[String] = File.find_files(dir, file => file.getName.endsWith(".jar")). map(File.absolute_name) @@ -106,24 +98,21 @@ def default_print_writer: PrintWriter = new NewLinePrintWriter(new ConsoleWriter, true) - class Context private [Compiler](val settings: GenericRunnerSettings) - { + class Context private [Compiler](val settings: GenericRunnerSettings) { override def toString: String = settings.toString def interpreter( print_writer: PrintWriter = default_print_writer, - class_loader: ClassLoader = null): IMain = - { - new IMain(settings, new ReplReporterImpl(settings, print_writer)) - { + class_loader: ClassLoader = null + ): IMain = { + new IMain(settings, new ReplReporterImpl(settings, print_writer)) { override def parentClassLoader: ClassLoader = if (class_loader == null) super.parentClassLoader else class_loader } } - def toplevel(interpret: Boolean, source: String): List[String] = - { + def toplevel(interpret: Boolean, source: String): List[String] = { val out = new StringWriter val interp = interpreter(new PrintWriter(out)) val marker = '\u000b' @@ -144,11 +133,9 @@ } } - object Toplevel extends Fun_String("scala_toplevel") - { + object Toplevel extends Fun_String("scala_toplevel") { val here = Scala_Project.here - def apply(arg: String): String = - { + def apply(arg: String): String = { val (interpret, source) = YXML.parse_body(arg) match { case Nil => (false, "") @@ -168,8 +155,7 @@ /* invoke function */ - object Tag extends Enumeration - { + object Tag extends Enumeration { val NULL, OK, ERROR, FAIL, INTERRUPT = Value } @@ -194,8 +180,7 @@ /* protocol handler */ - class Handler extends Session.Protocol_Handler - { + class Handler extends Session.Protocol_Handler { private var session: Session = null private var futures = Map.empty[String, Future[Unit]] @@ -215,8 +200,7 @@ } } - private def cancel(id: String, future: Future[Unit]): Unit = - { + private def cancel(id: String, future: Future[Unit]): Unit = { future.cancel() result(id, Scala.Tag.INTERRUPT, Nil) } @@ -224,8 +208,7 @@ private def invoke_scala(msg: Prover.Protocol_Output): Boolean = synchronized { msg.properties match { case Markup.Invoke_Scala(name, id) => - def body: Unit = - { + def body: Unit = { val (tag, res) = Scala.function_body(name, msg.chunks) result(id, tag, res) } diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/System/system_channel.scala --- a/src/Pure/System/system_channel.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/System/system_channel.scala Fri Apr 01 23:26:19 2022 +0200 @@ -11,13 +11,11 @@ import java.net.{ServerSocket, InetAddress} -object System_Channel -{ +object System_Channel { def apply(): System_Channel = new System_Channel } -class System_Channel private -{ +class System_Channel private { private val server = new ServerSocket(0, 50, Server.localhost) val address: String = Server.print_address(server.getLocalPort) @@ -27,8 +25,7 @@ def shutdown(): Unit = server.close() - def rendezvous(): (OutputStream, InputStream) = - { + def rendezvous(): (OutputStream, InputStream) = { val socket = server.accept try { val out_stream = socket.getOutputStream diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/System/tty_loop.scala --- a/src/Pure/System/tty_loop.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/System/tty_loop.scala Fri Apr 01 23:26:19 2022 +0200 @@ -13,8 +13,8 @@ class TTY_Loop( writer: Writer, reader: Reader, - writer_lock: AnyRef = new Object) -{ + writer_lock: AnyRef = new Object +) { private val console_output = Future.thread[Unit]("console_output", uninterruptible = true) { try { val result = new StringBuilder(100) diff -r 2c3eadf5ca6f -r cd9f2d382014 src/Pure/Thy/bibtex.scala --- a/src/Pure/Thy/bibtex.scala Fri Apr 01 10:54:40 2022 +0000 +++ b/src/Pure/Thy/bibtex.scala Fri Apr 01 23:26:19 2022 +0200 @@ -14,14 +14,12 @@ import scala.util.parsing.input.Reader -object Bibtex -{ +object Bibtex { /** file format **/ def is_bibtex(name: String): Boolean = name.endsWith(".bib") - class File_Format extends isabelle.File_Format - { + class File_Format extends isabelle.File_Format { val format_name: String = "bibtex" val file_ext: String = "bib" @@ -30,8 +28,7 @@ """theory "bib" imports Pure begin bibtex_file """ + Outer_Syntax.quote_string(name) + """ end""" - override def html_document(snapshot: Document.Snapshot): Option[Presentation.HTML_Document] = - { + override def html_document(snapshot: Document.Snapshot): Option[Presentation.HTML_Document] = { val name = snapshot.node_name if (detect(name.node)) { val title = "Bibliography " + quote(snapshot.node_name.path.file_name) @@ -50,8 +47,7 @@ /** bibtex errors **/ - def bibtex_errors(dir: Path, root_name: String): List[String] = - { + def bibtex_errors(dir: Path, root_name: String): List[String] = { val log_path = dir + Path.explode(root_name).ext("blg") if (log_path.is_file) { val Error1 = """^(I couldn't open database file .+)$""".r @@ -78,8 +74,7 @@ /** check database **/ - def check_database(database: String): (List[(String, Position.T)], List[(String, Position.T)]) = - { + def check_database(database: String): (List[(String, Position.T)], List[(String, Position.T)]) = { val chunks = parse(Line.normalize(database)) var chunk_pos = Map.empty[String, Position.T] val tokens = new mutable.ListBuffer[(Token, Position.T)] @@ -89,8 +84,7 @@ def make_pos(length: Int): Position.T = Position.Offset(offset) ::: Position.End_Offset(offset + length) ::: Position.Line(line) - def advance_pos(tok: Token): Unit = - { + def advance_pos(tok: Token): Unit = { for (s <- Symbol.iterator(tok.source)) { if (Symbol.is_newline(s)) line += 1 offset += 1 @@ -111,8 +105,7 @@ } } - Isabelle_System.with_tmp_dir("bibtex")(tmp_dir => - { + Isabelle_System.with_tmp_dir("bibtex") { tmp_dir => File.write(tmp_dir + Path.explode("root.bib"), tokens.iterator.map(p => p._1.source).mkString("", "\n", "\n")) File.write(tmp_dir + Path.explode("root.aux"), @@ -145,14 +138,12 @@ ).partition(_._1) } (errors.map(_._2), warnings.map(_._2)) - }) + } } - object Check_Database extends Scala.Fun_String("bibtex_check_database") - { + object Check_Database extends Scala.Fun_String("bibtex_check_database") { val here = Scala_Project.here - def apply(database: String): String = - { + def apply(database: String): String = { import XML.Encode._ YXML.string_of_body(pair(list(pair(string, properties)), list(pair(string, properties)))( check_database(database))) @@ -165,8 +156,7 @@ /* entries */ - def entries(text: String): List[Text.Info[String]] = - { + def entries(text: String): List[Text.Info[String]] = { val result = new mutable.ListBuffer[Text.Info[String]] var offset = 0 for (chunk <- Bibtex.parse(text)) { @@ -178,9 +168,9 @@ result.toList } - def entries_iterator[A, B <: Document.Model](models: Map[A, B]) - : Iterator[Text.Info[(String, B)]] = - { + def entries_iterator[A, B <: Document.Model]( + models: Map[A, B] + ): Iterator[Text.Info[(String, B)]] = { for { (_, model) <- models.iterator info <- model.bibtex_entries.iterator @@ -191,9 +181,11 @@ /* completion */ def completion[A, B <: Document.Model]( - history: Completion.History, rendering: Rendering, caret: Text.Offset, - models: Map[A, B]): Option[Completion.Result] = - { + history: Completion.History, + rendering: Rendering, + caret: Text.Offset, + models: Map[A, B] + ): Option[Completion.Result] = { for { Text.Info(r, name) <- rendering.citations(rendering.before_caret_range(caret)).headOption name1 <- Completion.clean_name(name) @@ -245,8 +237,8 @@ kind: String, required: List[String], optional_crossref: List[String], - optional_other: List[String]) - { + optional_other: List[String] + ) { val optional_standard: List[String] = List("url", "doi", "ee") def is_required(s: String): Boolean = required.contains(s.toLowerCase) @@ -329,10 +321,8 @@ /** tokens and chunks **/ - object Token - { - object Kind extends Enumeration - { + object Token { + object Kind extends Enumeration { val COMMAND = Value("command") val ENTRY = Value("entry") val KEYWORD = Value("keyword") @@ -346,8 +336,7 @@ } } - sealed case class Token(kind: Token.Kind.Value, source: String) - { + sealed case class Token(kind: Token.Kind.Value, source: String) { def is_kind: Boolean = kind == Token.Kind.COMMAND || kind == Token.Kind.ENTRY || @@ -364,8 +353,7 @@ kind == Token.Kind.KEYWORD && (source == "{" || source == "(") } - case class Chunk(kind: String, tokens: List[Token]) - { + case class Chunk(kind: String, tokens: List[Token]) { val source = tokens.map(_.source).mkString private val content: Option[List[Token]] = @@ -423,8 +411,7 @@ // See also https://ctan.org/tex-archive/biblio/bibtex/base/bibtex.web // module @. - object Parsers extends RegexParsers - { + object Parsers extends RegexParsers { /* white space and comments */ override val whiteSpace = "".r @@ -447,13 +434,11 @@ // see also bibtex.web: scan_a_field_token_and_eat_white, scan_balanced_braces private def delimited_depth(delim: Delimited): Parser[(String, Delimited)] = - new Parser[(String, Delimited)] - { + new Parser[(String, Delimited)] { require(if (delim.quoted) delim.depth > 0 else delim.depth >= 0, "bad delimiter depth") - def apply(in: Input) = - { + def apply(in: Input) = { val start = in.offset val end = in.source.length @@ -567,8 +552,7 @@ val chunk: Parser[Chunk] = ignored | (item | recover_item) - def chunk_line(ctxt: Line_Context): Parser[(Chunk, Line_Context)] = - { + def chunk_line(ctxt: Line_Context): Parser[(Chunk, Line_Context)] = { ctxt match { case Ignored => ignored_line | @@ -610,8 +594,7 @@ case _ => error("Unexpected failure to parse input:\n" + input.toString) } - def parse_line(input: CharSequence, context: Line_Context): (List[Chunk], Line_Context) = - { + def parse_line(input: CharSequence, context: Line_Context): (List[Chunk], Line_Context) = { var in: Reader[Char] = Scan.char_reader(input) val chunks = new mutable.ListBuffer[Chunk] var ctxt = context @@ -644,15 +627,13 @@ body: Boolean = false, citations: List[String] = List("*"), style: String = "", - chronological: Boolean = false): String = - { - Isabelle_System.with_tmp_dir("bibtex")(tmp_dir => - { + chronological: Boolean = false + ): String = { + Isabelle_System.with_tmp_dir("bibtex") { tmp_dir => /* database files */ val bib_files = bib.map(_.drop_ext) - val bib_names = - { + val bib_names = { val names0 = bib_files.map(_.file_name) if (Library.duplicates(names0).isEmpty) names0 else names0.zipWithIndex.map({ case (name, i) => (i + 1).toString + "-" + name }) @@ -701,6 +682,6 @@ dropWhile(line => !line.startsWith("