--- a/src/HOL/SPARK/Tools/spark.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/HOL/SPARK/Tools/spark.scala Fri Apr 01 17:06:10 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")
}
}
--- a/src/HOL/Tools/ATP/system_on_tptp.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/HOL/Tools/ATP/system_on_tptp.scala Fri Apr 01 17:06:10 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))
--- a/src/HOL/Tools/Mirabelle/mirabelle.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/HOL/Tools/Mirabelle/mirabelle.scala Fri Apr 01 17:06:10 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,8 +109,7 @@
/* Isabelle tool wrapper */
val isabelle_tool = Isabelle_Tool("mirabelle", "testing tool for automated proof tools",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
val build_options = Word.explode(Isabelle_System.getenv("ISABELLE_BUILD_OPTIONS"))
var options = Options.init(opts = build_options)
--- a/src/HOL/Tools/Nitpick/kodkod.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/HOL/Tools/Nitpick/kodkod.scala Fri Apr 01 17:06:10 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))
}
--- a/src/Pure/Admin/afp.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/afp.scala Fri Apr 01 17:06:10 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) }
--- a/src/Pure/Admin/build_csdp.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/build_csdp.scala Fri Apr 01 17:06:10 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
@@ -171,8 +167,7 @@
val isabelle_tool =
Isabelle_Tool("build_csdp", "build prover component from official download", Scala_Project.here,
- args =>
- {
+ args => {
var target_dir = Path.current
var mingw = MinGW.none
var download_url = default_download_url
--- a/src/Pure/Admin/build_cygwin.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/build_cygwin.scala Fri Apr 01 17:06:10 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"))
@@ -62,8 +60,7 @@
val isabelle_tool =
Isabelle_Tool("build_cygwin", "produce pre-canned Cygwin distribution for Isabelle",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
var mirror = default_mirror
var more_packages: List[String] = Nil
--- a/src/Pure/Admin/build_doc.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/build_doc.scala Fri Apr 01 17:06:10 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,8 +73,7 @@
/* Isabelle tool wrapper */
val isabelle_tool =
- Isabelle_Tool("build_doc", "build Isabelle documentation", Scala_Project.here, args =>
- {
+ Isabelle_Tool("build_doc", "build Isabelle documentation", Scala_Project.here, args => {
var all_docs = false
var max_jobs = 1
var sequential = false
--- a/src/Pure/Admin/build_e.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/build_e.scala Fri Apr 01 17:06:10 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 ""
}
@@ -110,8 +107,7 @@
val isabelle_tool =
Isabelle_Tool("build_e", "build prover component from source distribution", Scala_Project.here,
- args =>
- {
+ args => {
var target_dir = Path.current
var version = default_version
var download_url = default_download_url
--- a/src/Pure/Admin/build_fonts.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/build_fonts.scala Fri Apr 01 17:06:10 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 + " ...")
@@ -342,8 +336,7 @@
/* Isabelle tool wrapper */
val isabelle_tool =
- Isabelle_Tool("build_fonts", "construct Isabelle fonts", Scala_Project.here, args =>
- {
+ Isabelle_Tool("build_fonts", "construct Isabelle fonts", Scala_Project.here, args => {
var source_dirs: List[Path] = Nil
val getopts = Getopts("""
--- a/src/Pure/Admin/build_history.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/build_history.scala Fri Apr 01 17:06:10 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,13 +281,11 @@
build_out_progress.echo("Reading session build info ...")
val session_build_info =
- build_info.finished_sessions.flatMap(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 =>
- {
+ using(SQLite.open_database(database))(db => {
val theory_timings =
try {
store.read_theory_timings(db, session_name).map(ps =>
@@ -314,8 +308,7 @@
build_out_progress.echo("Reading ML statistics ...")
val ml_statistics =
- build_info.finished_sessions.flatMap(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)
@@ -342,8 +335,7 @@
build_out_progress.echo("Reading error messages ...")
val session_errors =
- build_info.failed_sessions.flatMap(session_name =>
- {
+ build_info.failed_sessions.flatMap(session_name => {
val database = isabelle_output + store.database(session_name)
val errors =
if (database.is_file) {
@@ -358,8 +350,7 @@
build_out_progress.echo("Reading heap sizes ...")
val heap_sizes =
- build_info.finished_sessions.flatMap(session_name =>
- {
+ 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)")
@@ -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)
--- a/src/Pure/Admin/build_jcef.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/build_jcef.scala Fri Apr 01 17:06:10 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)
@@ -144,8 +142,7 @@
val isabelle_tool =
Isabelle_Tool("build_jcef", "build component for Java Chromium Embedded Framework",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
var target_dir = Path.current
var base_url = default_url
var version = default_version
--- a/src/Pure/Admin/build_jdk.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/build_jdk.scala Fri Apr 01 17:06:10 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,12 +148,11 @@
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 =>
- {
+ Isabelle_System.with_tmp_dir("jdk")(dir => {
progress.echo("Extracting ...")
val platforms = archives.map(extract_archive(dir, _))
@@ -214,8 +209,7 @@
val isabelle_tool =
Isabelle_Tool("build_jdk", "build Isabelle jdk component from original archives",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
var target_dir = Path.current
val getopts = Getopts("""
--- a/src/Pure/Admin/build_jedit.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/build_jedit.scala Fri Apr 01 17:06:10 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")
@@ -210,8 +203,7 @@
}
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"
@@ -519,8 +511,7 @@
val isabelle_tool =
Isabelle_Tool("build_jedit", "build Isabelle component from the jEdit text-editor",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
var target_dir = Path.current
var java_home = default_java_home
var original = false
--- a/src/Pure/Admin/build_log.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/build_log.scala Fri Apr 01 17:06:10 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) {
@@ -945,11 +918,9 @@
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)
@@ -961,11 +932,9 @@
})
}
- 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
@@ -992,11 +961,9 @@
})
}
- 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
@@ -1016,11 +983,9 @@
})
}
- 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) },
@@ -1035,18 +1000,15 @@
})
}
- 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 {
@@ -1111,8 +1070,7 @@
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(
--- a/src/Pure/Admin/build_minisat.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/build_minisat.scala Fri Apr 01 17:06:10 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
@@ -123,8 +121,7 @@
val isabelle_tool =
Isabelle_Tool("build_minisat", "build prover component from sources", Scala_Project.here,
- args =>
- {
+ args => {
var target_dir = Path.current
var download_url = default_download_url
var component_name = ""
--- a/src/Pure/Admin/build_pdfjs.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/build_pdfjs.scala Fri Apr 01 17:06:10 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,8 +38,7 @@
/* 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),
@@ -74,8 +72,7 @@
val isabelle_tool =
Isabelle_Tool("build_pdfjs", "build component for Mozilla PDF.js",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
var target_dir = Path.current
var base_url = default_url
var version = default_version
--- a/src/Pure/Admin/build_polyml.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/build_polyml.scala Fri Apr 01 17:06:10 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,8 +201,7 @@
/** Isabelle tool wrappers **/
val isabelle_tool1 =
- Isabelle_Tool("build_polyml", "build Poly/ML from sources", Scala_Project.here, args =>
- {
+ 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
@@ -242,8 +239,7 @@
val isabelle_tool2 =
Isabelle_Tool("build_polyml_component", "make skeleton for Poly/ML component",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
var sha1_root: Option[Path] = None
val getopts = Getopts("""
--- a/src/Pure/Admin/build_release.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/build_release.scala Fri Apr 01 17:06:10 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)
@@ -136,8 +129,11 @@
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 =
@@ -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, """<?xml version="1.0" ?>
<!DOCTYPE plist PUBLIC "-//Apple//DTD PLIST 1.0//EN" "http://www.apple.com/DTDs/PropertyList-1.0.dtd">
<plist version="1.0">
@@ -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,8 +496,7 @@
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").
@@ -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,13 +562,11 @@
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
@@ -833,8 +815,7 @@
progress.echo_warning("Library archive already exists: " + context.isabelle_library_archive)
}
else {
- Isabelle_System.with_tmp_dir("build_release")(tmp_dir =>
- {
+ 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))
@@ -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
--- a/src/Pure/Admin/build_scala.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/build_scala.scala Fri Apr 01 17:06:10 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,8 +29,7 @@
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),
@@ -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,8 +139,7 @@
val isabelle_tool =
Isabelle_Tool("build_scala", "build Isabelle Scala component from official downloads",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
var target_dir = Path.current
val getopts = Getopts("""
--- a/src/Pure/Admin/build_spass.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/build_spass.scala Fri Apr 01 17:06:10 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")
@@ -150,8 +147,7 @@
val isabelle_tool =
Isabelle_Tool("build_spass", "build prover component from source distribution",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
var target_dir = Path.current
var download_url = default_download_url
var verbose = false
--- a/src/Pure/Admin/build_sqlite.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/build_sqlite.scala Fri Apr 01 17:06:10 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
@@ -85,8 +83,7 @@
val isabelle_tool =
Isabelle_Tool("build_sqlite", "build Isabelle sqlite-jdbc component from official download",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
var target_dir = Path.current
val getopts = Getopts("""
--- a/src/Pure/Admin/build_status.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/build_status.scala Fri Apr 01 17:06:10 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
@@ -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,8 +582,7 @@
val isabelle_tool =
Isabelle_Tool("build_status", "present recent build status information from database",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
var target_dir = default_target_dir
var ml_statistics = false
var only_sessions = Set.empty[String]
--- a/src/Pure/Admin/build_vampire.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/build_vampire.scala Fri Apr 01 17:06:10 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
@@ -128,8 +126,7 @@
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
--- a/src/Pure/Admin/build_verit.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/build_verit.scala Fri Apr 01 17:06:10 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
@@ -124,8 +122,7 @@
val isabelle_tool =
Isabelle_Tool("build_verit", "build prover component from official download",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
var target_dir = Path.current
var mingw = MinGW.none
var download_url = default_download_url
--- a/src/Pure/Admin/build_zipperposition.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/build_zipperposition.scala Fri Apr 01 17:06:10 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")
@@ -93,8 +91,7 @@
val isabelle_tool =
Isabelle_Tool("build_zipperposition", "build prover component from OPAM repository",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
var target_dir = Path.current
var version = default_version
var verbose = false
--- a/src/Pure/Admin/check_sources.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/check_sources.scala Fri Apr 01 17:06:10 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,8 +59,7 @@
val isabelle_tool =
Isabelle_Tool("check_sources", "some sanity checks for Isabelle sources",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
val getopts = Getopts("""
Usage: isabelle check_sources [ROOT_DIRS...]
--- a/src/Pure/Admin/ci_profile.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/ci_profile.scala Fri Apr 01 17:06:10 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()
--- a/src/Pure/Admin/isabelle_cronjob.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/isabelle_cronjob.scala Fri Apr 01 17:06:10 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,7 @@
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 +60,7 @@
})
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 +70,7 @@
/* 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,8 +79,7 @@
/* build release */
val build_release: Logger_Task =
- Logger_Task("build_release", logger =>
- {
+ Logger_Task("build_release", logger => {
Isabelle_Devel.release_snapshot(logger.options, get_rev(), get_afp_rev())
})
@@ -95,8 +90,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 +100,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 +122,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 +147,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 +165,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)
@@ -310,8 +305,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,14 +394,15 @@
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 =>
- {
+ Logger_Task(task_name, logger => {
+ using(r.ssh_session(logger.ssh_context))(ssh => {
val results =
Build_History.remote_build_history(ssh,
isabelle_repos,
@@ -438,20 +433,18 @@
/** 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 +463,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 +484,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 +506,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 +515,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,10 +544,8 @@
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 =
- {
+ 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)
@@ -576,8 +565,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 +609,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
--- a/src/Pure/Admin/isabelle_devel.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/isabelle_devel.scala Fri Apr 01 17:06:10 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,13 +32,10 @@
/* release snapshot */
- def release_snapshot(options: Options, rev: String, afp_rev: String): Unit =
- {
- Isabelle_System.with_tmp_dir("isadist")(target_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 =>
- {
+ 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,
@@ -53,11 +48,9 @@
/* 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))
@@ -67,8 +60,7 @@
/* 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))
}
--- a/src/Pure/Admin/jenkins.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/jenkins.scala Fri Apr 01 17:06:10 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 =
--- a/src/Pure/Admin/other_isabelle.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Admin/other_isabelle.scala Fri Apr 01 17:06:10 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
--- a/src/Pure/Concurrent/consumer_thread.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Concurrent/consumer_thread.scala Fri Apr 01 17:06:10 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()
}
--- a/src/Pure/Concurrent/counter.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Concurrent/counter.scala Fri Apr 01 17:06:10 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 {
--- a/src/Pure/Concurrent/delay.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Concurrent/delay.scala Fri Apr 01 17:06:10 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
}
--- a/src/Pure/Concurrent/event_timer.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Concurrent/event_timer.scala Fri Apr 01 17:06:10 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))
--- a/src/Pure/Concurrent/future.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Concurrent/future.scala Fri Apr 01 17:06:10 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
--- a/src/Pure/Concurrent/isabelle_thread.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Concurrent/isabelle_thread.scala Fri Apr 01 17:06:10 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)
--- a/src/Pure/Concurrent/mailbox.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Concurrent/mailbox.scala Fri Apr 01 17:06:10 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(", ",", ")")
--- a/src/Pure/Concurrent/par_list.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Concurrent/par_list.scala Fri Apr 01 17:06:10 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(
--- a/src/Pure/Concurrent/synchronized.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Concurrent/synchronized.scala Fri Apr 01 17:06:10 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 =>
--- a/src/Pure/GUI/color_value.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/GUI/color_value.scala Fri Apr 01 17:06:10 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)
--- a/src/Pure/GUI/desktop_app.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/GUI/desktop_app.scala Fri Apr 01 17:06:10 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
--- a/src/Pure/GUI/gui.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/GUI/gui.scala Fri Apr 01 17:06:10 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))
}
--- a/src/Pure/GUI/gui_thread.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/GUI/gui_thread.scala Fri Apr 01 17:06:10 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]
--- a/src/Pure/GUI/popup.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/GUI/popup.scala Fri Apr 01 17:06:10 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)
--- a/src/Pure/GUI/wrap_panel.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/GUI/wrap_panel.scala Fri Apr 01 17:06:10 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
--- a/src/Pure/General/antiquote.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/antiquote.scala Fri Apr 01 17:06:10 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
--- a/src/Pure/General/bytes.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/bytes.scala Fri Apr 01 17:06:10 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)
}
--- a/src/Pure/General/cache.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/cache.scala Fri Apr 01 17:06:10 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"
--- a/src/Pure/General/codepoint.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/codepoint.scala Fri Apr 01 17:06:10 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)
--- a/src/Pure/General/comment.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/comment.scala Fri Apr 01 17:06:10 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)
--- a/src/Pure/General/completion.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/completion.scala Fri Apr 01 17:06:10 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 =
@@ -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, _) =>
--- a/src/Pure/General/csv.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/csv.scala Fri Apr 01 17:06:10 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"
--- a/src/Pure/General/date.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/date.scala Fri Apr 01 17:06:10 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,8 +46,7 @@
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] =
@@ -60,9 +55,11 @@
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))
--- a/src/Pure/General/exn.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/exn.scala Fri Apr 01 17:06:10 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)
}
--- a/src/Pure/General/file.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/file.scala Fri Apr 01 17:06:10 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))
}
--- a/src/Pure/General/file_watcher.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/file_watcher.scala Fri Apr 01 17:06:10 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,20 +81,17 @@
/* 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 =>
- {
+ state.change_result(st => {
val (remove, changed) =
st.dirs.collectFirst({ case (dir, key1) if key == key1 => dir }) match {
case Some(dir) =>
@@ -126,8 +119,7 @@
/* shutdown */
- override def shutdown(): Unit =
- {
+ override def shutdown(): Unit = {
watcher_thread.interrupt()
watcher_thread.join()
delay_changed.revoke()
--- a/src/Pure/General/graph.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/graph.scala Fri Apr 01 17:06:10 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)
@@ -68,8 +66,7 @@
}
-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")
--- a/src/Pure/General/graph_display.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/graph_display.scala Fri Apr 01 17:06:10 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
--- a/src/Pure/General/graphics_file.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/graphics_file.scala Fri Apr 01 17:06:10 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))
--- a/src/Pure/General/http.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/http.scala Fri Apr 01 17:06:10 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()
@@ -109,10 +104,12 @@
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))
@@ -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]
@@ -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
--- a/src/Pure/General/json.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/json.scala Fri Apr 01 17:06:10 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)
}
--- a/src/Pure/General/json_api.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/json_api.scala Fri Apr 01 17:06:10 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 =
--- a/src/Pure/General/linear_set.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/linear_set.scala Fri Apr 01 17:06:10 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] =
--- a/src/Pure/General/logger.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/logger.scala Fri Apr 01 17:06:10 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
--- a/src/Pure/General/long_name.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/long_name.scala Fri Apr 01 17:06:10 2022 +0200
@@ -7,8 +7,7 @@
package isabelle
-object Long_Name
-{
+object Long_Name {
val separator = "."
val separator_char = '.'
--- a/src/Pure/General/mailman.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/mailman.scala Fri Apr 01 17:06:10 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 =
"""<title>The ([^</>]*) Archives</title>""".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 = """<li><a name="\d+" href="(msg\d+\.html)">""".r
private object Head extends
@@ -489,8 +473,7 @@
private object Body extends
Message_Chunk(bg = "<!--X-Body-of-Message-->", en = "<!--X-Body-of-Message-End-->")
- 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 = """<a href="[^"]*">""" + s + """</a>"""
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 = """<LI><A HREF="(\d+\.html)">""".r
private object Head extends Message_Chunk(en = "<!--beginarticle-->")
private object Body extends Message_Chunk(bg = "<!--beginarticle-->", en = "<!--endarticle-->")
- 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")
--- a/src/Pure/General/mercurial.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/mercurial.scala Fri Apr 01 17:06:10 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,11 +66,9 @@
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")
@@ -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,8 +397,7 @@
val isabelle_tool =
Isabelle_Tool("hg_setup", "setup remote vs. local Mercurial repository",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
var remote_name = ""
var path_name = default_path_name
var remote_exists = false
--- a/src/Pure/General/multi_map.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/multi_map.scala Fri Apr 01 17:06:10 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 {
--- a/src/Pure/General/output.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/output.scala Fri Apr 01 17:06:10 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")
--- a/src/Pure/General/path.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/path.scala Fri Apr 01 17:06:10 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
--- a/src/Pure/General/position.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/position.scala Fri Apr 01 17:06:10 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 -> _) ++
--- a/src/Pure/General/pretty.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/pretty.scala Fri Apr 01 17:06:10 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] =
--- a/src/Pure/General/properties.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/properties.scala Fri Apr 01 17:06:10 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 {
--- a/src/Pure/General/rdf.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/rdf.scala Fri Apr 01 17:06:10 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")
--- a/src/Pure/General/scan.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/scan.scala Fri Apr 01 17:06:10 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,12 +396,10 @@
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) =>
- {
+ (buf: Array[Char], offset: Int, length: Int) => {
var i = 0
var c = 0
var eof = false
@@ -441,8 +412,7 @@
})
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
--- a/src/Pure/General/sha1.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/sha1.scala Fri Apr 01 17:06:10 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,15 +27,13 @@
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 =>
- {
+ make_digest(sha => using(new FileInputStream(file))(stream => {
val buf = new Array[Byte](65536)
var m = 0
var cont = true
--- a/src/Pure/General/sql.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/sql.scala Fri Apr 01 17:06:10 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)
}
--- a/src/Pure/General/ssh.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/ssh.scala Fri Apr 01 17:06:10 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 {
@@ -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 = ""
--- a/src/Pure/General/symbol.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/symbol.scala Fri Apr 01 17:06:10 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 {
--- a/src/Pure/General/time.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/time.scala Fri Apr 01 17:06:10 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),
--- a/src/Pure/General/timing.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/timing.scala Fri Apr 01 17:06:10 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)
--- a/src/Pure/General/untyped.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/untyped.scala Fri Apr 01 17:06:10 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)
--- a/src/Pure/General/url.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/url.scala Fri Apr 01 17:06:10 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 =>
--- a/src/Pure/General/utf8.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/utf8.scala Fri Apr 01 17:06:10 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
--- a/src/Pure/General/uuid.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/uuid.scala Fri Apr 01 17:06:10 2022 +0200
@@ -7,8 +7,7 @@
package isabelle
-object UUID
-{
+object UUID {
type T = java.util.UUID
def random(): T = java.util.UUID.randomUUID()
--- a/src/Pure/General/value.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/value.scala Fri Apr 01 17:06:10 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
}
--- a/src/Pure/General/word.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/word.scala Fri Apr 01 17:06:10 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)
--- a/src/Pure/General/xz.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/General/xz.scala Fri Apr 01 17:06:10 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
--- a/src/Pure/Isar/document_structure.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Isar/document_structure.scala Fri Apr 01 17:06:10 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)
--- a/src/Pure/Isar/keyword.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Isar/keyword.scala Fri Apr 01 17:06:10 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) {
--- a/src/Pure/Isar/line_structure.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Isar/line_structure.scala Fri Apr 01 17:06:10 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)
--- a/src/Pure/Isar/outer_syntax.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Isar/outer_syntax.scala Fri Apr 01 17:06:10 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()
--- a/src/Pure/Isar/parse.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Isar/parse.scala Fri Apr 01 17:06:10 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)
--- a/src/Pure/Isar/token.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Isar/token.scala Fri Apr 01 17:06:10 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) &&
--- a/src/Pure/ML/ml_console.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/ML/ml_console.scala Fri Apr 01 17:06:10 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
--- a/src/Pure/ML/ml_lex.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/ML/ml_lex.scala Fri Apr 01 17:06:10 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
--- a/src/Pure/ML/ml_process.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/ML/ml_process.scala Fri Apr 01 17:06:10 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,8 +129,7 @@
cwd = cwd,
env = bash_env,
redirect = redirect,
- cleanup = () =>
- {
+ cleanup = () => {
isabelle_process_options.delete
init_session.delete
Isabelle_System.rm_tree(isabelle_tmp)
@@ -144,8 +141,7 @@
/* Isabelle tool wrapper */
val isabelle_tool = Isabelle_Tool("process", "raw ML process (batch mode)",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
var dirs: List[Path] = Nil
var eval_args: List[String] = Nil
var logic = Isabelle_System.getenv("ISABELLE_LOGIC")
--- a/src/Pure/ML/ml_profiling.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/ML/ml_profiling.scala Fri Apr 01 17:06:10 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) {
--- a/src/Pure/ML/ml_statistics.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/ML/ml_statistics.scala Fri Apr 01 17:06:10 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)
--- a/src/Pure/ML/ml_syntax.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/ML/ml_syntax.scala Fri Apr 01 17:06:10 2022 +0200
@@ -7,8 +7,7 @@
package isabelle
-object ML_Syntax
-{
+object ML_Syntax {
/* numbers */
private def signed(s: String): String =
--- a/src/Pure/PIDE/byte_message.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/PIDE/byte_message.scala Fri Apr 01 17:06:10 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))
--- a/src/Pure/PIDE/command.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/PIDE/command.scala Fri Apr 01 17:06:10 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 =>
--- a/src/Pure/PIDE/command_span.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/PIDE/command_span.scala Fri Apr 01 17:06:10 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)))
}
--- a/src/Pure/PIDE/document.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/PIDE/document.scala Fri Apr 01 17:06:10 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 => { case _ => 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
--- a/src/Pure/PIDE/document_id.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/PIDE/document_id.scala Fri Apr 01 17:06:10 2022 +0200
@@ -9,8 +9,7 @@
package isabelle
-object Document_ID
-{
+object Document_ID {
type Generic = Long
type Version = Generic
type Command = Generic
--- a/src/Pure/PIDE/document_status.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/PIDE/document_status.scala Fri Apr 01 17:06:10 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
--- a/src/Pure/PIDE/editor.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/PIDE/editor.scala Fri Apr 01 17:06:10 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
}
--- a/src/Pure/PIDE/headless.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/PIDE/headless.scala Fri Apr 01 17:06:10 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)
}
@@ -478,9 +467,10 @@
(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,40 +601,35 @@
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)
@@ -650,11 +637,11 @@
})
}
- 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)
})
--- a/src/Pure/PIDE/line.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/PIDE/line.scala Fri Apr 01 17:06:10 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 =
--- a/src/Pure/PIDE/markup.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/PIDE/markup.scala Fri Apr 01 17:06:10 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)
--- a/src/Pure/PIDE/markup_tree.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/PIDE/markup_tree.scala Fri Apr 01 17:06:10 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) {
--- a/src/Pure/PIDE/protocol.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/PIDE/protocol.scala Fri Apr 01 17:06:10 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(
@@ -374,11 +357,9 @@
toks_yxml :: toks_sources)
}
- def define_commands(resources: Resources, commands: List[Command]): Unit =
- {
+ 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)
@@ -389,8 +370,7 @@
}))
}
- 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 +392,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 +427,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))) }
--- a/src/Pure/PIDE/protocol_handlers.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/PIDE/protocol_handlers.scala Fri Apr 01 17:06:10 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 =
--- a/src/Pure/PIDE/protocol_message.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/PIDE/protocol_message.scala Fri Apr 01 17:06:10 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 + " = "
--- a/src/Pure/PIDE/prover.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/PIDE/prover.scala Fri Apr 01 17:06:10 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(_)))
}
--- a/src/Pure/PIDE/query_operation.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/PIDE/query_operation.scala Fri Apr 01 17:06:10 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)
--- a/src/Pure/PIDE/rendering.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/PIDE/rendering.scala Fri Apr 01 17:06:10 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, _ =>
{
--- a/src/Pure/PIDE/resources.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/PIDE/resources.scala Fri Apr 01 17:06:10 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)
--- a/src/Pure/PIDE/session.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/PIDE/session.scala Fri Apr 01 17:06:10 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(
{
--- a/src/Pure/PIDE/text.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/PIDE/text.scala Fri Apr 01 17:06:10 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 + ")"
--- a/src/Pure/PIDE/xml.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/PIDE/xml.scala Fri Apr 01 17:06:10 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 = "<?xml version=\"1.0\" encoding=\"utf-8\"?>\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 =
--- a/src/Pure/PIDE/yxml.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/PIDE/yxml.scala Fri Apr 01 17:06:10 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) }
}
--- a/src/Pure/ROOT.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/ROOT.scala Fri Apr 01 17:06:10 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:_*)
--- a/src/Pure/System/bash.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/System/bash.scala Fri Apr 01 17:06:10 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 +
--- a/src/Pure/System/command_line.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/System/command_line.scala Fri Apr 01 17:06:10 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 =
--- a/src/Pure/System/components.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/System/components.scala Fri Apr 01 17:06:10 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,8 +247,7 @@
// 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))
})
@@ -321,8 +310,7 @@
val isabelle_tool =
Isabelle_Tool("build_components", "build and publish Isabelle components",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
var publish = false
var update_components_sha1 = false
var force = false
--- a/src/Pure/System/executable.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/System/executable.scala Fri Apr 01 17:06:10 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)
--- a/src/Pure/System/getopts.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/System/getopts.scala Fri Apr 01 17:06:10 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)
}
--- a/src/Pure/System/isabelle_charset.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/System/isabelle_charset.scala Fri Apr 01 17:06:10 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()
--- a/src/Pure/System/isabelle_fonts.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/System/isabelle_fonts.scala Fri Apr 01 17:06:10 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)
}
--- a/src/Pure/System/isabelle_platform.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/System/isabelle_platform.scala Fri Apr 01 17:06:10 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)))
--- a/src/Pure/System/isabelle_process.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/System/isabelle_process.scala Fri Apr 01 17:06:10 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
--- a/src/Pure/System/isabelle_system.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/System/isabelle_system.scala Fri Apr 01 17:06:10 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,10 +417,8 @@
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),
@@ -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) }
--- a/src/Pure/System/isabelle_tool.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/System/isabelle_tool.scala Fri Apr 01 17:06:10 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 &&
@@ -74,12 +71,11 @@
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)
- }
+ (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)
+ }
})
@@ -97,8 +93,7 @@
/* list tools */
- abstract class Entry
- {
+ abstract class Entry {
def name: String
def position: Properties.T
def description: String
@@ -109,18 +104,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 +126,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 +140,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 +165,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
}
--- a/src/Pure/System/java_statistics.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/System/java_statistics.scala Fri Apr 01 17:06:10 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
--- a/src/Pure/System/linux.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/System/linux.scala Fri Apr 01 17:06:10 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
}
--- a/src/Pure/System/mingw.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/System/mingw.scala Fri Apr 01 17:06:10 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")) }
--- a/src/Pure/System/numa.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/System/numa.scala Fri Apr 01 17:06:10 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
--- a/src/Pure/System/options.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/System/options.scala Fri Apr 01 17:06:10 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,8 +151,7 @@
/* Isabelle tool wrapper */
val isabelle_tool = Isabelle_Tool("options", "print Isabelle system options",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
var build_options = false
var get_option = ""
var list_options = false
@@ -184,8 +177,7 @@
val more_options = getopts(args)
if (get_option == "" && !list_options && export_file == "") getopts.usage()
- val options =
- {
+ val options = {
val options0 = Options.init()
val options1 =
if (build_options)
@@ -208,8 +200,8 @@
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 +220,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 +229,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 +247,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 +282,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 +300,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 +331,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 +339,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 +374,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 +386,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 +405,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 +413,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))
}
--- a/src/Pure/System/platform.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/System/platform.scala Fri Apr 01 17:06:10 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)
--- a/src/Pure/System/posix_interrupt.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/System/posix_interrupt.scala Fri Apr 01 17:06:10 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 }
}
--- a/src/Pure/System/process_result.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/System/process_result.scala Fri Apr 01 17:06:10 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)
--- a/src/Pure/System/progress.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/System/progress.scala Fri Apr 01 17:06:10 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")
--- a/src/Pure/System/scala.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/System/scala.scala Fri Apr 01 17:06:10 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)
}
--- a/src/Pure/System/system_channel.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/System/system_channel.scala Fri Apr 01 17:06:10 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
--- a/src/Pure/System/tty_loop.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/System/tty_loop.scala Fri Apr 01 17:06:10 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)
--- a/src/Pure/Thy/bibtex.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Thy/bibtex.scala Fri Apr 01 17:06:10 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"),
@@ -148,11 +141,9 @@
})
}
- 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 @<Scan for and process a \.{.bib} command or database entry@>.
- 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 })
--- a/src/Pure/Thy/document_build.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Thy/document_build.scala Fri Apr 01 17:06:10 2022 +0200
@@ -7,20 +7,17 @@
package isabelle
-object Document_Build
-{
+object Document_Build {
/* document variants */
- abstract class Document_Name
- {
+ abstract class Document_Name {
def name: String
def path: Path = Path.basic(name)
override def toString: String = name
}
- object Document_Variant
- {
+ object Document_Variant {
def parse(opt: String): Document_Variant =
Library.space_explode('=', opt) match {
case List(name) => Document_Variant(name, Latex.Tags.empty)
@@ -29,25 +26,22 @@
}
}
- sealed case class Document_Variant(name: String, tags: Latex.Tags) extends Document_Name
- {
+ sealed case class Document_Variant(name: String, tags: Latex.Tags) extends Document_Name {
def print: String = if (tags.toString.isEmpty) name else name + "=" + tags.toString
}
sealed case class Document_Input(name: String, sources: SHA1.Digest)
- extends Document_Name
+ extends Document_Name
sealed case class Document_Output(name: String, sources: SHA1.Digest, log_xz: Bytes, pdf: Bytes)
- extends Document_Name
- {
+ extends Document_Name {
def log: String = log_xz.uncompress().text
def log_lines: List[String] = split_lines(log)
def write(db: SQL.Database, session_name: String): Unit =
write_document(db, session_name, this)
- def write(dir: Path): Path =
- {
+ def write(dir: Path): Path = {
val path = dir + Path.basic(name).pdf
Isabelle_System.make_directory(path.expand.dir)
Bytes.write(path, pdf)
@@ -58,8 +52,7 @@
/* SQL data model */
- object Data
- {
+ object Data {
val session_name = SQL.Column.string("session_name").make_primary_key
val name = SQL.Column.string("name").make_primary_key
val sources = SQL.Column.string("sources")
@@ -73,23 +66,23 @@
(if (name == "") "" else " AND " + Data.name.equal(name))
}
- def read_documents(db: SQL.Database, session_name: String): List[Document_Input] =
- {
+ def read_documents(db: SQL.Database, session_name: String): List[Document_Input] = {
val select = Data.table.select(List(Data.name, Data.sources), Data.where_equal(session_name))
db.using_statement(select)(stmt =>
- stmt.execute_query().iterator(res =>
- {
+ stmt.execute_query().iterator(res => {
val name = res.string(Data.name)
val sources = res.string(Data.sources)
Document_Input(name, SHA1.fake_digest(sources))
}).toList)
}
- def read_document(db: SQL.Database, session_name: String, name: String): Option[Document_Output] =
- {
+ def read_document(
+ db: SQL.Database,
+ session_name: String,
+ name: String
+ ): Option[Document_Output] = {
val select = Data.table.select(sql = Data.where_equal(session_name, name))
- db.using_statement(select)(stmt =>
- {
+ db.using_statement(select)(stmt => {
val res = stmt.execute_query()
if (res.next()) {
val name = res.string(Data.name)
@@ -102,10 +95,8 @@
})
}
- def write_document(db: SQL.Database, session_name: String, doc: Document_Output): Unit =
- {
- db.using_statement(Data.table.insert())(stmt =>
- {
+ def write_document(db: SQL.Database, session_name: String, doc: Document_Output): Unit = {
+ db.using_statement(Data.table.insert())(stmt => {
stmt.string(1) = session_name
stmt.string(2) = doc.name
stmt.string(3) = doc.sources.toString
@@ -128,8 +119,8 @@
session: String,
deps: Sessions.Deps,
db_context: Sessions.Database_Context,
- progress: Progress = new Progress): Context =
- {
+ progress: Progress = new Progress
+ ): Context = {
val info = deps.sessions_structure(session)
val base = deps(session)
val hierarchy = deps.sessions_structure.build_hierarchy(session)
@@ -141,8 +132,8 @@
base: Sessions.Base,
hierarchy: List[String],
db_context: Sessions.Database_Context,
- val progress: Progress = new Progress)
- {
+ val progress: Progress = new Progress
+ ) {
/* session info */
def session: String = info.name
@@ -159,8 +150,7 @@
def document_build: String = options.string("document_build")
- def get_engine(): Engine =
- {
+ def get_engine(): Engine = {
val name = document_build
engines.find(_.name == name).getOrElse(error("Bad document_build engine " + quote(name)))
}
@@ -184,15 +174,13 @@
File.Content(path, content)
}
- lazy val session_graph: File.Content =
- {
+ lazy val session_graph: File.Content = {
val path = Presentation.session_graph_path
val content = graphview.Graph_File.make_pdf(options, base.session_graph_display)
File.Content(path, content)
}
- lazy val session_tex: File.Content =
- {
+ lazy val session_tex: File.Content = {
val path = Path.basic("session.tex")
val content =
Library.terminate_lines(
@@ -200,11 +188,9 @@
File.Content(path, content)
}
- lazy val isabelle_logo: Option[File.Content] =
- {
+ lazy val isabelle_logo: Option[File.Content] = {
document_logo.map(logo_name =>
- Isabelle_System.with_tmp_file("logo", ext = "pdf")(tmp_path =>
- {
+ Isabelle_System.with_tmp_file("logo", ext = "pdf")(tmp_path => {
Logo.create_logo(logo_name, output_file = tmp_path, quiet = true)
val path = Path.basic("isabelle_logo.pdf")
val content = Bytes.read(tmp_path)
@@ -215,8 +201,11 @@
/* document directory */
- def prepare_directory(dir: Path, doc: Document_Variant, latex_output: Latex.Output): Directory =
- {
+ def prepare_directory(
+ dir: Path,
+ doc: Document_Variant,
+ latex_output: Latex.Output
+ ): Directory = {
val doc_dir = Isabelle_System.make_directory(dir + Path.basic(doc.name))
@@ -270,8 +259,8 @@
doc_dir: Path,
doc: Document_Variant,
root_name: String,
- sources: SHA1.Digest)
- {
+ sources: SHA1.Digest
+ ) {
def root_name_script(ext: String = ""): String =
Bash.string(if (ext.isEmpty) root_name else root_name + "." + ext)
@@ -286,8 +275,7 @@
Latex.latex_errors(doc_dir, root_name) :::
Bibtex.bibtex_errors(doc_dir, root_name)
- def make_document(log: List[String], errors: List[String]): Document_Output =
- {
+ def make_document(log: List[String], errors: List[String]): Document_Output = {
val root_pdf = Path.basic(root_name).pdf
val result_pdf = doc_dir + root_pdf
@@ -312,16 +300,14 @@
lazy val engines: List[Engine] = Isabelle_System.make_services(classOf[Engine])
- abstract class Engine(val name: String) extends Isabelle_System.Service
- {
+ abstract class Engine(val name: String) extends Isabelle_System.Service {
override def toString: String = name
def prepare_directory(context: Context, dir: Path, doc: Document_Variant): Directory
def build_document(context: Context, directory: Directory, verbose: Boolean): Document_Output
}
- abstract class Bash_Engine(name: String) extends Engine(name)
- {
+ abstract class Bash_Engine(name: String) extends Engine(name) {
def prepare_directory(context: Context, dir: Path, doc: Document_Variant): Directory =
context.prepare_directory(dir, doc, new Latex.Output(context.options))
@@ -330,8 +316,7 @@
(if (use_pdflatex) "$ISABELLE_PDFLATEX" else "$ISABELLE_LUALATEX") +
" " + directory.root_name_script() + "\n"
- def bibtex_script(context: Context, directory: Directory, latex: Boolean = false): String =
- {
+ def bibtex_script(context: Context, directory: Directory, latex: Boolean = false): String = {
val ext = if (context.document_bibliography) "aux" else "bib"
directory.conditional_script(ext, "$ISABELLE_BIBTEX",
after = if (latex) latex_script(context, directory) else "")
@@ -342,8 +327,7 @@
after = if (latex) latex_script(context, directory) else "")
def use_build_script: Boolean = false
- def build_script(context: Context, directory: Directory): String =
- {
+ def build_script(context: Context, directory: Directory): String = {
val has_build_script = (directory.doc_dir + Path.explode("build")).is_file
if (!use_build_script && has_build_script) {
@@ -362,8 +346,11 @@
}
}
- def build_document(context: Context, directory: Directory, verbose: Boolean): Document_Output =
- {
+ def build_document(
+ context: Context,
+ directory: Directory,
+ verbose: Boolean
+ ): Document_Output = {
val result =
context.progress.bash(
build_script(context, directory),
@@ -393,16 +380,15 @@
context: Context,
output_sources: Option[Path] = None,
output_pdf: Option[Path] = None,
- verbose: Boolean = false): List[Document_Output] =
- {
+ verbose: Boolean = false
+ ): List[Document_Output] = {
val progress = context.progress
val engine = context.get_engine()
val documents =
for (doc <- context.documents)
yield {
- Isabelle_System.with_tmp_dir("document")(tmp_dir =>
- {
+ Isabelle_System.with_tmp_dir("document")(tmp_dir => {
progress.echo("Preparing " + context.session + "/" + doc.name + " ...")
val start = Time.now()
@@ -435,8 +421,7 @@
/* Isabelle tool wrapper */
val isabelle_tool =
- Isabelle_Tool("document", "prepare session theory document", Scala_Project.here, args =>
- {
+ Isabelle_Tool("document", "prepare session theory document", Scala_Project.here, args => {
var output_sources: Option[Path] = None
var output_pdf: Option[Path] = None
var verbose_latex = false
@@ -496,8 +481,7 @@
progress.echo_warning("No output directory")
}
- using(store.open_database_context())(db_context =>
- {
+ using(store.open_database_context())(db_context => {
build_documents(context(session, deps, db_context, progress = progress),
output_sources = output_sources, output_pdf = output_pdf,
verbose = verbose_latex)
--- a/src/Pure/Thy/export.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Thy/export.scala Fri Apr 01 17:06:10 2022 +0200
@@ -11,8 +11,7 @@
import scala.util.matching.Regex
-object Export
-{
+object Export {
/* artefact names */
val DOCUMENT_ID = "PIDE/document_id"
@@ -31,8 +30,7 @@
/* SQL data model */
- object Data
- {
+ object Data {
val session_name = SQL.Column.string("session_name").make_primary_key
val theory_name = SQL.Column.string("theory_name").make_primary_key
val name = SQL.Column.string("name").make_primary_key
@@ -50,30 +48,31 @@
(if (name == "") "" else " AND " + Data.name.equal(name))
}
- def read_name(db: SQL.Database, session_name: String, theory_name: String, name: String): Boolean =
- {
+ def read_name(
+ db: SQL.Database,
+ session_name: String,
+ theory_name: String,
+ name: String
+ ): Boolean = {
val select =
Data.table.select(List(Data.name), Data.where_equal(session_name, theory_name, name))
db.using_statement(select)(stmt => stmt.execute_query().next())
}
- def read_names(db: SQL.Database, session_name: String, theory_name: String): List[String] =
- {
+ def read_names(db: SQL.Database, session_name: String, theory_name: String): List[String] = {
val select = Data.table.select(List(Data.name), Data.where_equal(session_name, theory_name))
db.using_statement(select)(stmt =>
stmt.execute_query().iterator(res => res.string(Data.name)).toList)
}
- def read_theory_names(db: SQL.Database, session_name: String): List[String] =
- {
+ def read_theory_names(db: SQL.Database, session_name: String): List[String] = {
val select =
Data.table.select(List(Data.theory_name), Data.where_equal(session_name), distinct = true)
db.using_statement(select)(stmt =>
stmt.execute_query().iterator(_.string(Data.theory_name)).toList)
}
- def read_theory_exports(db: SQL.Database, session_name: String): List[(String, String)] =
- {
+ def read_theory_exports(db: SQL.Database, session_name: String): List[(String, String)] = {
val select = Data.table.select(List(Data.theory_name, Data.name), Data.where_equal(session_name))
db.using_statement(select)(stmt =>
stmt.execute_query().iterator(res =>
@@ -95,8 +94,8 @@
name: String,
executable: Boolean,
body: Future[(Boolean, Bytes)],
- cache: XML.Cache)
- {
+ cache: XML.Cache
+ ) {
override def toString: String = name
def compound_name: String = Export.compound_name(theory_name, name)
@@ -109,8 +108,7 @@
def text: String = uncompressed.text
- def uncompressed: Bytes =
- {
+ def uncompressed: Bytes = {
val (compressed, bytes) = body.join
if (compressed) bytes.uncompress(cache = cache.xz) else bytes
}
@@ -118,11 +116,9 @@
def uncompressed_yxml: XML.Body =
YXML.parse_body(UTF8.decode_permissive(uncompressed), cache = cache)
- def write(db: SQL.Database): Unit =
- {
+ def write(db: SQL.Database): Unit = {
val (compressed, bytes) = body.join
- db.using_statement(Data.table.insert())(stmt =>
- {
+ db.using_statement(Data.table.insert())(stmt => {
stmt.string(1) = session_name
stmt.string(2) = theory_name
stmt.string(3) = name
@@ -134,8 +130,7 @@
}
}
- def make_regex(pattern: String): Regex =
- {
+ def make_regex(pattern: String): Regex = {
@tailrec def make(result: List[String], depth: Int, chs: List[Char]): Regex =
chs match {
case '*' :: '*' :: rest => make("[^:]*" :: result, depth, rest)
@@ -152,30 +147,35 @@
make(Nil, 0, pattern.toList)
}
- def make_matcher(pattern: String): (String, String) => Boolean =
- {
+ def make_matcher(pattern: String): (String, String) => Boolean = {
val regex = make_regex(pattern)
(theory_name: String, name: String) =>
regex.pattern.matcher(compound_name(theory_name, name)).matches
}
def make_entry(
- session_name: String, args: Protocol.Export.Args, bytes: Bytes, cache: XML.Cache): Entry =
- {
+ session_name: String,
+ args: Protocol.Export.Args,
+ bytes: Bytes,
+ cache: XML.Cache
+ ): Entry = {
val body =
if (args.compress) Future.fork(bytes.maybe_compress(cache = cache.xz))
else Future.value((false, bytes))
Entry(session_name, args.theory_name, args.name, args.executable, body, cache)
}
- def read_entry(db: SQL.Database, cache: XML.Cache,
- session_name: String, theory_name: String, name: String): Option[Entry] =
- {
+ def read_entry(
+ db: SQL.Database,
+ cache: XML.Cache,
+ session_name: String,
+ theory_name: String,
+ name: String
+ ): Option[Entry] = {
val select =
Data.table.select(List(Data.executable, Data.compressed, Data.body),
Data.where_equal(session_name, theory_name, name))
- db.using_statement(select)(stmt =>
- {
+ db.using_statement(select)(stmt => {
val res = stmt.execute_query()
if (res.next()) {
val executable = res.bool(Data.executable)
@@ -188,9 +188,13 @@
})
}
- def read_entry(dir: Path, cache: XML.Cache,
- session_name: String, theory_name: String, name: String): Option[Entry] =
- {
+ def read_entry(
+ dir: Path,
+ cache: XML.Cache,
+ session_name: String,
+ theory_name: String,
+ name: String
+ ): Option[Entry] = {
val path = dir + Path.basic(theory_name) + Path.explode(name)
if (path.is_file) {
val executable = File.is_executable(path)
@@ -207,16 +211,14 @@
def consumer(db: SQL.Database, cache: XML.Cache, progress: Progress = new Progress): Consumer =
new Consumer(db, cache, progress)
- class Consumer private[Export](db: SQL.Database, cache: XML.Cache, progress: Progress)
- {
+ class Consumer private[Export](db: SQL.Database, cache: XML.Cache, progress: Progress) {
private val errors = Synchronized[List[String]](Nil)
private val consumer =
Consumer_Thread.fork_bulk[(Entry, Boolean)](name = "export")(
bulk = { case (entry, _) => entry.body.is_finished },
consume =
- (args: List[(Entry, Boolean)]) =>
- {
+ (args: List[(Entry, Boolean)]) => {
val results =
db.transaction {
for ((entry, strict) <- args)
@@ -238,15 +240,13 @@
(results, true)
})
- def apply(session_name: String, args: Protocol.Export.Args, body: Bytes): Unit =
- {
+ def apply(session_name: String, args: Protocol.Export.Args, body: Bytes): Unit = {
if (!progress.stopped) {
consumer.send(make_entry(session_name, args, body, cache) -> args.strict)
}
}
- def shutdown(close: Boolean = false): List[String] =
- {
+ def shutdown(close: Boolean = false): List[String] = {
consumer.shutdown()
if (close) db.close()
errors.value.reverse ::: (if (progress.stopped) List("Export stopped") else Nil)
@@ -256,8 +256,7 @@
/* abstract provider */
- object Provider
- {
+ object Provider {
def none: Provider =
new Provider {
def apply(export_name: String): Option[Entry] = None
@@ -279,9 +278,12 @@
override def toString: String = context.toString
}
- def database(db: SQL.Database, cache: XML.Cache, session_name: String, theory_name: String)
- : Provider =
- {
+ def database(
+ db: SQL.Database,
+ cache: XML.Cache,
+ session_name: String,
+ theory_name: String
+ ) : Provider = {
new Provider {
def apply(export_name: String): Option[Entry] =
read_entry(db, cache, session_name, theory_name, export_name)
@@ -311,9 +313,12 @@
override def toString: String = snapshot.toString
}
- def directory(dir: Path, cache: XML.Cache, session_name: String, theory_name: String)
- : Provider =
- {
+ def directory(
+ dir: Path,
+ cache: XML.Cache,
+ session_name: String,
+ theory_name: String
+ ) : Provider = {
new Provider {
def apply(export_name: String): Option[Entry] =
read_entry(dir, cache, session_name, theory_name, export_name)
@@ -327,8 +332,7 @@
}
}
- trait Provider
- {
+ trait Provider {
def apply(export_name: String): Option[Entry]
def uncompressed_yxml(export_name: String): XML.Body =
@@ -350,10 +354,9 @@
progress: Progress = new Progress,
export_prune: Int = 0,
export_list: Boolean = false,
- export_patterns: List[String] = Nil): Unit =
- {
- using(store.open_database(session_name))(db =>
- {
+ export_patterns: List[String] = Nil
+ ): Unit = {
+ using(store.open_database(session_name))(db => {
db.transaction {
val export_names = read_theory_exports(db, session_name)
@@ -400,8 +403,7 @@
val default_export_dir: Path = Path.explode("export")
val isabelle_tool = Isabelle_Tool("export", "retrieve theory exports",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
/* arguments */
var export_dir = default_export_dir
--- a/src/Pure/Thy/export_theory.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Thy/export_theory.scala Fri Apr 01 17:06:10 2022 +0200
@@ -10,12 +10,10 @@
import scala.collection.immutable.SortedMap
-object Export_Theory
-{
+object Export_Theory {
/** session content **/
- sealed case class Session(name: String, theory_graph: Graph[String, Option[Theory]])
- {
+ sealed case class Session(name: String, theory_graph: Graph[String, Option[Theory]]) {
override def toString: String = name
def theory(theory_name: String): Option[Theory] =
@@ -31,12 +29,10 @@
sessions_structure: Sessions.Structure,
session_name: String,
progress: Progress = new Progress,
- cache: Term.Cache = Term.Cache.make()): Session =
- {
+ cache: Term.Cache = Term.Cache.make()): Session = {
val thys =
sessions_structure.build_requirements(List(session_name)).flatMap(session =>
- using(store.open_database(session))(db =>
- {
+ using(store.open_database(session))(db => {
db.transaction {
for (theory <- Export.read_theory_names(db, session))
yield {
@@ -80,8 +76,8 @@
typedefs: List[Typedef],
datatypes: List[Datatype],
spec_rules: List[Spec_Rule],
- others: Map[String, List[Entity[Other]]])
- {
+ others: Map[String, List[Entity[Other]]]
+ ) {
override def toString: String = name
def entity_iterator: Iterator[Entity[No_Content]] =
@@ -113,8 +109,7 @@
(for ((k, xs) <- others.iterator) yield cache.string(k) -> xs.map(_.cache(cache))).toMap)
}
- def read_theory_parents(provider: Export.Provider, theory_name: String): Option[List[String]] =
- {
+ def read_theory_parents(provider: Export.Provider, theory_name: String): Option[List[String]] = {
if (theory_name == Thy_Header.PURE) Some(Nil)
else {
provider(Export.THEORY_PREFIX + "parents")
@@ -125,9 +120,12 @@
def no_theory: Theory =
Theory("", Nil, Nil, Nil, Nil, Nil, Nil, Nil, Nil, Nil, Nil, Nil, Nil, Nil, Nil, Map.empty)
- def read_theory(provider: Export.Provider, session_name: String, theory_name: String,
- cache: Term.Cache = Term.Cache.none): Theory =
- {
+ def read_theory(
+ provider: Export.Provider,
+ session_name: String,
+ theory_name: String,
+ cache: Term.Cache = Term.Cache.none
+ ): Theory = {
val parents =
read_theory_parents(provider, theory_name) getOrElse
error("Missing theory export in session " + quote(session_name) + ": " + quote(theory_name))
@@ -150,13 +148,11 @@
if (cache.no_cache) theory else theory.cache(cache)
}
- def read_pure[A](store: Sessions.Store, read: (Export.Provider, String, String) => A): A =
- {
+ def read_pure[A](store: Sessions.Store, read: (Export.Provider, String, String) => A): A = {
val session_name = Thy_Header.PURE
val theory_name = Thy_Header.PURE
- using(store.open_database(session_name))(db =>
- {
+ using(store.open_database(session_name))(db => {
db.transaction {
val provider = Export.Provider.database(db, store.cache, session_name, theory_name)
read(provider, session_name, theory_name)
@@ -174,8 +170,7 @@
/* entities */
- object Kind
- {
+ object Kind {
val TYPE = "type"
val CONST = "const"
val THM = "thm"
@@ -194,12 +189,10 @@
def export_kind_name(kind: String, name: String): String =
name + "|" + export_kind(kind)
- abstract class Content[T]
- {
+ abstract class Content[T] {
def cache(cache: Term.Cache): T
}
- sealed case class No_Content() extends Content[No_Content]
- {
+ sealed case class No_Content() extends Content[No_Content] {
def cache(cache: Term.Cache): No_Content = this
}
sealed case class Entity[A <: Content[A]](
@@ -209,8 +202,8 @@
pos: Position.T,
id: Option[Long],
serial: Long,
- content: Option[A])
- {
+ content: Option[A]
+ ) {
val kname: String = export_kind_name(kind, name)
val range: Symbol.Range = Position.Range.unapply(pos).getOrElse(Text.Range.offside)
@@ -237,10 +230,9 @@
provider: Export.Provider,
export_name: String,
kind: String,
- decode: XML.Decode.T[A]): List[Entity[A]] =
- {
- def decode_entity(tree: XML.Tree): Entity[A] =
- {
+ decode: XML.Decode.T[A]
+ ): List[Entity[A]] = {
+ def decode_entity(tree: XML.Tree): Entity[A] = {
def err(): Nothing = throw new XML.XML_Body(List(tree))
tree match {
@@ -261,8 +253,7 @@
/* approximative syntax */
- object Assoc extends Enumeration
- {
+ object Assoc extends Enumeration {
val NO_ASSOC, LEFT_ASSOC, RIGHT_ASSOC = Value
}
@@ -284,8 +275,7 @@
/* types */
sealed case class Type(syntax: Syntax, args: List[String], abbrev: Option[Term.Typ])
- extends Content[Type]
- {
+ extends Content[Type] {
override def cache(cache: Term.Cache): Type =
Type(
syntax,
@@ -294,8 +284,8 @@
}
def read_types(provider: Export.Provider): List[Entity[Type]] =
- read_entities(provider, Export.THEORY_PREFIX + "types", Markup.TYPE_NAME, body =>
- {
+ read_entities(provider, Export.THEORY_PREFIX + "types", Markup.TYPE_NAME,
+ body => {
import XML.Decode._
val (syntax, args, abbrev) =
triple(decode_syntax, list(string), option(Term_XML.Decode.typ))(body)
@@ -310,8 +300,8 @@
typargs: List[String],
typ: Term.Typ,
abbrev: Option[Term.Term],
- propositional: Boolean) extends Content[Const]
- {
+ propositional: Boolean
+ ) extends Content[Const] {
override def cache(cache: Term.Cache): Const =
Const(
syntax,
@@ -322,8 +312,8 @@
}
def read_consts(provider: Export.Provider): List[Entity[Const]] =
- read_entities(provider, Export.THEORY_PREFIX + "consts", Markup.CONSTANT, body =>
- {
+ read_entities(provider, Export.THEORY_PREFIX + "consts", Markup.CONSTANT,
+ body => {
import XML.Decode._
val (syntax, (typargs, (typ, (abbrev, propositional)))) =
pair(decode_syntax,
@@ -339,8 +329,8 @@
sealed case class Prop(
typargs: List[(String, Term.Sort)],
args: List[(String, Term.Typ)],
- term: Term.Term) extends Content[Prop]
- {
+ term: Term.Term
+ ) extends Content[Prop] {
override def cache(cache: Term.Cache): Prop =
Prop(
typargs.map({ case (name, sort) => (cache.string(name), cache.sort(sort)) }),
@@ -348,10 +338,8 @@
cache.term(term))
}
- def decode_prop(body: XML.Body): Prop =
- {
- val (typargs, args, t) =
- {
+ def decode_prop(body: XML.Body): Prop = {
+ val (typargs, args, t) = {
import XML.Decode._
import Term_XML.Decode._
triple(list(pair(string, sort)), list(pair(string, typ)), term)(body)
@@ -359,8 +347,7 @@
Prop(typargs, args, t)
}
- sealed case class Axiom(prop: Prop) extends Content[Axiom]
- {
+ sealed case class Axiom(prop: Prop) extends Content[Axiom] {
override def cache(cache: Term.Cache): Axiom = Axiom(prop.cache(cache))
}
@@ -371,16 +358,15 @@
/* theorems */
- sealed case class Thm_Id(serial: Long, theory_name: String)
- {
+ sealed case class Thm_Id(serial: Long, theory_name: String) {
def pure: Boolean = theory_name == Thy_Header.PURE
}
sealed case class Thm(
prop: Prop,
deps: List[String],
- proof: Term.Proof) extends Content[Thm]
- {
+ proof: Term.Proof)
+ extends Content[Thm] {
override def cache(cache: Term.Cache): Thm =
Thm(
prop.cache(cache),
@@ -389,8 +375,8 @@
}
def read_thms(provider: Export.Provider): List[Entity[Thm]] =
- read_entities(provider, Export.THEORY_PREFIX + "thms", Kind.THM, body =>
- {
+ read_entities(provider, Export.THEORY_PREFIX + "thms", Kind.THM,
+ body => {
import XML.Decode._
import Term_XML.Decode._
val (prop, deps, prf) = triple(decode_prop, list(string), proof)(body)
@@ -401,8 +387,8 @@
typargs: List[(String, Term.Sort)],
args: List[(String, Term.Typ)],
term: Term.Term,
- proof: Term.Proof)
- {
+ proof: Term.Proof
+ ) {
def prop: Prop = Prop(typargs, args, term)
def cache(cache: Term.Cache): Proof =
@@ -414,13 +400,14 @@
}
def read_proof(
- provider: Export.Provider, id: Thm_Id, cache: Term.Cache = Term.Cache.none): Option[Proof] =
- {
+ provider: Export.Provider,
+ id: Thm_Id,
+ cache: Term.Cache = Term.Cache.none
+ ): Option[Proof] = {
for { entry <- provider.focus(id.theory_name)(Export.PROOFS_PREFIX + id.serial) }
yield {
val body = entry.uncompressed_yxml
- val (typargs, (args, (prop_body, proof_body))) =
- {
+ val (typargs, (args, (prop_body, proof_body))) = {
import XML.Decode._
import Term_XML.Decode._
pair(list(pair(string, sort)), pair(list(pair(string, typ)), pair(x => x, x => x)))(body)
@@ -439,13 +426,12 @@
provider: Export.Provider,
proof: Term.Proof,
suppress: Thm_Id => Boolean = _ => false,
- cache: Term.Cache = Term.Cache.none): List[(Thm_Id, Proof)] =
- {
+ cache: Term.Cache = Term.Cache.none
+ ): List[(Thm_Id, Proof)] = {
var seen = Set.empty[Long]
var result = SortedMap.empty[Long, (Thm_Id, Proof)]
- def boxes(context: Option[(Long, Term.Proof)], prf: Term.Proof): Unit =
- {
+ def boxes(context: Option[(Long, Term.Proof)], prf: Term.Proof): Unit = {
prf match {
case Term.Abst(_, _, p) => boxes(context, p)
case Term.AbsP(_, _, p) => boxes(context, p)
@@ -482,8 +468,7 @@
/* type classes */
sealed case class Class(params: List[(String, Term.Typ)], axioms: List[Prop])
- extends Content[Class]
- {
+ extends Content[Class] {
override def cache(cache: Term.Cache): Class =
Class(
params.map({ case (name, typ) => (cache.string(name), cache.typ(typ)) }),
@@ -491,8 +476,8 @@
}
def read_classes(provider: Export.Provider): List[Entity[Class]] =
- read_entities(provider, Export.THEORY_PREFIX + "classes", Markup.CLASS, body =>
- {
+ read_entities(provider, Export.THEORY_PREFIX + "classes", Markup.CLASS,
+ body => {
import XML.Decode._
import Term_XML.Decode._
val (params, axioms) = pair(list(pair(string, typ)), list(decode_prop))(body)
@@ -505,8 +490,8 @@
sealed case class Locale(
typargs: List[(String, Term.Sort)],
args: List[((String, Term.Typ), Syntax)],
- axioms: List[Prop]) extends Content[Locale]
- {
+ axioms: List[Prop]
+ ) extends Content[Locale] {
override def cache(cache: Term.Cache): Locale =
Locale(
typargs.map({ case (name, sort) => (cache.string(name), cache.sort(sort)) }),
@@ -515,8 +500,8 @@
}
def read_locales(provider: Export.Provider): List[Entity[Locale]] =
- read_entities(provider, Export.THEORY_PREFIX + "locales", Markup.LOCALE, body =>
- {
+ read_entities(provider, Export.THEORY_PREFIX + "locales", Markup.LOCALE,
+ body => {
import XML.Decode._
import Term_XML.Decode._
val (typargs, args, axioms) =
@@ -533,8 +518,8 @@
target: String,
prefix: List[(String, Boolean)],
subst_types: List[((String, Term.Sort), Term.Typ)],
- subst_terms: List[((String, Term.Typ), Term.Term)]) extends Content[Locale_Dependency]
- {
+ subst_terms: List[((String, Term.Typ), Term.Term)]
+ ) extends Content[Locale_Dependency] {
override def cache(cache: Term.Cache): Locale_Dependency =
Locale_Dependency(
cache.string(source),
@@ -549,48 +534,46 @@
def read_locale_dependencies(provider: Export.Provider): List[Entity[Locale_Dependency]] =
read_entities(provider, Export.THEORY_PREFIX + "locale_dependencies", Kind.LOCALE_DEPENDENCY,
- body =>
- {
- import XML.Decode._
- import Term_XML.Decode._
- val (source, (target, (prefix, (subst_types, subst_terms)))) =
- pair(string, pair(string, pair(list(pair(string, bool)),
- pair(list(pair(pair(string, sort), typ)), list(pair(pair(string, typ), term))))))(body)
- Locale_Dependency(source, target, prefix, subst_types, subst_terms)
- })
+ body => {
+ import XML.Decode._
+ import Term_XML.Decode._
+ val (source, (target, (prefix, (subst_types, subst_terms)))) =
+ pair(string, pair(string, pair(list(pair(string, bool)),
+ pair(list(pair(pair(string, sort), typ)), list(pair(pair(string, typ), term))))))(body)
+ Locale_Dependency(source, target, prefix, subst_types, subst_terms)
+ })
/* sort algebra */
- sealed case class Classrel(class1: String, class2: String, prop: Prop)
- {
+ sealed case class Classrel(class1: String, class2: String, prop: Prop) {
def cache(cache: Term.Cache): Classrel =
Classrel(cache.string(class1), cache.string(class2), prop.cache(cache))
}
- def read_classrel(provider: Export.Provider): List[Classrel] =
- {
+ def read_classrel(provider: Export.Provider): List[Classrel] = {
val body = provider.uncompressed_yxml(Export.THEORY_PREFIX + "classrel")
- val classrel =
- {
+ val classrel = {
import XML.Decode._
list(pair(decode_prop, pair(string, string)))(body)
}
for ((prop, (c1, c2)) <- classrel) yield Classrel(c1, c2, prop)
}
- sealed case class Arity(type_name: String, domain: List[Term.Sort], codomain: String, prop: Prop)
- {
+ sealed case class Arity(
+ type_name: String,
+ domain: List[Term.Sort],
+ codomain: String,
+ prop: Prop
+ ) {
def cache(cache: Term.Cache): Arity =
Arity(cache.string(type_name), domain.map(cache.sort), cache.string(codomain),
prop.cache(cache))
}
- def read_arities(provider: Export.Provider): List[Arity] =
- {
+ def read_arities(provider: Export.Provider): List[Arity] = {
val body = provider.uncompressed_yxml(Export.THEORY_PREFIX + "arities")
- val arities =
- {
+ val arities = {
import XML.Decode._
import Term_XML.Decode._
list(pair(decode_prop, triple(string, list(sort), string)))(body)
@@ -601,17 +584,14 @@
/* Pure constdefs */
- sealed case class Constdef(name: String, axiom_name: String)
- {
+ sealed case class Constdef(name: String, axiom_name: String) {
def cache(cache: Term.Cache): Constdef =
Constdef(cache.string(name), cache.string(axiom_name))
}
- def read_constdefs(provider: Export.Provider): List[Constdef] =
- {
+ def read_constdefs(provider: Export.Provider): List[Constdef] = {
val body = provider.uncompressed_yxml(Export.THEORY_PREFIX + "constdefs")
- val constdefs =
- {
+ val constdefs = {
import XML.Decode._
list(pair(string, string))(body)
}
@@ -621,9 +601,14 @@
/* HOL typedefs */
- sealed case class Typedef(name: String,
- rep_type: Term.Typ, abs_type: Term.Typ, rep_name: String, abs_name: String, axiom_name: String)
- {
+ sealed case class Typedef(
+ name: String,
+ rep_type: Term.Typ,
+ abs_type: Term.Typ,
+ rep_name: String,
+ abs_name: String,
+ axiom_name: String
+ ) {
def cache(cache: Term.Cache): Typedef =
Typedef(cache.string(name),
cache.typ(rep_type),
@@ -633,11 +618,9 @@
cache.string(axiom_name))
}
- def read_typedefs(provider: Export.Provider): List[Typedef] =
- {
+ def read_typedefs(provider: Export.Provider): List[Typedef] = {
val body = provider.uncompressed_yxml(Export.THEORY_PREFIX + "typedefs")
- val typedefs =
- {
+ val typedefs = {
import XML.Decode._
import Term_XML.Decode._
list(pair(string, pair(typ, pair(typ, pair(string, pair(string, string))))))(body)
@@ -655,8 +638,8 @@
co: Boolean,
typargs: List[(String, Term.Sort)],
typ: Term.Typ,
- constructors: List[(Term.Term, Term.Typ)])
- {
+ constructors: List[(Term.Term, Term.Typ)]
+ ) {
def id: Option[Long] = Position.Id.unapply(pos)
def cache(cache: Term.Cache): Datatype =
@@ -669,11 +652,9 @@
constructors.map({ case (term, typ) => (cache.term(term), cache.typ(typ)) }))
}
- def read_datatypes(provider: Export.Provider): List[Datatype] =
- {
+ def read_datatypes(provider: Export.Provider): List[Datatype] = {
val body = provider.uncompressed_yxml(Export.THEORY_PREFIX + "datatypes")
- val datatypes =
- {
+ val datatypes = {
import XML.Decode._
import Term_XML.Decode._
list(pair(properties, pair(string, pair(bool, pair(list(pair(string, sort)),
@@ -686,8 +667,7 @@
/* Pure spec rules */
- sealed abstract class Recursion
- {
+ sealed abstract class Recursion {
def cache(cache: Term.Cache): Recursion =
this match {
case Primrec(types) => Primrec(types.map(cache.string))
@@ -701,8 +681,7 @@
case object Corec extends Recursion
case object Unknown_Recursion extends Recursion
- val decode_recursion: XML.Decode.T[Recursion] =
- {
+ val decode_recursion: XML.Decode.T[Recursion] = {
import XML.Decode._
variant(List(
{ case (Nil, a) => Primrec(list(string)(a)) },
@@ -713,8 +692,7 @@
}
- sealed abstract class Rough_Classification
- {
+ sealed abstract class Rough_Classification {
def is_equational: Boolean = this.isInstanceOf[Equational]
def is_inductive: Boolean = this == Inductive
def is_co_inductive: Boolean = this == Co_Inductive
@@ -732,8 +710,7 @@
case object Co_Inductive extends Rough_Classification
case object Unknown extends Rough_Classification
- val decode_rough_classification: XML.Decode.T[Rough_Classification] =
- {
+ val decode_rough_classification: XML.Decode.T[Rough_Classification] = {
import XML.Decode._
variant(List(
{ case (Nil, a) => Equational(decode_recursion(a)) },
@@ -750,8 +727,8 @@
typargs: List[(String, Term.Sort)],
args: List[(String, Term.Typ)],
terms: List[(Term.Term, Term.Typ)],
- rules: List[Term.Term])
- {
+ rules: List[Term.Term]
+ ) {
def id: Option[Long] = Position.Id.unapply(pos)
def cache(cache: Term.Cache): Spec_Rule =
@@ -765,11 +742,9 @@
rules.map(cache.term))
}
- def read_spec_rules(provider: Export.Provider): List[Spec_Rule] =
- {
+ def read_spec_rules(provider: Export.Provider): List[Spec_Rule] = {
val body = provider.uncompressed_yxml(Export.THEORY_PREFIX + "spec_rules")
- val spec_rules =
- {
+ val spec_rules = {
import XML.Decode._
import Term_XML.Decode._
list(
@@ -784,13 +759,11 @@
/* other entities */
- sealed case class Other() extends Content[Other]
- {
+ sealed case class Other() extends Content[Other] {
override def cache(cache: Term.Cache): Other = this
}
- def read_others(provider: Export.Provider): Map[String, List[Entity[Other]]] =
- {
+ def read_others(provider: Export.Provider): Map[String, List[Entity[Other]]] = {
val kinds =
provider(Export.THEORY_PREFIX + "other_kinds") match {
case Some(entry) => split_lines(entry.uncompressed.text)
--- a/src/Pure/Thy/file_format.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Thy/file_format.scala Fri Apr 01 17:06:10 2022 +0200
@@ -7,8 +7,7 @@
package isabelle
-object File_Format
-{
+object File_Format {
sealed case class Theory_Context(name: Document.Node.Name, content: String)
@@ -17,8 +16,7 @@
lazy val registry: Registry =
new Registry(Isabelle_System.make_services(classOf[File_Format]))
- final class Registry private [File_Format](file_formats: List[File_Format])
- {
+ final class Registry private [File_Format](file_formats: List[File_Format]) {
override def toString: String = file_formats.mkString("File_Format.Environment(", ",", ")")
def get(name: String): Option[File_Format] = file_formats.find(_.detect(name))
@@ -33,8 +31,7 @@
/* session */
- final class Session private[File_Format](agents: List[Agent])
- {
+ final class Session private[File_Format](agents: List[Agent]) {
override def toString: String =
agents.mkString("File_Format.Session(", ", ", ")")
@@ -44,8 +41,7 @@
def stop_session: Unit = agents.foreach(_.stop())
}
- trait Agent
- {
+ trait Agent {
def prover_options(options: Options): Options = options
def stop(): Unit = {}
}
@@ -53,8 +49,7 @@
object No_Agent extends Agent
}
-abstract class File_Format extends Isabelle_System.Service
-{
+abstract class File_Format extends Isabelle_System.Service {
def format_name: String
override def toString: String = "File_Format(" + format_name + ")"
@@ -67,8 +62,10 @@
def theory_suffix: String = ""
def theory_content(name: String): String = ""
- def make_theory_name(resources: Resources, name: Document.Node.Name): Option[Document.Node.Name] =
- {
+ def make_theory_name(
+ resources: Resources,
+ name: Document.Node.Name
+ ): Option[Document.Node.Name] = {
for {
(_, thy) <- Thy_Header.split_file_name(name.node)
if detect(name.node) && theory_suffix.nonEmpty
@@ -79,8 +76,7 @@
}
}
- def make_theory_content(resources: Resources, thy_name: Document.Node.Name): Option[String] =
- {
+ def make_theory_content(resources: Resources, thy_name: Document.Node.Name): Option[String] = {
for {
(prefix, suffix) <- Thy_Header.split_file_name(thy_name.node)
if detect(prefix) && suffix == theory_suffix
--- a/src/Pure/Thy/html.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Thy/html.scala Fri Apr 01 17:06:10 2022 +0200
@@ -7,12 +7,10 @@
package isabelle
-object HTML
-{
+object HTML {
/* attributes */
- class Attribute(val name: String, value: String)
- {
+ class Attribute(val name: String, value: String) {
def xml: XML.Attribute = name -> value
def apply(elem: XML.Elem): XML.Elem = elem + xml
}
@@ -34,15 +32,13 @@
val break: XML.Body = List(XML.elem("br"))
val nl: XML.Body = List(XML.Text("\n"))
- class Operator(val name: String)
- {
+ class Operator(val name: String) {
def apply(body: XML.Body): XML.Elem = XML.elem(name, body)
def apply(att: Attribute, body: XML.Body): XML.Elem = att(apply(body))
def apply(c: String, body: XML.Body): XML.Elem = apply(class_(c), body)
}
- class Heading(name: String) extends Operator(name)
- {
+ class Heading(name: String) extends Operator(name) {
def apply(txt: String): XML.Elem = super.apply(text(txt))
def apply(att: Attribute, txt: String): XML.Elem = super.apply(att, text(txt))
def apply(c: String, txt: String): XML.Elem = super.apply(c, text(txt))
@@ -120,16 +116,14 @@
def is_control_block_begin(sym: Symbol.Symbol): Boolean = control_block_begin.isDefinedAt(sym)
def is_control_block_end(sym: Symbol.Symbol): Boolean = control_block_end.isDefinedAt(sym)
- def is_control_block_pair(bg: Symbol.Symbol, en: Symbol.Symbol): Boolean =
- {
+ def is_control_block_pair(bg: Symbol.Symbol, en: Symbol.Symbol): Boolean = {
val bg_decoded = Symbol.decode(bg)
val en_decoded = Symbol.decode(en)
bg_decoded == Symbol.bsub_decoded && en_decoded == Symbol.esub_decoded ||
bg_decoded == Symbol.bsup_decoded && en_decoded == Symbol.esup_decoded
}
- def check_control_blocks(body: XML.Body): Boolean =
- {
+ def check_control_blocks(body: XML.Body): Boolean = {
var ok = true
var open = List.empty[Symbol.Symbol]
for { XML.Text(text) <- body; sym <- Symbol.iterator(text) } {
@@ -144,9 +138,13 @@
ok && open.isEmpty
}
- def output(s: StringBuilder, text: String,
- control_blocks: Boolean, hidden: Boolean, permissive: Boolean): Unit =
- {
+ def output(
+ s: StringBuilder,
+ text: String,
+ control_blocks: Boolean,
+ hidden: Boolean,
+ permissive: Boolean
+ ): Unit = {
def output_string(str: String): Unit =
XML.output_string(s, str, permissive = permissive)
@@ -197,8 +195,7 @@
output_symbol(ctrl)
}
- def output(text: String): String =
- {
+ def output(text: String): String = {
val control_blocks = check_control_blocks(List(XML.Text(text)))
Library.make_string(output(_, text,
control_blocks = control_blocks, hidden = false, permissive = true))
@@ -211,10 +208,8 @@
Set("head", "body", "meta", "div", "pre", "p", "title", "h1", "h2", "h3", "h4", "h5", "h6",
"ul", "ol", "dl", "li", "dt", "dd")
- def output(s: StringBuilder, xml: XML.Body, hidden: Boolean, structural: Boolean): Unit =
- {
- def output_body(body: XML.Body): Unit =
- {
+ def output(s: StringBuilder, xml: XML.Body, hidden: Boolean, structural: Boolean): Unit = {
+ def output_body(body: XML.Body): Unit = {
val control_blocks = check_control_blocks(body)
body foreach {
case XML.Elem(markup, Nil) =>
@@ -269,8 +264,7 @@
/* GUI elements */
- object GUI
- {
+ object GUI {
private def optional_value(text: String): XML.Attributes =
proper_string(text).map(a => "value" -> a).toList
@@ -333,16 +327,17 @@
/* GUI layout */
- object Wrap_Panel
- {
- object Alignment extends Enumeration
- {
+ object Wrap_Panel {
+ object Alignment extends Enumeration {
val left, right, center = Value
}
- def apply(contents: List[XML.Elem], name: String = "", action: String = "",
- alignment: Alignment.Value = Alignment.right): XML.Elem =
- {
+ def apply(
+ contents: List[XML.Elem],
+ name: String = "",
+ action: String = "",
+ alignment: Alignment.Value = Alignment.right
+ ): XML.Elem = {
val body = Library.separate(XML.Text(" "), contents)
GUI.form(List(div(body) + ("style" -> ("text-align: " + alignment))),
name = name, action = action)
@@ -363,11 +358,13 @@
XML.Elem(Markup("meta",
List("http-equiv" -> "Content-Type", "content" -> "text/html; charset=utf-8")), Nil)
- def output_document(head: XML.Body, body: XML.Body,
+ def output_document(
+ head: XML.Body,
+ body: XML.Body,
css: String = "isabelle.css",
hidden: Boolean = true,
- structural: Boolean = true): String =
- {
+ structural: Boolean = true
+ ): String = {
cat_lines(
List(
header,
@@ -384,8 +381,7 @@
val fonts_path: Path = Path.explode("fonts")
- def init_fonts(dir: Path): Unit =
- {
+ def init_fonts(dir: Path): Unit = {
val fonts_dir = Isabelle_System.make_directory(dir + HTML.fonts_path)
for (entry <- Isabelle_Fonts.fonts(hidden = true))
Isabelle_System.copy_file(entry.path, fonts_dir)
@@ -397,8 +393,7 @@
(for (entry <- Isabelle_Fonts.fonts(hidden = true))
yield (entry.path.file_name -> Url.print_file(entry.path.file))).toMap
- def fonts_css(make_url: String => String = fonts_url()): String =
- {
+ def fonts_css(make_url: String => String = fonts_url()): String = {
def font_face(entry: Isabelle_Fonts.Entry): String =
cat_lines(
List(
@@ -413,8 +408,7 @@
.mkString("", "\n\n", "\n")
}
- def fonts_css_dir(prefix: String = ""): String =
- {
+ def fonts_css_dir(prefix: String = ""): String = {
val prefix1 = if (prefix.isEmpty || prefix.endsWith("/")) prefix else prefix + "/"
fonts_css(fonts_dir(prefix1 + fonts_path.implode))
}
@@ -436,8 +430,8 @@
base: Option[Path] = None,
css: String = isabelle_css.file_name,
hidden: Boolean = true,
- structural: Boolean = true): Unit =
- {
+ structural: Boolean = true
+ ): Unit = {
Isabelle_System.make_directory(dir)
val prefix = relative_prefix(dir, base)
File.write(dir + isabelle_css.base, fonts_css_dir(prefix) + "\n\n" + File.read(isabelle_css))
--- a/src/Pure/Thy/latex.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Thy/latex.scala Fri Apr 01 17:06:10 2022 +0200
@@ -15,8 +15,7 @@
import scala.util.matching.Regex
-object Latex
-{
+object Latex {
/* output name for LaTeX macros */
private val output_name_map: Map[Char, String] =
@@ -49,8 +48,7 @@
/* index entries */
- def index_escape(str: String): String =
- {
+ def index_escape(str: String): String = {
val special1 = "!\"@|"
val special2 = "\\{}#"
if (str.exists(c => special1.contains(c) || special2.contains(c))) {
@@ -70,8 +68,7 @@
else str
}
- object Index_Item
- {
+ object Index_Item {
sealed case class Value(text: Text, like: String)
def unapply(tree: XML.Tree): Option[Value] =
tree match {
@@ -81,8 +78,7 @@
}
}
- object Index_Entry
- {
+ object Index_Entry {
sealed case class Value(items: List[Index_Item.Value], kind: String)
def unapply(tree: XML.Tree): Option[Value] =
tree match {
@@ -96,10 +92,8 @@
/* tags */
- object Tags
- {
- object Op extends Enumeration
- {
+ object Tags {
+ object Op extends Enumeration {
val fold, drop, keep = Value
}
@@ -123,14 +117,12 @@
val empty: Tags = apply("")
}
- class Tags private(spec: String, map: TreeMap[String, Tags.Op.Value])
- {
+ class Tags private(spec: String, map: TreeMap[String, Tags.Op.Value]) {
override def toString: String = spec
def get(name: String): Option[Tags.Op.Value] = map.get(name)
- def sty(comment_latex: Boolean): File.Content =
- {
+ def sty(comment_latex: Boolean): File.Content = {
val path = Path.explode("isabelletags.sty")
val comment =
if (comment_latex) """\usepackage{comment}"""
@@ -165,8 +157,7 @@
if (file_pos.isEmpty) Nil
else List("\\endinput\n", position(Markup.FILE, file_pos))
- class Output(options: Options)
- {
+ class Output(options: Options) {
def latex_output(latex_text: Text): String = apply(latex_text)
def latex_macro0(name: String, optional_argument: String = ""): Text =
@@ -188,8 +179,7 @@
def latex_body(kind: String, body: Text, optional_argument: String = ""): Text =
latex_environment("isamarkup" + kind, body, optional_argument)
- def latex_tag(name: String, body: Text, delim: Boolean = false): Text =
- {
+ def latex_tag(name: String, body: Text, delim: Boolean = false): Text = {
val s = output_name(name)
val kind = if (delim) "delim" else "tag"
val end = if (delim) "" else "{\\isafold" + s + "}%\n"
@@ -205,15 +195,13 @@
}
}
- def index_item(item: Index_Item.Value): String =
- {
+ def index_item(item: Index_Item.Value): String = {
val like = if (item.like.isEmpty) "" else index_escape(item.like) + "@"
val text = index_escape(latex_output(item.text))
like + text
}
- def index_entry(entry: Index_Entry.Value): Text =
- {
+ def index_entry(entry: Index_Entry.Value): Text = {
val items = entry.items.map(index_item).mkString("!")
val kind = if (entry.kind.isEmpty) "" else "|" + index_escape(entry.kind)
latex_macro("index", XML.string(items + kind))
@@ -226,16 +214,14 @@
error("Unknown latex markup element " + quote(elem.name) + Position.here(pos) +
":\n" + XML.string_of_tree(elem))
- def apply(latex_text: Text, file_pos: String = ""): String =
- {
+ def apply(latex_text: Text, file_pos: String = ""): String = {
var line = 1
val result = new mutable.ListBuffer[String]
val positions = new mutable.ListBuffer[String] ++= init_position(file_pos)
val file_position = if (file_pos.isEmpty) Position.none else Position.File(file_pos)
- def traverse(xml: XML.Body): Unit =
- {
+ def traverse(xml: XML.Body): Unit = {
xml.foreach {
case XML.Text(s) =>
line += s.count(_ == '\n')
@@ -281,8 +267,7 @@
private val File_Pattern = """^%:%file=(.+)%:%$""".r
private val Line_Pattern = """^*%:%(\d+)=(\d+)%:%$""".r
- def read_tex_file(tex_file: Path): Tex_File =
- {
+ def read_tex_file(tex_file: Path): Tex_File = {
val positions =
Line.logical_lines(File.read(tex_file)).reverse.
takeWhile(_ != "\\endinput").reverse
@@ -306,8 +291,10 @@
}
final class Tex_File private[Latex](
- tex_file: Path, source_file: Option[String], source_lines: List[(Int, Int)])
- {
+ tex_file: Path,
+ source_file: Option[String],
+ source_lines: List[(Int, Int)]
+ ) {
override def toString: String = tex_file.toString
def source_position(l: Int): Option[Position.T] =
@@ -327,8 +314,7 @@
/* latex log */
- def latex_errors(dir: Path, root_name: String): List[String] =
- {
+ def latex_errors(dir: Path, root_name: String): List[String] = {
val root_log_path = dir + Path.explode(root_name).ext("log")
if (root_log_path.is_file) {
for { (msg, pos) <- filter_errors(dir, File.read(root_log_path)) }
@@ -337,8 +323,7 @@
else Nil
}
- def filter_errors(dir: Path, root_log: String): List[(String, Position.T)] =
- {
+ def filter_errors(dir: Path, root_log: String): List[(String, Position.T)] = {
val seen_files = Synchronized(Map.empty[JFile, Tex_File])
def check_tex_file(path: Path): Option[Tex_File] =
@@ -359,8 +344,7 @@
case None => Position.Line_File(line, path.implode)
}
- object File_Line_Error
- {
+ object File_Line_Error {
val Pattern: Regex = """^(.*?\.\w\w\w):(\d+): (.*)$""".r
def unapply(line: String): Option[(Path, Int, String)] =
line match {
@@ -404,8 +388,7 @@
"See the LaTeX manual or LaTeX Companion for explanation.",
"Type H <return> for immediate help.")
- def error_suffix1(lines: List[String]): Option[String] =
- {
+ def error_suffix1(lines: List[String]): Option[String] = {
val lines1 =
lines.take(20).takeWhile({ case File_Line_Error((_, _, _)) => false case _ => true })
lines1.zipWithIndex.collectFirst({
@@ -418,9 +401,10 @@
case _ => None
}
- @tailrec def filter(lines: List[String], result: List[(String, Position.T)])
- : List[(String, Position.T)] =
- {
+ @tailrec def filter(
+ lines: List[String],
+ result: List[(String, Position.T)]
+ ) : List[(String, Position.T)] = {
lines match {
case File_Line_Error((file, line, msg1)) :: rest1 =>
val pos = tex_file_position(file, line)
--- a/src/Pure/Thy/presentation.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Thy/presentation.scala Fri Apr 01 17:06:10 2022 +0200
@@ -12,16 +12,14 @@
import scala.collection.mutable
-object Presentation
-{
+object Presentation {
/** HTML documents **/
/* HTML context */
sealed case class HTML_Document(title: String, content: String)
- abstract class HTML_Context
- {
+ abstract class HTML_Context {
/* directory structure and resources */
def root_dir: Path
@@ -47,17 +45,18 @@
def source(body: XML.Body): XML.Tree = HTML.pre("source", body)
- def contents(heading: String, items: List[XML.Body], css_class: String = "contents")
- : List[XML.Elem] =
- {
+ def contents(
+ heading: String,
+ items: List[XML.Body],
+ css_class: String = "contents"
+ ) : List[XML.Elem] = {
if (items.isEmpty) Nil
else List(HTML.div(css_class, List(HTML.section(heading), HTML.itemize(items))))
}
val isabelle_css: String = File.read(HTML.isabelle_css)
- def html_document(title: String, body: XML.Body, fonts_css: String): HTML_Document =
- {
+ def html_document(title: String, body: XML.Body, fonts_css: String): HTML_Document = {
val content =
HTML.output_document(
List(
@@ -93,8 +92,7 @@
type Entity = Export_Theory.Entity[Export_Theory.No_Content]
- object Entity_Context
- {
+ object Entity_Context {
sealed case class Theory_Export(
entity_by_range: Map[Symbol.Range, List[Export_Theory.Entity[Export_Theory.No_Content]]],
entity_by_kind_name: Map[(String, String), Export_Theory.Entity[Export_Theory.No_Content]],
@@ -102,8 +100,7 @@
val no_theory_export: Theory_Export = Theory_Export(Map.empty, Map.empty, Nil)
- object Theory_Ref
- {
+ object Theory_Ref {
def unapply(props: Properties.T): Option[Document.Node.Name] =
(props, props, props) match {
case (Markup.Kind(Markup.THEORY), Markup.Name(theory), Position.Def_File(thy_file)) =>
@@ -112,8 +109,7 @@
}
}
- object Entity_Ref
- {
+ object Entity_Ref {
def unapply(props: Properties.T): Option[(Path, Option[String], String, String)] =
(props, props, props, props) match {
case (Markup.Entity.Ref.Prop(_), Position.Def_File(def_file),
@@ -134,8 +130,7 @@
new Entity_Context {
private val seen_ranges: mutable.Set[Symbol.Range] = mutable.Set.empty
- override def make_def(range: Symbol.Range, body: XML.Body): Option[XML.Elem] =
- {
+ override def make_def(range: Symbol.Range, body: XML.Body): Option[XML.Elem] = {
body match {
case List(XML.Elem(Markup("span", List("id" -> _)), _)) => None
case _ =>
@@ -159,8 +154,7 @@
private def offset_id(range: Text.Range): String =
"offset_" + range.start + ".." + range.stop
- private def physical_ref(thy_name: String, props: Properties.T): Option[String] =
- {
+ private def physical_ref(thy_name: String, props: Properties.T): Option[String] = {
for {
range <- Position.Def_Range.unapply(props)
if thy_name == node.theory
@@ -176,8 +170,7 @@
entity <- thy.entity_by_kind_name.get((kind, name))
} yield entity.kname
- override def make_ref(props: Properties.T, body: XML.Body): Option[XML.Elem] =
- {
+ override def make_ref(props: Properties.T, body: XML.Body): Option[XML.Elem] = {
props match {
case Theory_Ref(node_name) =>
node_relative(deps, session, node_name).map(html_dir =>
@@ -200,8 +193,7 @@
}
}
- class Entity_Context
- {
+ class Entity_Context {
def make_def(range: Symbol.Range, body: XML.Body): Option[XML.Elem] = None
def make_ref(props: Properties.T, body: XML.Body): Option[XML.Elem] = None
}
@@ -216,8 +208,8 @@
def make_html(
entity_context: Entity_Context,
elements: Elements,
- xml: XML.Body): XML.Body =
- {
+ xml: XML.Body
+ ): XML.Body = {
def html_div(html: XML.Body): Boolean =
html exists {
case XML.Elem(markup, body) => div_elements.contains(markup.name) || html_div(body)
@@ -297,8 +289,8 @@
html_context: HTML_Context,
elements: Elements,
plain_text: Boolean = false,
- fonts_css: String = HTML.fonts_css()): HTML_Document =
- {
+ fonts_css: String = HTML.fonts_css()
+ ): HTML_Document = {
require(!snapshot.is_outdated, "document snapshot outdated")
val name = snapshot.node_name
@@ -325,8 +317,7 @@
/* presentation context */
- object Context
- {
+ object Context {
val none: Context = new Context { def enabled: Boolean = false }
val standard: Context = new Context { def enabled: Boolean = true }
@@ -340,8 +331,7 @@
if (s == ":") standard else dir(Path.explode(s))
}
- abstract class Context private
- {
+ abstract class Context private {
def enabled: Boolean
def enabled(info: Sessions.Info): Boolean = enabled || info.browser_info
def dir(store: Sessions.Store): Path = store.presentation_dir
@@ -354,8 +344,7 @@
private val sessions_path = Path.basic(".sessions")
- private def read_sessions(dir: Path): List[(String, String)] =
- {
+ private def read_sessions(dir: Path): List[(String, String)] = {
val path = dir + sessions_path
if (path.is_file) {
import XML.Decode._
@@ -365,8 +354,10 @@
}
def update_chapter(
- presentation_dir: Path, chapter: String, new_sessions: List[(String, String)]): Unit =
- {
+ presentation_dir: Path,
+ chapter: String,
+ new_sessions: List[(String, String)]
+ ): Unit = {
val dir = Isabelle_System.make_directory(presentation_dir + Path.basic(chapter))
val sessions0 =
@@ -396,8 +387,7 @@
base = Some(presentation_dir))
}
- def update_root(presentation_dir: Path): Unit =
- {
+ def update_root(presentation_dir: Path): Unit = {
Isabelle_System.make_directory(presentation_dir)
HTML.init_fonts(presentation_dir)
Isabelle_System.copy_file(Path.explode("~~/lib/logo/isabelle.gif"),
@@ -446,8 +436,11 @@
private def node_file(name: Document.Node.Name): String =
if (name.node.endsWith(".thy")) html_name(name) else files_path(name.path)
- private def session_relative(deps: Sessions.Deps, session0: String, session1: String): Option[String] =
- {
+ private def session_relative(
+ deps: Sessions.Deps,
+ session0: String,
+ session1: String
+ ): Option[String] = {
for {
info0 <- deps.sessions_structure.get(session0)
info1 <- deps.sessions_structure.get(session1)
@@ -457,15 +450,19 @@
def node_relative(
deps: Sessions.Deps,
session0: String,
- node_name: Document.Node.Name): Option[String] =
- {
+ node_name: Document.Node.Name
+ ): Option[String] = {
val session1 = deps(session0).theory_qualifier(node_name)
session_relative(deps, session0, session1)
}
- def theory_link(deps: Sessions.Deps, session0: String,
- name: Document.Node.Name, body: XML.Body, anchor: Option[String] = None): Option[XML.Tree] =
- {
+ def theory_link(
+ deps: Sessions.Deps,
+ session0: String,
+ name: Document.Node.Name,
+ body: XML.Body,
+ anchor: Option[String] = None
+ ): Option[XML.Tree] = {
val session1 = deps(session0).theory_qualifier(name)
val info0 = deps.sessions_structure.get(session0)
val info1 = deps.sessions_structure.get(session1)
@@ -479,8 +476,8 @@
def read_exports(
sessions: List[String],
deps: Sessions.Deps,
- db_context: Sessions.Database_Context): Map[String, Entity_Context.Theory_Export] =
- {
+ db_context: Sessions.Database_Context
+ ): Map[String, Entity_Context.Theory_Export] = {
type Batch = (String, List[String])
val batches =
sessions.foldLeft((Set.empty[String], List.empty[Batch]))(
@@ -517,8 +514,8 @@
progress: Progress = new Progress,
verbose: Boolean = false,
html_context: HTML_Context,
- session_elements: Elements): Unit =
- {
+ session_elements: Elements
+ ): Unit = {
val info = deps.sessions_structure(session)
val options = info.options
val base = deps(session)
@@ -543,8 +540,7 @@
doc
}
- val view_links =
- {
+ val view_links = {
val deps_link =
HTML.link(session_graph_path, HTML.text("theory dependencies"))
@@ -568,24 +564,23 @@
sealed case class Seen_File(
- src_path: Path, thy_name: Document.Node.Name, thy_session: String)
- {
+ src_path: Path,
+ thy_name: Document.Node.Name,
+ thy_session: String
+ ) {
val files_path: Path = html_context.files_path(thy_name, src_path)
- def check(src_path1: Path, thy_name1: Document.Node.Name, thy_session1: String): Boolean =
- {
+ def check(src_path1: Path, thy_name1: Document.Node.Name, thy_session1: String): Boolean = {
val files_path1 = html_context.files_path(thy_name1, src_path1)
(src_path == src_path1 || files_path == files_path1) && thy_session == thy_session1
}
}
var seen_files = List.empty[Seen_File]
- def present_theory(name: Document.Node.Name): Option[XML.Body] =
- {
+ def present_theory(name: Document.Node.Name): Option[XML.Body] = {
progress.expose_interrupt()
- Build_Job.read_theory(db_context, hierarchy, name.theory).flatMap(command =>
- {
+ Build_Job.read_theory(db_context, hierarchy, name.theory).flatMap(command => {
if (verbose) progress.echo("Presenting theory " + name)
val snapshot = Document.State.init.snippet(command)
--- a/src/Pure/Thy/sessions.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Thy/sessions.scala Fri Apr 01 17:06:10 2022 +0200
@@ -15,8 +15,7 @@
import scala.collection.mutable
-object Sessions
-{
+object Sessions {
/* session and theory names */
val ROOTS: Path = Path.explode("ROOTS")
@@ -40,8 +39,7 @@
/* ROOTS file format */
- class File_Format extends isabelle.File_Format
- {
+ class File_Format extends isabelle.File_Format {
val format_name: String = roots_name
val file_ext = ""
@@ -75,8 +73,8 @@
imported_sources: List[(Path, SHA1.Digest)] = Nil,
sources: List[(Path, SHA1.Digest)] = Nil,
session_graph_display: Graph_Display.Graph = Graph_Display.empty_graph,
- errors: List[String] = Nil)
- {
+ errors: List[String] = Nil
+ ) {
override def toString: String =
"Sessions.Base(loaded_theories = " + loaded_theories.size +
", used_theories = " + used_theories.length + ")"
@@ -100,8 +98,7 @@
nodes(name).syntax orElse loaded_theory_syntax(name) getOrElse overall_syntax
}
- sealed case class Deps(sessions_structure: Structure, session_bases: Map[String, Base])
- {
+ sealed case class Deps(sessions_structure: Structure, session_bases: Map[String, Base]) {
override def toString: String = "Sessions.Deps(" + sessions_structure + ")"
def is_empty: Boolean = session_bases.isEmpty
@@ -130,15 +127,14 @@
}
def deps(sessions_structure: Structure,
- progress: Progress = new Progress,
- inlined_files: Boolean = false,
- verbose: Boolean = false,
- list_files: Boolean = false,
- check_keywords: Set[String] = Set.empty): Deps =
- {
+ progress: Progress = new Progress,
+ inlined_files: Boolean = false,
+ verbose: Boolean = false,
+ list_files: Boolean = false,
+ check_keywords: Set[String] = Set.empty
+ ): Deps = {
var cache_sources = Map.empty[JFile, SHA1.Digest]
- def check_sources(paths: List[Path]): List[(Path, SHA1.Digest)] =
- {
+ def check_sources(paths: List[Path]): List[(Path, SHA1.Digest)] = {
for {
path <- paths
file = path.file
@@ -204,13 +200,11 @@
progress, overall_syntax.keywords, check_keywords, theory_files)
}
- val session_graph_display: Graph_Display.Graph =
- {
+ val session_graph_display: Graph_Display.Graph = {
def session_node(name: String): Graph_Display.Node =
Graph_Display.Node("[" + name + "]", "session." + name)
- def node(name: Document.Node.Name): Graph_Display.Node =
- {
+ def node(name: Document.Node.Name): Graph_Display.Node = {
val qualifier = deps_base.theory_qualifier(name)
if (qualifier == info.name)
Graph_Display.Node(name.theory_base_name, "theory." + name.theory)
@@ -251,8 +245,7 @@
val known_loaded_files = deps_base.known_loaded_files ++ loaded_files
- val import_errors =
- {
+ val import_errors = {
val known_sessions =
sessions_structure.imports_requirements(List(session_name)).toSet
for {
@@ -291,8 +284,7 @@
val document_theories =
info.document_theories.map({ case (thy, _) => known_theories(thy).name })
- val dir_errors =
- {
+ val dir_errors = {
val ok = info.dirs.map(_.canonical_file).toSet
val bad =
(for {
@@ -372,8 +364,8 @@
sessions_structure: Structure,
errors: List[String],
base: Base,
- infos: List[Info])
- {
+ infos: List[Info]
+ ) {
def check: Base_Info = if (errors.isEmpty) this else error(cat_lines(errors))
}
@@ -383,8 +375,8 @@
dirs: List[Path] = Nil,
include_sessions: List[String] = Nil,
session_ancestor: Option[String] = None,
- session_requirements: Boolean = false): Base_Info =
- {
+ session_requirements: Boolean = false
+ ): Base_Info = {
val full_sessions = load_structure(options, dirs = dirs)
val selected_sessions =
@@ -474,8 +466,8 @@
document_theories: List[(String, Position.T)],
document_files: List[(Path, Path)],
export_files: List[(Path, Int, List[String])],
- meta_digest: SHA1.Digest)
- {
+ meta_digest: SHA1.Digest
+ ) {
def chapter_session: String = chapter + "/" + name
def relative_path(info1: Info): String =
@@ -485,8 +477,7 @@
def deps: List[String] = parent.toList ::: imports
- def deps_base(session_bases: String => Base): Base =
- {
+ def deps_base(session_bases: String => Base): Base = {
val parent_base = session_bases(parent.getOrElse(""))
val imports_bases = imports.map(session_bases)
parent_base.copy(
@@ -514,8 +505,7 @@
case doc => error("Bad document specification " + quote(doc))
}
- def document_variants: List[Document_Build.Document_Variant] =
- {
+ def document_variants: List[Document_Build.Document_Variant] = {
val variants =
Library.space_explode(':', options.string("document_variants")).
map(Document_Build.Document_Variant.parse)
@@ -526,8 +516,7 @@
variants
}
- def documents: List[Document_Build.Document_Variant] =
- {
+ def documents: List[Document_Build.Document_Variant] = {
val variants = document_variants
if (!document_enabled || document_files.isEmpty) Nil else variants
}
@@ -553,9 +542,13 @@
def is_afp_bulky: Boolean = is_afp && groups.exists(AFP.groups_bulky.contains)
}
- def make_info(options: Options, dir_selected: Boolean, dir: Path, chapter: String,
- entry: Session_Entry): Info =
- {
+ def make_info(
+ options: Options,
+ dir_selected: Boolean,
+ dir: Path,
+ chapter: String,
+ entry: Session_Entry
+ ): Info = {
try {
val name = entry.name
@@ -615,8 +608,7 @@
}
}
- object Selection
- {
+ object Selection {
val empty: Selection = Selection()
val all: Selection = Selection(all_sessions = true)
def session(session: String): Selection = Selection(sessions = List(session))
@@ -629,8 +621,8 @@
exclude_session_groups: List[String] = Nil,
exclude_sessions: List[String] = Nil,
session_groups: List[String] = Nil,
- sessions: List[String] = Nil)
- {
+ sessions: List[String] = Nil
+ ) {
def ++ (other: Selection): Selection =
Selection(
requirements = requirements || other.requirements,
@@ -642,17 +634,16 @@
sessions = Library.merge(sessions, other.sessions))
}
- object Structure
- {
+ object Structure {
val empty: Structure = make(Nil)
- def make(infos: List[Info]): Structure =
- {
- def add_edges(graph: Graph[String, Info], kind: String, edges: Info => Iterable[String])
- : Graph[String, Info] =
- {
- def add_edge(pos: Position.T, name: String, g: Graph[String, Info], parent: String) =
- {
+ def make(infos: List[Info]): Structure = {
+ def add_edges(
+ graph: Graph[String, Info],
+ kind: String,
+ edges: Info => Iterable[String]
+ ) : Graph[String, Info] = {
+ def add_edge(pos: Position.T, name: String, g: Graph[String, Info], parent: String) = {
if (!g.defined(parent))
error("Bad " + kind + " session " + quote(parent) + " for " +
quote(name) + Position.here(pos))
@@ -725,12 +716,12 @@
}
final class Structure private[Sessions](
- val session_positions: List[(String, Position.T)],
- val session_directories: Map[JFile, String],
- val global_theories: Map[String, String],
- val build_graph: Graph[String, Info],
- val imports_graph: Graph[String, Info])
- {
+ val session_positions: List[(String, Position.T)],
+ val session_directories: Map[JFile, String],
+ val global_theories: Map[String, String],
+ val build_graph: Graph[String, Info],
+ val imports_graph: Graph[String, Info]
+ ) {
sessions_structure =>
def bootstrap: Base =
@@ -759,8 +750,7 @@
def theory_qualifier(name: String): String =
global_theories.getOrElse(name, Long_Name.qualifier(name))
- def check_sessions(names: List[String]): Unit =
- {
+ def check_sessions(names: List[String]): Unit = {
val bad_sessions = SortedSet(names.filterNot(defined): _*).toList
if (bad_sessions.nonEmpty)
error("Undefined session(s): " + commas_quote(bad_sessions))
@@ -769,8 +759,7 @@
def check_sessions(sel: Selection): Unit =
check_sessions(sel.base_sessions ::: sel.exclude_sessions ::: sel.sessions)
- private def selected(graph: Graph[String, Info], sel: Selection): List[String] =
- {
+ private def selected(graph: Graph[String, Info], sel: Selection): List[String] = {
check_sessions(sel)
val select_group = sel.session_groups.toSet
@@ -789,12 +778,10 @@
else selected0
}
- def selection(sel: Selection): Structure =
- {
+ def selection(sel: Selection): Structure = {
check_sessions(sel)
- val excluded =
- {
+ val excluded = {
val exclude_group = sel.exclude_session_groups.toSet
val exclude_group_sessions =
(for {
@@ -804,8 +791,7 @@
imports_graph.all_succs(exclude_group_sessions ::: sel.exclude_sessions).toSet
}
- def restrict(graph: Graph[String, Info]): Graph[String, Info] =
- {
+ def restrict(graph: Graph[String, Info]): Graph[String, Info] = {
val sessions = graph.all_preds(selected(graph, sel)).filterNot(excluded)
graph.restrict(graph.all_preds(sessions).toSet)
}
@@ -822,8 +808,8 @@
progress: Progress = new Progress,
loading_sessions: Boolean = false,
inlined_files: Boolean = false,
- verbose: Boolean = false): Deps =
- {
+ verbose: Boolean = false
+ ): Deps = {
val deps =
Sessions.deps(sessions_structure.selection(selection),
progress = progress, inlined_files = inlined_files, verbose = verbose)
@@ -904,25 +890,22 @@
theories: List[(List[Options.Spec], List[((String, Position.T), Boolean)])],
document_theories: List[(String, Position.T)],
document_files: List[(String, String)],
- export_files: List[(String, Int, List[String])]) extends Entry
- {
+ export_files: List[(String, Int, List[String])]
+ ) extends Entry {
def theories_no_position: List[(List[Options.Spec], List[(String, Boolean)])] =
theories.map({ case (a, b) => (a, b.map({ case ((c, _), d) => (c, d) })) })
def document_theories_no_position: List[String] =
document_theories.map(_._1)
}
- private object Parser extends Options.Parser
- {
- private val chapter: Parser[Chapter] =
- {
+ private object Parser extends Options.Parser {
+ private val chapter: Parser[Chapter] = {
val chapter_name = atom("chapter name", _.is_name)
command(CHAPTER) ~! chapter_name ^^ { case _ ~ a => Chapter(a) }
}
- private val session_entry: Parser[Session_Entry] =
- {
+ private val session_entry: Parser[Session_Entry] = {
val option =
option_name ~ opt($$$("=") ~! option_value ^^ { case _ ~ x => x }) ^^
{ case x ~ y => (x, y) }
@@ -969,8 +952,7 @@
Session_Entry(pos, a, b, c, d, e, f, g, h, i, j, k, l) }
}
- def parse_root(path: Path): List[Entry] =
- {
+ def parse_root(path: Path): List[Entry] = {
val toks = Token.explode(root_syntax.keywords, File.read(path))
val start = Token.Pos.file(path.implode)
@@ -987,8 +969,7 @@
for (entry <- Parser.parse_root(path) if entry.isInstanceOf[Session_Entry])
yield entry.asInstanceOf[Session_Entry]
- def read_root(options: Options, select: Boolean, path: Path): List[Info] =
- {
+ def read_root(options: Options, select: Boolean, path: Path): List[Info] = {
var entry_chapter = UNSORTED
val infos = new mutable.ListBuffer[Info]
parse_root(path).foreach {
@@ -999,8 +980,7 @@
infos.toList
}
- def parse_roots(roots: Path): List[String] =
- {
+ def parse_roots(roots: Path): List[String] = {
for {
line <- split_lines(File.read(roots))
if !(line == "" || line.startsWith("#"))
@@ -1017,29 +997,27 @@
if (is_session_dir(dir)) File.pwd() + dir.expand
else error("Bad session root directory (missing ROOT or ROOTS): " + dir.expand.toString)
- def directories(dirs: List[Path], select_dirs: List[Path]): List[(Boolean, Path)] =
- {
+ def directories(dirs: List[Path], select_dirs: List[Path]): List[(Boolean, Path)] = {
val default_dirs = Components.directories().filter(is_session_dir)
for { (select, dir) <- (default_dirs ::: dirs).map((false, _)) ::: select_dirs.map((true, _)) }
yield (select, dir.canonical)
}
- def load_structure(options: Options,
+ def load_structure(
+ options: Options,
dirs: List[Path] = Nil,
select_dirs: List[Path] = Nil,
- infos: List[Info] = Nil): Structure =
- {
+ infos: List[Info] = Nil
+ ): Structure = {
def load_dir(select: Boolean, dir: Path): List[(Boolean, Path)] =
load_root(select, dir) ::: load_roots(select, dir)
- def load_root(select: Boolean, dir: Path): List[(Boolean, Path)] =
- {
+ def load_root(select: Boolean, dir: Path): List[(Boolean, Path)] = {
val root = dir + ROOT
if (root.is_file) List((select, root)) else Nil
}
- def load_roots(select: Boolean, dir: Path): List[(Boolean, Path)] =
- {
+ def load_roots(select: Boolean, dir: Path): List[(Boolean, Path)] = {
val roots = dir + ROOTS
if (roots.is_file) {
for {
@@ -1079,8 +1057,7 @@
/* Isabelle tool wrapper */
val isabelle_tool = Isabelle_Tool("sessions", "explore structure of Isabelle sessions",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
var base_sessions: List[String] = Nil
var select_dirs: List[Path] = Nil
var requirements = false
@@ -1137,11 +1114,9 @@
private val sha1_prefix = "SHA1:"
- def read_heap_digest(heap: Path): Option[String] =
- {
+ def read_heap_digest(heap: Path): Option[String] = {
if (heap.is_file) {
- using(FileChannel.open(heap.java_path, StandardOpenOption.READ))(file =>
- {
+ using(FileChannel.open(heap.java_path, StandardOpenOption.READ))(file => {
val len = file.size
val n = sha1_prefix.length + SHA1.digest_length
if (len >= n) {
@@ -1183,8 +1158,7 @@
/** persistent store **/
- object Session_Info
- {
+ object Session_Info {
val session_name = SQL.Column.string("session_name").make_primary_key
// Build_Log.Session_Info
@@ -1210,8 +1184,8 @@
class Database_Context private[Sessions](
val store: Sessions.Store,
- database_server: Option[SQL.Database]) extends AutoCloseable
- {
+ database_server: Option[SQL.Database]
+ ) extends AutoCloseable {
def cache: Term.Cache = store.cache
def close(): Unit = database_server.foreach(_.close())
@@ -1233,8 +1207,10 @@
}
def read_export(
- sessions: List[String], theory_name: String, name: String): Option[Export.Entry] =
- {
+ sessions: List[String],
+ theory_name: String,
+ name: String
+ ): Option[Export.Entry] = {
val attempts =
database_server match {
case Some(db) =>
@@ -1256,8 +1232,7 @@
read_export(session_hierarchy, theory_name, name) getOrElse
Export.empty_entry(theory_name, name)
- override def toString: String =
- {
+ override def toString: String = {
val s =
database_server match {
case Some(db) => db.toString
@@ -1270,8 +1245,7 @@
def store(options: Options, cache: Term.Cache = Term.Cache.make()): Store =
new Store(options, cache)
- class Store private[Sessions](val options: Options, val cache: Term.Cache)
- {
+ class Store private[Sessions](val options: Options, val cache: Term.Cache) {
store =>
override def toString: String = "Store(output_dir = " + output_dir.absolute + ")"
@@ -1347,8 +1321,7 @@
def open_database_context(): Database_Context =
new Database_Context(store, if (database_server) Some(open_database_server()) else None)
- def try_open_database(name: String, output: Boolean = false): Option[SQL.Database] =
- {
+ def try_open_database(name: String, output: Boolean = false): Option[SQL.Database] = {
def check(db: SQL.Database): Option[SQL.Database] =
if (output || session_info_exists(db)) Some(db) else { db.close(); None }
@@ -1367,11 +1340,9 @@
try_open_database(name, output = output) getOrElse
error("Missing build database for session " + quote(name))
- def clean_output(name: String): (Boolean, Boolean) =
- {
+ def clean_output(name: String): (Boolean, Boolean) = {
val relevant_db =
- database_server &&
- {
+ database_server && {
try_open_database(name) match {
case Some(db) =>
try {
@@ -1403,8 +1374,7 @@
def read_bytes(db: SQL.Database, name: String, column: SQL.Column): Bytes =
db.using_statement(Session_Info.table.select(List(column),
- Session_Info.session_name.where_equal(name)))(stmt =>
- {
+ Session_Info.session_name.where_equal(name)))(stmt => {
val res = stmt.execute_query()
if (!res.next()) Bytes.empty else res.bytes(column)
})
@@ -1415,8 +1385,7 @@
/* session info */
- def init_session_info(db: SQL.Database, name: String): Unit =
- {
+ def init_session_info(db: SQL.Database, name: String): Unit = {
db.transaction {
db.create_table(Session_Info.table)
db.using_statement(
@@ -1433,8 +1402,7 @@
}
}
- def session_info_exists(db: SQL.Database): Boolean =
- {
+ def session_info_exists(db: SQL.Database): Boolean = {
val tables = db.tables
tables.contains(Session_Info.table.name) &&
tables.contains(Export.Data.table.name)
@@ -1442,8 +1410,7 @@
def session_info_defined(db: SQL.Database, name: String): Boolean =
db.transaction {
- session_info_exists(db) &&
- {
+ session_info_exists(db) && {
db.using_statement(
Session_Info.table.select(List(Session_Info.session_name),
Session_Info.session_name.where_equal(name)))(stmt => stmt.execute_query().next())
@@ -1454,11 +1421,10 @@
db: SQL.Database,
name: String,
build_log: Build_Log.Session_Info,
- build: Build.Session_Info): Unit =
- {
+ build: Build.Session_Info
+ ): Unit = {
db.transaction {
- db.using_statement(Session_Info.table.insert())(stmt =>
- {
+ db.using_statement(Session_Info.table.insert())(stmt => {
stmt.string(1) = name
stmt.bytes(2) = Properties.encode(build_log.session_timing)
stmt.bytes(3) = Properties.compress(build_log.command_timings, cache = cache.xz)
@@ -1496,12 +1462,10 @@
def read_errors(db: SQL.Database, name: String): List[String] =
Build_Log.uncompress_errors(read_bytes(db, name, Session_Info.errors), cache = cache)
- def read_build(db: SQL.Database, name: String): Option[Build.Session_Info] =
- {
+ def read_build(db: SQL.Database, name: String): Option[Build.Session_Info] = {
if (db.tables.contains(Session_Info.table.name)) {
db.using_statement(Session_Info.table.select(Session_Info.build_columns,
- Session_Info.session_name.where_equal(name)))(stmt =>
- {
+ Session_Info.session_name.where_equal(name)))(stmt => {
val res = stmt.execute_query()
if (!res.next()) None
else {
--- a/src/Pure/Thy/thy_element.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Thy/thy_element.scala Fri Apr 01 17:06:10 2022 +0200
@@ -10,13 +10,11 @@
import scala.util.parsing.input
-object Thy_Element
-{
+object Thy_Element {
/* datatype element */
type Proof[A] = (List[Element[A]], A)
- sealed case class Element[A](head: A, proof: Option[Proof[A]])
- {
+ sealed case class Element[A](head: A, proof: Option[Proof[A]]) {
def iterator: Iterator[A] =
Iterator(head) ++
(for {
@@ -52,24 +50,20 @@
type Element_Command = Element[Command]
- def parse_elements(keywords: Keyword.Keywords, commands: List[Command]): List[Element_Command] =
- {
- case class Reader(in: List[Command]) extends input.Reader[Command]
- {
+ def parse_elements(keywords: Keyword.Keywords, commands: List[Command]): List[Element_Command] = {
+ case class Reader(in: List[Command]) extends input.Reader[Command] {
def first: Command = in.head
def rest: Reader = Reader(in.tail)
def pos: input.Position = input.NoPosition
def atEnd: Boolean = in.isEmpty
}
- object Parser extends Parsers
- {
+ object Parser extends Parsers {
type Elem = Command
def command(pred: Command => Boolean): Parser[Command] =
new Parser[Elem] {
- def apply(in: Input) =
- {
+ def apply(in: Input) = {
if (in.atEnd) Failure("end of input", in)
else {
val command = in.first
--- a/src/Pure/Thy/thy_header.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Thy/thy_header.scala Fri Apr 01 17:06:10 2022 +0200
@@ -11,8 +11,7 @@
import scala.util.matching.Regex
-object Thy_Header
-{
+object Thy_Header {
/* bootstrap keywords */
type Keywords = List[(String, Keyword.Spec)]
@@ -114,8 +113,7 @@
def bootstrap_name(theory: String): String =
bootstrap_thys.collectFirst({ case (a, b) if a == theory => b }).getOrElse(theory)
- def try_read_dir(dir: Path): List[String] =
- {
+ def try_read_dir(dir: Path): List[String] = {
val files =
try { File.read_dir(dir) }
catch { case ERROR(_) => Nil }
@@ -125,10 +123,8 @@
/* parser */
- trait Parser extends Parse.Parser
- {
- val header: Parser[Thy_Header] =
- {
+ trait Parser extends Parse.Parser {
+ val header: Parser[Thy_Header] = {
val load_command =
($$$("(") ~! (position(name) <~ $$$(")")) ^^ { case _ ~ x => x }) |
success(("", Position.none))
@@ -183,8 +179,7 @@
/* read -- lazy scanning */
- private def read_tokens(reader: Reader[Char], strict: Boolean): (List[Token], List[Token]) =
- {
+ private def read_tokens(reader: Reader[Char], strict: Boolean): (List[Token], List[Token]) = {
val token = Token.Parsers.token(bootstrap_keywords)
def make_tokens(in: Reader[Char]): LazyList[Token] =
token(in) match {
@@ -204,8 +199,7 @@
(drop_tokens, tokens1 ::: tokens2)
}
- private object Parser extends Parser
- {
+ private object Parser extends Parser {
def parse_header(tokens: List[Token], pos: Token.Pos): Thy_Header =
parse(commit(header), Token.reader(tokens, pos)) match {
case Success(result, _) => result
@@ -215,8 +209,8 @@
def read(node_name: Document.Node.Name, reader: Reader[Char],
command: Boolean = true,
- strict: Boolean = true): Thy_Header =
- {
+ strict: Boolean = true
+ ): Thy_Header = {
val (_, tokens0) = read_tokens(reader, true)
val text = Scan.reader_decode_utf8(reader, Token.implode(tokens0))
@@ -234,16 +228,15 @@
pos: Position.T,
imports: List[(String, Position.T)],
keywords: Thy_Header.Keywords,
- abbrevs: Thy_Header.Abbrevs)
-{
+ abbrevs: Thy_Header.Abbrevs
+) {
def map(f: String => String): Thy_Header =
Thy_Header(f(name), pos,
imports.map({ case (a, b) => (f(a), b) }),
keywords.map({ case (a, spec) => (f(a), spec.map(f)) }),
abbrevs.map({ case (a, b) => (f(a), f(b)) }))
- def check(node_name: Document.Node.Name): Thy_Header =
- {
+ def check(node_name: Document.Node.Name): Thy_Header = {
val base_name = node_name.theory_base_name
if (Long_Name.is_qualified(name)) {
error("Bad theory name " + quote(name) + " with qualification" + Position.here(pos))
--- a/src/Pure/Thy/thy_syntax.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Thy/thy_syntax.scala Fri Apr 01 17:06:10 2022 +0200
@@ -11,24 +11,24 @@
import scala.annotation.tailrec
-object Thy_Syntax
-{
+object Thy_Syntax {
/** perspective **/
def command_perspective(
- node: Document.Node,
- perspective: Text.Perspective,
- overlays: Document.Node.Overlays): (Command.Perspective, Command.Perspective) =
- {
+ node: Document.Node,
+ perspective: Text.Perspective,
+ overlays: Document.Node.Overlays
+ ): (Command.Perspective, Command.Perspective) = {
if (perspective.is_empty && overlays.is_empty)
(Command.Perspective.empty, Command.Perspective.empty)
else {
val has_overlay = overlays.commands
val visible = new mutable.ListBuffer[Command]
val visible_overlay = new mutable.ListBuffer[Command]
- @tailrec
- def check_ranges(ranges: List[Text.Range], commands: LazyList[(Command, Text.Offset)]): Unit =
- {
+ @tailrec def check_ranges(
+ ranges: List[Text.Range],
+ commands: LazyList[(Command, Text.Offset)]
+ ): Unit = {
(ranges, commands) match {
case (range :: more_ranges, (command, offset) #:: more_commands) =>
val command_range = command.range + offset
@@ -68,9 +68,8 @@
private def header_edits(
resources: Resources,
previous: Document.Version,
- edits: List[Document.Edit_Text]):
- (List[Document.Node.Name], Document.Nodes, List[Document.Edit_Command]) =
- {
+ edits: List[Document.Edit_Text]
+ ): (List[Document.Node.Name], Document.Nodes, List[Document.Edit_Command]) = {
val syntax_changed0 = new mutable.ListBuffer[Document.Node.Name]
var nodes = previous.nodes
val doc_edits = new mutable.ListBuffer[Document.Edit_Command]
@@ -119,8 +118,10 @@
/* edit individual command source */
- @tailrec def edit_text(eds: List[Text.Edit], commands: Linear_Set[Command]): Linear_Set[Command] =
- {
+ @tailrec def edit_text(
+ eds: List[Text.Edit],
+ commands: Linear_Set[Command]
+ ): Linear_Set[Command] = {
eds match {
case e :: es =>
def insert_text(cmd: Option[Command], text: String): Linear_Set[Command] =
@@ -151,10 +152,9 @@
/* reparse range of command spans */
@tailrec private def chop_common(
- cmds: List[Command],
- blobs_spans: List[(Command.Blobs_Info, Command_Span.Span)])
- : (List[Command], List[(Command.Blobs_Info, Command_Span.Span)]) =
- {
+ cmds: List[Command],
+ blobs_spans: List[(Command.Blobs_Info, Command_Span.Span)]
+ ) : (List[Command], List[(Command.Blobs_Info, Command_Span.Span)]) = {
(cmds, blobs_spans) match {
case (cmd :: cmds, (blobs_info, span) :: rest)
if cmd.blobs_info == blobs_info && cmd.span == span => chop_common(cmds, rest)
@@ -169,8 +169,8 @@
can_import: Document.Node.Name => Boolean,
node_name: Document.Node.Name,
commands: Linear_Set[Command],
- first: Command, last: Command): Linear_Set[Command] =
- {
+ first: Command, last: Command
+ ): Linear_Set[Command] = {
val cmds0 = commands.iterator(first, last).toList
val blobs_spans0 =
syntax.parse_spans(cmds0.iterator.map(_.source).mkString).map(span =>
@@ -198,9 +198,10 @@
/* main */
- def diff_commands(old_cmds: Linear_Set[Command], new_cmds: Linear_Set[Command])
- : List[Command.Edit] =
- {
+ def diff_commands(
+ old_cmds: Linear_Set[Command],
+ new_cmds: Linear_Set[Command]
+ ) : List[Command.Edit] = {
val removed = old_cmds.iterator.filter(!new_cmds.contains(_)).toList
val inserted = new_cmds.iterator.filter(!old_cmds.contains(_)).toList
@@ -214,15 +215,15 @@
get_blob: Document.Node.Name => Option[Document.Blob],
can_import: Document.Node.Name => Boolean,
reparse_limit: Int,
- node: Document.Node, edit: Document.Edit_Text): Document.Node =
- {
+ node: Document.Node, edit: Document.Edit_Text
+ ): Document.Node = {
/* recover command spans after edits */
// FIXME somewhat slow
def recover_spans(
name: Document.Node.Name,
perspective: Command.Perspective,
- commands: Linear_Set[Command]): Linear_Set[Command] =
- {
+ commands: Linear_Set[Command]
+ ): Linear_Set[Command] = {
val is_visible = perspective.commands.toSet
def next_invisible(cmds: Linear_Set[Command], from: Command): Command =
@@ -291,13 +292,13 @@
}
def parse_change(
- resources: Resources,
- reparse_limit: Int,
- previous: Document.Version,
- doc_blobs: Document.Blobs,
- edits: List[Document.Edit_Text],
- consolidate: List[Document.Node.Name]): Session.Change =
- {
+ resources: Resources,
+ reparse_limit: Int,
+ previous: Document.Version,
+ doc_blobs: Document.Blobs,
+ edits: List[Document.Edit_Text],
+ consolidate: List[Document.Node.Name]
+ ): Session.Change = {
val (syntax_changed, nodes0, doc_edits0) = header_edits(resources, previous, edits)
def get_blob(name: Document.Node.Name): Option[Document.Blob] =
--- a/src/Pure/Tools/build.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/build.scala Fri Apr 01 17:06:10 2022 +0200
@@ -12,8 +12,7 @@
import scala.annotation.tailrec
-object Build
-{
+object Build {
/** auxiliary **/
/* persistent build info */
@@ -22,27 +21,24 @@
sources: String,
input_heaps: List[String],
output_heap: Option[String],
- return_code: Int)
- {
+ return_code: Int
+ ) {
def ok: Boolean = return_code == 0
}
/* queue with scheduling information */
- private object Queue
- {
+ private object Queue {
type Timings = (List[Properties.T], Double)
- def load_timings(progress: Progress, store: Sessions.Store, session_name: String): Timings =
- {
+ def load_timings(progress: Progress, store: Sessions.Store, session_name: String): Timings = {
val no_timings: Timings = (Nil, 0.0)
store.try_open_database(session_name) match {
case None => no_timings
case Some(db) =>
- def ignore_error(msg: String) =
- {
+ def ignore_error(msg: String) = {
progress.echo_warning("Ignoring bad database " + db + (if (msg == "") "" else "\n" + msg))
no_timings
}
@@ -64,12 +60,12 @@
}
}
- def make_session_timing(sessions_structure: Sessions.Structure, timing: Map[String, Double])
- : Map[String, Double] =
- {
+ def make_session_timing(
+ sessions_structure: Sessions.Structure,
+ timing: Map[String, Double]
+ ) : Map[String, Double] = {
val maximals = sessions_structure.build_graph.maximals.toSet
- def desc_timing(session_name: String): Double =
- {
+ def desc_timing(session_name: String): Double = {
if (maximals.contains(session_name)) timing(session_name)
else {
val descendants = sessions_structure.build_descendants(List(session_name)).toSet
@@ -84,9 +80,11 @@
timing.keySet.iterator.map(name => (name -> desc_timing(name))).toMap.withDefaultValue(0.0)
}
- def apply(progress: Progress, sessions_structure: Sessions.Structure, store: Sessions.Store)
- : Queue =
- {
+ def apply(
+ progress: Progress,
+ sessions_structure: Sessions.Structure,
+ store: Sessions.Store
+ ) : Queue = {
val graph = sessions_structure.build_graph
val names = graph.keys
@@ -97,8 +95,7 @@
make_session_timing(sessions_structure,
timings.map({ case (name, (_, t)) => (name, t) }).toMap)
- object Ordering extends scala.math.Ordering[String]
- {
+ object Ordering extends scala.math.Ordering[String] {
def compare(name1: String, name2: String): Int =
session_timing(name2) compare session_timing(name1) match {
case 0 =>
@@ -117,8 +114,8 @@
private class Queue(
graph: Graph[String, Sessions.Info],
order: SortedSet[String],
- val command_timings: String => List[Properties.T])
- {
+ val command_timings: String => List[Properties.T]
+ ) {
def is_inner(name: String): Boolean = !graph.is_maximal(name)
def is_empty: Boolean = graph.is_empty
@@ -126,8 +123,7 @@
def - (name: String): Queue =
new Queue(graph.del_node(name), order - name, command_timings)
- def dequeue(skip: String => Boolean): Option[(String, Sessions.Info)] =
- {
+ def dequeue(skip: String => Boolean): Option[(String, Sessions.Info)] = {
val it = order.iterator.dropWhile(name => skip(name) || !graph.is_minimal(name))
if (it.hasNext) { val name = it.next(); Some((name, graph.get_node(name))) }
else None
@@ -138,8 +134,7 @@
/** build with results **/
- class Results private[Build](results: Map[String, (Option[Process_Result], Sessions.Info)])
- {
+ class Results private[Build](results: Map[String, (Option[Process_Result], Sessions.Info)]) {
def sessions: Set[String] = results.keySet
def infos: List[Sessions.Info] = results.values.map(_._2).toList
def cancelled(name: String): Boolean = results(name)._1.isEmpty
@@ -157,8 +152,7 @@
def session_finished(session_name: String, process_result: Process_Result): String =
"Finished " + session_name + " (" + process_result.timing.message_resources + ")"
- def session_timing(session_name: String, build_log: Build_Log.Session_Info): String =
- {
+ def session_timing(session_name: String, build_log: Build_Log.Session_Info): String = {
val props = build_log.session_timing
val threads = Markup.Session_Timing.Threads.unapply(props) getOrElse 1
val timing = Markup.Timing_Properties.get(props)
@@ -185,8 +179,8 @@
soft_build: Boolean = false,
verbose: Boolean = false,
export_files: Boolean = false,
- session_setup: (String, Session) => Unit = (_, _) => ()): Results =
- {
+ session_setup: (String, Session) => Unit = (_, _) => ()
+ ): Results = {
val build_options =
options +
"completion_limit=0" +
@@ -206,8 +200,7 @@
val full_sessions_selection = full_sessions.imports_selection(selection)
val full_sessions_selected = full_sessions_selection.toSet
- def sources_stamp(deps: Sessions.Deps, session_name: String): String =
- {
+ def sources_stamp(deps: Sessions.Deps, session_name: String): String = {
val digests =
full_sessions(session_name).meta_digest ::
deps.sources(session_name) :::
@@ -215,8 +208,7 @@
SHA1.digest_set(digests).toString
}
- val deps =
- {
+ val deps = {
val deps0 =
Sessions.deps(full_sessions.selection(selection),
progress = progress, inlined_files = true, verbose = verbose,
@@ -289,8 +281,8 @@
current: Boolean,
heap_digest: Option[String],
process: Option[Process_Result],
- info: Sessions.Info)
- {
+ info: Sessions.Info
+ ) {
def ok: Boolean =
process match {
case None => false
@@ -313,8 +305,8 @@
@tailrec def loop(
pending: Queue,
running: Map[String, (List[String], Build_Job)],
- results: Map[String, Result]): Map[String, Result] =
- {
+ results: Map[String, Result]
+ ): Map[String, Result] = {
def used_node(i: Int): Boolean =
running.iterator.exists(
{ case (_, (_, job)) => job.numa_node.isDefined && job.numa_node.get == i })
@@ -332,8 +324,7 @@
val (process_result, heap_digest) = job.join
val log_lines = process_result.out_lines.filterNot(Protocol_Message.Marker.test)
- val process_result_tail =
- {
+ val process_result_tail = {
val tail = job.info.options.int("process_output_tail")
process_result.copy(
out_lines =
@@ -392,8 +383,7 @@
val do_store =
build_heap || Sessions.is_pure(session_name) || queue.is_inner(session_name)
- val (current, heap_digest) =
- {
+ val (current, heap_digest) = {
store.try_open_database(session_name) match {
case Some(db) =>
using(db)(store.read_build(_, session_name)) match {
@@ -453,8 +443,7 @@
/* build results */
- val results =
- {
+ val results = {
val results0 =
if (deps.is_empty) {
progress.echo_warning("Nothing to build")
@@ -505,13 +494,11 @@
Presentation.update_chapter(presentation_dir, chapter, entries)
}
- using(store.open_database_context())(db_context =>
- {
+ using(store.open_database_context())(db_context => {
val exports =
Presentation.read_exports(presentation_sessions.map(_.name), deps, db_context)
- Par_List.map((session: String) =>
- {
+ Par_List.map((session: String) => {
progress.expose_interrupt()
progress.echo("Presenting " + session + " ...")
@@ -538,8 +525,7 @@
/* Isabelle tool wrapper */
val isabelle_tool = Isabelle_Tool("build", "build and manage Isabelle sessions",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
val build_options = Word.explode(Isabelle_System.getenv("ISABELLE_BUILD_OPTIONS"))
var base_sessions: List[String] = Nil
@@ -679,8 +665,8 @@
build_heap: Boolean = false,
dirs: List[Path] = Nil,
fresh: Boolean = false,
- strict: Boolean = false): Int =
- {
+ strict: Boolean = false
+ ): Int = {
val selection = Sessions.Selection.session(logic)
val rc =
if (!fresh && build(options, selection = selection,
--- a/src/Pure/Tools/build_docker.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/build_docker.scala Fri Apr 01 17:06:10 2022 +0200
@@ -7,8 +7,7 @@
package isabelle
-object Build_Docker
-{
+object Build_Docker {
private val default_base = "ubuntu"
private lazy val default_logic = Isabelle_System.getenv("ISABELLE_LOGIC")
@@ -31,8 +30,8 @@
output: Option[Path] = None,
more_packages: List[String] = Nil,
tag: String = "",
- verbose: Boolean = false): Unit =
- {
+ verbose: Boolean = false
+ ): Unit = {
val isabelle_name =
app_archive match {
case Isabelle_Name(name) => name
@@ -78,8 +77,7 @@
output.foreach(File.write(_, dockerfile))
if (!no_build) {
- Isabelle_System.with_tmp_dir("docker")(tmp_dir =>
- {
+ Isabelle_System.with_tmp_dir("docker")(tmp_dir => {
File.write(tmp_dir + Path.explode("Dockerfile"), dockerfile)
if (is_remote) {
@@ -102,8 +100,7 @@
val isabelle_tool =
Isabelle_Tool("build_docker", "build Isabelle docker image",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
var base = default_base
var entrypoint = false
var logic = default_logic
--- a/src/Pure/Tools/build_job.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/build_job.scala Fri Apr 01 17:06:10 2022 +0200
@@ -12,16 +12,15 @@
import scala.collection.mutable
-object Build_Job
-{
+object Build_Job {
/* theory markup/messages from database */
def read_theory(
db_context: Sessions.Database_Context,
session_hierarchy: List[String],
theory: String,
- unicode_symbols: Boolean = false): Option[Command] =
- {
+ unicode_symbols: Boolean = false
+ ): Option[Command] = {
def read(name: String): Export.Entry =
db_context.get_export(session_hierarchy, theory, name)
@@ -40,8 +39,7 @@
yield i -> elem)
val blobs =
- blobs_files.map(file =>
- {
+ blobs_files.map(file => {
val path = Path.explode(file)
val name = Resources.file_node(path)
val src_path = File.relative_path(node_name.master_dir_path, path).getOrElse(path)
@@ -91,16 +89,14 @@
margin: Double = Pretty.default_margin,
breakgain: Double = Pretty.default_breakgain,
metric: Pretty.Metric = Symbol.Metric,
- unicode_symbols: Boolean = false): Unit =
- {
+ unicode_symbols: Boolean = false
+ ): Unit = {
val store = Sessions.store(options)
val session = new Session(options, Resources.empty)
- using(store.open_database_context())(db_context =>
- {
+ using(store.open_database_context())(db_context => {
val result =
- db_context.input_database(session_name)((db, _) =>
- {
+ db_context.input_database(session_name)((db, _) => {
val theories = store.read_theories(db, session_name)
val errors = store.read_errors(db, session_name)
store.read_build(db, session_name).map(info => (theories, errors, info.return_code))
@@ -154,8 +150,7 @@
/* Isabelle tool wrapper */
val isabelle_tool = Isabelle_Tool("log", "print messages from build database",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
/* arguments */
var unicode_symbols = false
@@ -207,8 +202,8 @@
log: Logger,
session_setup: (String, Session) => Unit,
val numa_node: Option[Int],
- command_timings0: List[Properties.T])
-{
+ command_timings0: List[Properties.T]
+) {
val options: Options = NUMA.policy_options(info.options, numa_node)
private val sessions_structure = deps.sessions_structure
@@ -240,8 +235,7 @@
new Session(options, resources) {
override val cache: Term.Cache = store.cache
- override def build_blobs_info(name: Document.Node.Name): Command.Blobs_Info =
- {
+ override def build_blobs_info(name: Document.Node.Name): Command.Blobs_Info = {
result_base.load_commands.get(name.expand) match {
case Some(spans) =>
val syntax = result_base.theory_syntax(name)
@@ -251,14 +245,12 @@
}
}
- object Build_Session_Errors
- {
+ object Build_Session_Errors {
private val promise: Promise[List[String]] = Future.promise
def result: Exn.Result[List[String]] = promise.join_result
def cancel(): Unit = promise.cancel()
- def apply(errs: List[String]): Unit =
- {
+ def apply(errs: List[String]): Unit = {
try { promise.fulfill(errs) }
catch { case _: IllegalStateException => }
}
@@ -279,8 +271,8 @@
def fun(
name: String,
acc: mutable.ListBuffer[Properties.T],
- unapply: Properties.T => Option[Properties.T]): (String, Session.Protocol_Function) =
- {
+ unapply: Properties.T => Option[Properties.T]
+ ): (String, Session.Protocol_Function) = {
name -> ((msg: Prover.Protocol_Output) =>
unapply(msg.properties) match {
case Some(props) => acc += props; true
@@ -288,16 +280,13 @@
})
}
- session.init_protocol_handler(new Session.Protocol_Handler
- {
+ session.init_protocol_handler(new Session.Protocol_Handler {
override def exit(): Unit = Build_Session_Errors.cancel()
- private def build_session_finished(msg: Prover.Protocol_Output): Boolean =
- {
+ private def build_session_finished(msg: Prover.Protocol_Output): Boolean = {
val (rc, errors) =
try {
- val (rc, errs) =
- {
+ val (rc, errs) = {
import XML.Decode._
pair(int, list(x => x))(Symbol.decode_yxml(msg.text))
}
@@ -341,81 +330,76 @@
fun(Markup.Task_Statistics.name, task_statistics, Markup.Task_Statistics.unapply))
})
- session.command_timings += Session.Consumer("command_timings")
- {
- case Session.Command_Timing(props) =>
- for {
- elapsed <- Markup.Elapsed.unapply(props)
- elapsed_time = Time.seconds(elapsed)
- if elapsed_time.is_relevant && elapsed_time >= options.seconds("command_timing_threshold")
- } command_timings += props.filter(Markup.command_timing_property)
- }
+ session.command_timings += Session.Consumer("command_timings") {
+ case Session.Command_Timing(props) =>
+ for {
+ elapsed <- Markup.Elapsed.unapply(props)
+ elapsed_time = Time.seconds(elapsed)
+ if elapsed_time.is_relevant && elapsed_time >= options.seconds("command_timing_threshold")
+ } command_timings += props.filter(Markup.command_timing_property)
+ }
- session.runtime_statistics += Session.Consumer("ML_statistics")
- {
- case Session.Runtime_Statistics(props) => runtime_statistics += props
- }
+ session.runtime_statistics += Session.Consumer("ML_statistics") {
+ case Session.Runtime_Statistics(props) => runtime_statistics += props
+ }
- session.finished_theories += Session.Consumer[Document.Snapshot]("finished_theories")
- {
- case snapshot =>
- if (!progress.stopped) {
- val rendering = new Rendering(snapshot, options, session)
+ session.finished_theories += Session.Consumer[Document.Snapshot]("finished_theories") {
+ case snapshot =>
+ if (!progress.stopped) {
+ val rendering = new Rendering(snapshot, options, session)
- def export_(name: String, xml: XML.Body, compress: Boolean = true): Unit =
- {
- if (!progress.stopped) {
- val theory_name = snapshot.node_name.theory
- val args =
- Protocol.Export.Args(theory_name = theory_name, name = name, compress = compress)
- val bytes = Bytes(Symbol.encode(YXML.string_of_body(xml)))
- if (!bytes.is_empty) export_consumer(session_name, args, bytes)
- }
+ def export_(name: String, xml: XML.Body, compress: Boolean = true): Unit = {
+ if (!progress.stopped) {
+ val theory_name = snapshot.node_name.theory
+ val args =
+ Protocol.Export.Args(theory_name = theory_name, name = name, compress = compress)
+ val bytes = Bytes(Symbol.encode(YXML.string_of_body(xml)))
+ if (!bytes.is_empty) export_consumer(session_name, args, bytes)
}
- def export_text(name: String, text: String, compress: Boolean = true): Unit =
- export_(name, List(XML.Text(text)), compress = compress)
+ }
+ def export_text(name: String, text: String, compress: Boolean = true): Unit =
+ export_(name, List(XML.Text(text)), compress = compress)
- for (command <- snapshot.snippet_command) {
- export_text(Export.DOCUMENT_ID, command.id.toString, compress = false)
- }
+ for (command <- snapshot.snippet_command) {
+ export_text(Export.DOCUMENT_ID, command.id.toString, compress = false)
+ }
- export_text(Export.FILES,
- cat_lines(snapshot.node_files.map(_.symbolic.node)), compress = false)
+ export_text(Export.FILES,
+ cat_lines(snapshot.node_files.map(_.symbolic.node)), compress = false)
- for (((_, xml), i) <- snapshot.xml_markup_blobs().zipWithIndex) {
- export_(Export.MARKUP + (i + 1), xml)
- }
- export_(Export.MARKUP, snapshot.xml_markup())
- export_(Export.MESSAGES, snapshot.messages.map(_._1))
+ for (((_, xml), i) <- snapshot.xml_markup_blobs().zipWithIndex) {
+ export_(Export.MARKUP + (i + 1), xml)
+ }
+ export_(Export.MARKUP, snapshot.xml_markup())
+ export_(Export.MESSAGES, snapshot.messages.map(_._1))
- val citations = Library.distinct(rendering.citations(Text.Range.full).map(_.info))
- export_text(Export.DOCUMENT_CITATIONS, cat_lines(citations))
- }
- }
+ val citations = Library.distinct(rendering.citations(Text.Range.full).map(_.info))
+ export_text(Export.DOCUMENT_CITATIONS, cat_lines(citations))
+ }
+ }
- session.all_messages += Session.Consumer[Any]("build_session_output")
- {
- case msg: Prover.Output =>
- val message = msg.message
- if (msg.is_system) resources.log(Protocol.message_text(message))
+ session.all_messages += Session.Consumer[Any]("build_session_output") {
+ case msg: Prover.Output =>
+ val message = msg.message
+ if (msg.is_system) resources.log(Protocol.message_text(message))
- if (msg.is_stdout) {
- stdout ++= Symbol.encode(XML.content(message))
- }
- else if (msg.is_stderr) {
- stderr ++= Symbol.encode(XML.content(message))
- }
- else if (msg.is_exit) {
- val err =
- "Prover terminated" +
- (msg.properties match {
- case Markup.Process_Result(result) => ": " + result.print_rc
- case _ => ""
- })
- Build_Session_Errors(List(err))
- }
- case _ =>
- }
+ if (msg.is_stdout) {
+ stdout ++= Symbol.encode(XML.content(message))
+ }
+ else if (msg.is_stderr) {
+ stderr ++= Symbol.encode(XML.content(message))
+ }
+ else if (msg.is_exit) {
+ val err =
+ "Prover terminated" +
+ (msg.properties match {
+ case Markup.Process_Result(result) => ": " + result.print_rc
+ case _ => ""
+ })
+ Build_Session_Errors(List(err))
+ }
+ case _ =>
+ }
session_setup(session_name, session)
@@ -458,8 +442,7 @@
val (document_output, document_errors) =
try {
if (build_errors.isInstanceOf[Exn.Res[_]] && process_result.ok && info.documents.nonEmpty) {
- using(store.open_database_context())(db_context =>
- {
+ using(store.open_database_context())(db_context => {
val documents =
Document_Build.build_documents(
Document_Build.context(session_name, deps, db_context, progress = progress),
@@ -477,11 +460,9 @@
case Exn.Interrupt.ERROR(msg) => (Nil, List(msg))
}
- val result =
- {
+ val result = {
val theory_timing =
- theory_timings.iterator.map(
- { case props @ Markup.Name(name) => name -> props }).toMap
+ theory_timings.iterator.map({ case props @ Markup.Name(name) => name -> props }).toMap
val used_theory_timings =
for { (name, _) <- deps(session_name).used_theories }
yield theory_timing.getOrElse(name.theory, Markup.Name(name.theory))
@@ -520,14 +501,12 @@
def terminate(): Unit = future_result.cancel()
def is_finished: Boolean = future_result.is_finished
- private val timeout_request: Option[Event_Timer.Request] =
- {
+ private val timeout_request: Option[Event_Timer.Request] = {
if (info.timeout_ignored) None
else Some(Event_Timer.request(Time.now() + info.timeout) { terminate() })
}
- def join: (Process_Result, Option[String]) =
- {
+ def join: (Process_Result, Option[String]) = {
val result1 = future_result.join
val was_timeout =
--- a/src/Pure/Tools/check_keywords.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/check_keywords.scala Fri Apr 01 17:06:10 2022 +0200
@@ -7,16 +7,14 @@
package isabelle
-object Check_Keywords
-{
+object Check_Keywords {
def conflicts(
keywords: Keyword.Keywords,
check: Set[String],
input: CharSequence,
- start: Token.Pos): List[(Token, Position.T)] =
- {
- object Parser extends Parse.Parser
- {
+ start: Token.Pos
+ ): List[(Token, Position.T)] = {
+ object Parser extends Parse.Parser {
private val conflict =
position(token("token", tok => !(tok.is_command || tok.is_keyword) && check(tok.source)))
private val other = token("token", _ => true)
@@ -35,8 +33,8 @@
progress: Progress,
keywords: Keyword.Keywords,
check: Set[String],
- paths: List[Path]): Unit =
- {
+ paths: List[Path]
+ ): Unit = {
val parallel_args = paths.map(path => (File.read(path), Token.Pos.file(path.expand.implode)))
val bad =
--- a/src/Pure/Tools/debugger.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/debugger.scala Fri Apr 01 17:06:10 2022 +0200
@@ -10,8 +10,7 @@
import scala.collection.immutable.SortedMap
-object Debugger
-{
+object Debugger {
/* thread context */
case object Update
@@ -19,8 +18,7 @@
sealed case class Debug_State(pos: Position.T, function: String)
type Threads = SortedMap[String, List[Debug_State]]
- sealed case class Context(thread_name: String, debug_states: List[Debug_State], index: Int = 0)
- {
+ sealed case class Context(thread_name: String, debug_states: List[Debug_State], index: Int = 0) {
def size: Int = debug_states.length + 1
def reset: Context = copy(index = 0)
def select(i: Int): Context = copy(index = i + 1)
@@ -56,16 +54,15 @@
active_breakpoints: Set[Long] = Set.empty, // explicit breakpoint state
threads: Threads = SortedMap.empty, // thread name ~> stack of debug states
focus: Map[String, Context] = Map.empty, // thread name ~> focus
- output: Map[String, Command.Results] = Map.empty) // thread name ~> output messages
- {
+ output: Map[String, Command.Results] = Map.empty // thread name ~> output messages
+ ) {
def set_break(b: Boolean): State = copy(break = b)
def is_active: Boolean = active > 0
def inc_active: State = copy(active = active + 1)
def dec_active: State = copy(active = active - 1)
- def toggle_breakpoint(breakpoint: Long): (Boolean, State) =
- {
+ def toggle_breakpoint(breakpoint: Long): (Boolean, State) = {
val active_breakpoints1 =
if (active_breakpoints(breakpoint)) active_breakpoints - breakpoint
else active_breakpoints + breakpoint
@@ -75,8 +72,7 @@
def get_thread(thread_name: String): List[Debug_State] =
threads.getOrElse(thread_name, Nil)
- def update_thread(thread_name: String, debug_states: List[Debug_State]): State =
- {
+ def update_thread(thread_name: String, debug_states: List[Debug_State]): State = {
val threads1 =
if (debug_states.nonEmpty) threads + (thread_name -> debug_states)
else threads - thread_name
@@ -104,17 +100,14 @@
/* protocol handler */
- class Handler(session: Session) extends Session.Protocol_Handler
- {
+ class Handler(session: Session) extends Session.Protocol_Handler {
val debugger: Debugger = new Debugger(session)
- private def debugger_state(msg: Prover.Protocol_Output): Boolean =
- {
+ private def debugger_state(msg: Prover.Protocol_Output): Boolean = {
msg.properties match {
case Markup.Debugger_State(thread_name) =>
val msg_body = Symbol.decode_yxml_failsafe(msg.text)
- val debug_states =
- {
+ val debug_states = {
import XML.Decode._
list(pair(properties, string))(msg_body).map({
case (pos, function) => Debug_State(pos, function)
@@ -126,8 +119,7 @@
}
}
- private def debugger_output(msg: Prover.Protocol_Output): Boolean =
- {
+ private def debugger_output(msg: Prover.Protocol_Output): Boolean = {
msg.properties match {
case Markup.Debugger_Output(thread_name) =>
Symbol.decode_yxml_failsafe(msg.text) match {
@@ -148,8 +140,7 @@
}
}
-class Debugger private(session: Session)
-{
+class Debugger private(session: Session) {
/* debugger state */
private val state = Synchronized(Debugger.State())
@@ -162,29 +153,25 @@
/* protocol commands */
- def update_thread(thread_name: String, debug_states: List[Debugger.Debug_State]): Unit =
- {
+ def update_thread(thread_name: String, debug_states: List[Debugger.Debug_State]): Unit = {
state.change(_.update_thread(thread_name, debug_states))
delay_update.invoke()
}
- def add_output(thread_name: String, entry: Command.Results.Entry): Unit =
- {
+ def add_output(thread_name: String, entry: Command.Results.Entry): Unit = {
state.change(_.add_output(thread_name, entry))
delay_update.invoke()
}
def is_active(): Boolean = session.is_ready && state.value.is_active
- def ready(): Unit =
- {
+ def ready(): Unit = {
if (is_active())
session.protocol_command("Debugger.init")
}
def init(): Unit =
- state.change(st =>
- {
+ state.change(st => {
val st1 = st.inc_active
if (session.is_ready && !st.is_active && st1.is_active)
session.protocol_command("Debugger.init")
@@ -192,8 +179,7 @@
})
def exit(): Unit =
- state.change(st =>
- {
+ state.change(st => {
val st1 = st.dec_active
if (session.is_ready && st.is_active && !st1.is_active)
session.protocol_command("Debugger.exit")
@@ -201,8 +187,7 @@
})
def is_break(): Boolean = state.value.break
- def set_break(b: Boolean): Unit =
- {
+ def set_break(b: Boolean): Unit = {
state.change(st => {
val st1 = st.set_break(b)
session.protocol_command("Debugger.break", b.toString)
@@ -211,8 +196,7 @@
delay_update.invoke()
}
- def active_breakpoint_state(breakpoint: Long): Option[Boolean] =
- {
+ def active_breakpoint_state(breakpoint: Long): Option[Boolean] = {
val st = state.value
if (st.is_active) Some(st.active_breakpoints(breakpoint)) else None
}
@@ -220,23 +204,20 @@
def breakpoint_state(breakpoint: Long): Boolean =
state.value.active_breakpoints(breakpoint)
- def toggle_breakpoint(command: Command, breakpoint: Long): Unit =
- {
- state.change(st =>
- {
- val (breakpoint_state, st1) = st.toggle_breakpoint(breakpoint)
- session.protocol_command(
- "Debugger.breakpoint",
- Symbol.encode(command.node_name.node),
- Document_ID(command.id),
- Value.Long(breakpoint),
- Value.Boolean(breakpoint_state))
- st1
- })
+ def toggle_breakpoint(command: Command, breakpoint: Long): Unit = {
+ state.change(st => {
+ val (breakpoint_state, st1) = st.toggle_breakpoint(breakpoint)
+ session.protocol_command(
+ "Debugger.breakpoint",
+ Symbol.encode(command.node_name.node),
+ Document_ID(command.id),
+ Value.Long(breakpoint),
+ Value.Boolean(breakpoint_state))
+ st1
+ })
}
- def status(focus: Option[Debugger.Context]): (Debugger.Threads, List[XML.Tree]) =
- {
+ def status(focus: Option[Debugger.Context]): (Debugger.Threads, List[XML.Tree]) = {
val st = state.value
val output =
focus match {
@@ -252,8 +233,7 @@
}
def focus(): List[Debugger.Context] = state.value.focus.toList.map(_._2)
- def set_focus(c: Debugger.Context): Unit =
- {
+ def set_focus(c: Debugger.Context): Unit = {
state.change(_.set_focus(c))
delay_update.invoke()
}
@@ -266,14 +246,12 @@
def step_over(thread_name: String): Unit = input(thread_name, "step_over")
def step_out(thread_name: String): Unit = input(thread_name, "step_out")
- def clear_output(thread_name: String): Unit =
- {
+ def clear_output(thread_name: String): Unit = {
state.change(_.clear_output(thread_name))
delay_update.invoke()
}
- def eval(c: Debugger.Context, SML: Boolean, context: String, expression: String): Unit =
- {
+ def eval(c: Debugger.Context, SML: Boolean, context: String, expression: String): Unit = {
state.change(st => {
input(c.thread_name, "eval", c.debug_index.getOrElse(0).toString,
SML.toString, Symbol.encode(context), Symbol.encode(expression))
@@ -282,8 +260,7 @@
delay_update.invoke()
}
- def print_vals(c: Debugger.Context, SML: Boolean, context: String): Unit =
- {
+ def print_vals(c: Debugger.Context, SML: Boolean, context: String): Unit = {
require(c.debug_index.isDefined)
state.change(st => {
--- a/src/Pure/Tools/doc.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/doc.scala Fri Apr 01 17:06:10 2022 +0200
@@ -10,8 +10,7 @@
import scala.collection.mutable
-object Doc
-{
+object Doc {
/* dirs */
def dirs(): List[Path] =
@@ -28,23 +27,20 @@
line <- Library.trim_split_lines(File.read(catalog))
} yield (dir, line)
- object Contents
- {
+ object Contents {
def apply(sections: List[Section]): Contents = new Contents(sections)
def section(title: String, important: Boolean, entries: List[Entry]): Contents =
apply(List(Section(title, important, entries)))
}
- class Contents private(val sections: List[Section])
- {
+ class Contents private(val sections: List[Section]) {
def ++ (other: Contents): Contents = new Contents(sections ::: other.sections)
def entries: List[Entry] = sections.flatMap(_.entries)
def docs: List[Doc] = entries.collect({ case doc: Doc => doc })
}
case class Section(title: String, important: Boolean, entries: List[Entry])
- sealed abstract class Entry
- {
+ sealed abstract class Entry {
def name: String
def path: Path
}
@@ -71,8 +67,7 @@
case None => error("Bad entry in ISABELLE_DOCS_EXAMPLES: " + file)
}))
- def main_contents(): Contents =
- {
+ def main_contents(): Contents = {
val result = new mutable.ListBuffer[Section]
val entries = new mutable.ListBuffer[Entry]
var section: Option[Section] = None
@@ -84,8 +79,7 @@
section = None
}
- def begin(s: Section): Unit =
- {
+ def begin(s: Section): Unit = {
flush()
section = Some(s)
}
@@ -110,13 +104,11 @@
Contents(result.toList)
}
- def contents(): Contents =
- {
+ def contents(): Contents = {
examples() ++ release_notes() ++ main_contents()
}
- object Doc_Names extends Scala.Fun_String("doc_names")
- {
+ object Doc_Names extends Scala.Fun_String("doc_names") {
val here = Scala_Project.here
def apply(arg: String): String =
if (arg.nonEmpty) error("Bad argument: " + quote(arg))
@@ -126,8 +118,7 @@
/* view */
- def view(path: Path): Unit =
- {
+ def view(path: Path): Unit = {
if (!path.is_file) error("Bad Isabelle documentation file: " + path)
else if (path.is_pdf) Isabelle_System.pdf_viewer(path)
else Output.writeln(Library.trim_line(File.read(path)), stdout = true)
@@ -137,8 +128,7 @@
/* Isabelle tool wrapper */
val isabelle_tool = Isabelle_Tool("doc", "view Isabelle PDF documentation",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
val getopts = Getopts("""
Usage: isabelle doc [DOC ...]
--- a/src/Pure/Tools/dump.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/dump.scala Fri Apr 01 17:06:10 2022 +0200
@@ -9,8 +9,7 @@
import java.io.{BufferedWriter, FileOutputStream, OutputStreamWriter}
-object Dump
-{
+object Dump {
/* aspects */
sealed case class Aspect_Args(
@@ -19,10 +18,9 @@
progress: Progress,
output_dir: Path,
snapshot: Document.Snapshot,
- status: Document_Status.Node_Status)
- {
- def write_path(file_name: Path): Path =
- {
+ status: Document_Status.Node_Status
+ ) {
+ def write_path(file_name: Path): Path = {
val path = output_dir + Path.basic(snapshot.node_name.theory) + file_name
Isabelle_System.make_directory(path.dir)
path
@@ -39,9 +37,12 @@
writer => YXML.traversal(s => writer.write(Symbol.encode(s)), body))
}
- sealed case class Aspect(name: String, description: String, operation: Aspect_Args => Unit,
- options: List[String] = Nil)
- {
+ sealed case class Aspect(
+ name: String,
+ description: String,
+ operation: Aspect_Args => Unit,
+ options: List[String] = Nil
+ ) {
override def toString: String = name
}
@@ -79,13 +80,12 @@
sealed case class Args(
session: Headless.Session,
snapshot: Document.Snapshot,
- status: Document_Status.Node_Status)
- {
+ status: Document_Status.Node_Status
+ ) {
def print_node: String = snapshot.node_name.toString
}
- object Context
- {
+ object Context {
def apply(
options: Options,
aspects: List[Aspect] = Nil,
@@ -94,10 +94,9 @@
select_dirs: List[Path] = Nil,
selection: Sessions.Selection = Sessions.Selection.empty,
pure_base: Boolean = false,
- skip_base: Boolean = false): Context =
- {
- val session_options: Options =
- {
+ skip_base: Boolean = false
+ ): Context = {
+ val session_options: Options = {
val options0 = if (NUMA.enabled) NUMA.policy_options(options) else options
val options1 =
options0 +
@@ -132,22 +131,21 @@
val pure_base: Boolean,
val skip_base: Boolean,
val session_options: Options,
- val deps: Sessions.Deps)
- {
+ val deps: Sessions.Deps
+ ) {
context =>
def session_dirs: List[Path] = dirs ::: select_dirs
- def build_logic(logic: String): Unit =
- {
+ def build_logic(logic: String): Unit = {
Build.build_logic(options, logic, build_heap = true, progress = progress,
dirs = session_dirs, strict = true)
}
def sessions(
logic: String = default_logic,
- log: Logger = No_Logger): List[Session] =
- {
+ log: Logger = No_Logger
+ ): List[Session] = {
/* partitions */
def session_info(session_name: String): Sessions.Info =
@@ -176,8 +174,8 @@
selected_sessions: List[String],
session_logic: String = logic,
strict: Boolean = false,
- record_proofs: Boolean = false): List[Session] =
- {
+ record_proofs: Boolean = false
+ ): List[Session] = {
if (selected_sessions.isEmpty && !strict) Nil
else List(new Session(context, session_logic, log, selected_sessions, record_proofs))
}
@@ -200,8 +198,7 @@
val afp =
if (afp_sessions.isEmpty) Nil
else {
- val (part1, part2) =
- {
+ val (part1, part2) = {
val graph = session_graph.restrict(afp_sessions -- afp_bulky_sessions)
val force_partition1 = AFP.force_partition1.filter(graph.defined)
val force_part1 = graph.all_preds(graph.all_succs(force_partition1)).toSet
@@ -226,13 +223,11 @@
private val errors = Synchronized(List.empty[String])
- def add_errors(more_errs: List[String]): Unit =
- {
+ def add_errors(more_errs: List[String]): Unit = {
errors.change(errs => errs ::: more_errs)
}
- def check_errors: Unit =
- {
+ def check_errors: Unit = {
val errs = errors.value
if (errs.nonEmpty) error(errs.mkString("\n\n"))
}
@@ -243,8 +238,8 @@
val logic: String,
log: Logger,
selected_sessions: List[String],
- record_proofs: Boolean)
- {
+ record_proofs: Boolean
+ ) {
/* resources */
val options: Options =
@@ -259,8 +254,7 @@
session_dirs = context.session_dirs,
include_sessions = deps.sessions_structure.imports_topological_order)
- val used_theories: List[Document.Node.Name] =
- {
+ val used_theories: List[Document.Node.Name] = {
for {
session_name <-
deps.sessions_structure.build_graph.restrict(selected_sessions.toSet).topological_order
@@ -289,15 +283,13 @@
/* process */
- def process(process_theory: Args => Unit, unicode_symbols: Boolean = false): Unit =
- {
+ def process(process_theory: Args => Unit, unicode_symbols: Boolean = false): Unit = {
val session = resources.start_session(progress = progress)
// asynchronous consumer
- object Consumer
- {
+ object Consumer {
sealed case class Bad_Theory(
name: Document.Node.Name,
status: Document_Status.Node_Status,
@@ -307,8 +299,7 @@
private val consumer =
Consumer_Thread.fork(name = "dump")(
- consume = (args: (Document.Snapshot, Document_Status.Node_Status)) =>
- {
+ consume = (args: (Document.Snapshot, Document_Status.Node_Status)) => {
val (snapshot, status) = args
val name = snapshot.node_name
if (status.ok) {
@@ -342,8 +333,7 @@
def apply(snapshot: Document.Snapshot, status: Document_Status.Node_Status): Unit =
consumer.send((snapshot, status))
- def shutdown(): List[Bad_Theory] =
- {
+ def shutdown(): List[Bad_Theory] = {
consumer.shutdown()
consumer_bad_theories.value.reverse
}
@@ -394,8 +384,8 @@
dirs: List[Path] = Nil,
select_dirs: List[Path] = Nil,
output_dir: Path = default_output_dir,
- selection: Sessions.Selection = Sessions.Selection.empty): Unit =
- {
+ selection: Sessions.Selection = Sessions.Selection.empty
+ ): Unit = {
val context =
Context(options, aspects = aspects, progress = progress, dirs = dirs,
select_dirs = select_dirs, selection = selection)
@@ -403,8 +393,7 @@
context.build_logic(logic)
for (session <- context.sessions(logic = logic, log = log)) {
- session.process((args: Args) =>
- {
+ session.process((args: Args) => {
progress.echo("Processing theory " + args.print_node + " ...")
val aspect_args =
Aspect_Args(session.options, context.deps, progress, output_dir,
@@ -420,8 +409,8 @@
/* Isabelle tool wrapper */
val isabelle_tool =
- Isabelle_Tool("dump", "dump cumulative PIDE session database", Scala_Project.here, args =>
- {
+ Isabelle_Tool("dump", "dump cumulative PIDE session database", Scala_Project.here,
+ args => {
var aspects: List[Aspect] = known_aspects
var base_sessions: List[String] = Nil
var select_dirs: List[Path] = Nil
--- a/src/Pure/Tools/flarum.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/flarum.scala Fri Apr 01 17:06:10 2022 +0200
@@ -10,15 +10,13 @@
import java.net.URL
-object Flarum
-{
+object Flarum {
// see RFC3339 in https://www.php.net/manual/en/class.datetimeinterface.php
val Date_Format: Date.Format = Date.Format("uuuu-MM-dd'T'HH:mm:ssxxx")
def server(url: URL): Server = new Server(url)
- class Server private[Flarum](url: URL) extends JSON_API.Service(url)
- {
+ class Server private[Flarum](url: URL) extends JSON_API.Service(url) {
def get_api(route: String): JSON_API.Root =
get_root("api" + (if (route.isEmpty) "" else "/" + route))
--- a/src/Pure/Tools/fontforge.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/fontforge.scala Fri Apr 01 17:06:10 2022 +0200
@@ -12,8 +12,7 @@
import java.util.Locale
-object Fontforge
-{
+object Fontforge {
/** scripting language **/
type Script = String
@@ -21,16 +20,14 @@
/* concrete syntax */
- def string(s: String): Script =
- {
+ def string(s: String): Script = {
def err(c: Char): Nothing =
error("Bad character \\u" + String.format(Locale.ROOT, "%04x", Integer.valueOf(c)) +
" in fontforge string " + quote(s))
val q = if (s.contains('"')) '\'' else '"'
- def escape(c: Char): String =
- {
+ def escape(c: Char): String = {
if (c == '\u0000' || c == '\r' || c == q) err(c)
else if (c == '\n') "\\n"
else if (c == '\\') "\\\\"
@@ -70,15 +67,13 @@
/** execute fontforge program **/
def execute(script: Script, args: String = "", cwd: JFile = null): Process_Result =
- Isabelle_System.with_tmp_file("fontforge")(script_file =>
- {
+ Isabelle_System.with_tmp_file("fontforge")(script_file => {
File.write(script_file, script)
Isabelle_System.bash(File.bash_path(Path.explode("$ISABELLE_FONTFORGE")) +
" -lang=ff -script " + File.bash_path(script_file) + " " + args)
})
- def font_domain(path: Path, strict: Boolean = false): List[Int] =
- {
+ def font_domain(path: Path, strict: Boolean = false): List[Int] = {
val script =
commands(
open(path),
@@ -100,14 +95,13 @@
fullname: String = "",
weight: String = "",
copyright: String = "",
- fontversion: String = "")
- {
+ fontversion: String = ""
+ ) {
override def toString: String = fontname
def ttf: Path = Path.explode(fontname).ext("ttf")
- def update(prefix: String = "", version: String = ""): Font_Names =
- {
+ def update(prefix: String = "", version: String = ""): Font_Names = {
val fontversion1 = proper_string(version) getOrElse fontversion
if (prefix == "") copy(fontversion = fontversion1)
else {
@@ -125,8 +119,7 @@
mkString("SetFontNames(", ",", ")")
}
- def font_names(path: Path): Font_Names =
- {
+ def font_names(path: Path): Font_Names = {
val script =
commands(
open(path),
--- a/src/Pure/Tools/java_monitor.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/java_monitor.scala Fri Apr 01 17:06:10 2022 +0200
@@ -13,8 +13,7 @@
Messages, Resources => JConsoleResources}
-object Java_Monitor
-{
+object Java_Monitor {
/* default arguments */
def default_pid: Long = ProcessHandle.current().pid
@@ -26,8 +25,8 @@
def java_monitor_internal(
pid: Long = default_pid,
look_and_feel: String = "",
- update_interval: Time = default_update_interval): Unit =
- {
+ update_interval: Time = default_update_interval
+ ): Unit = {
val vm =
if (pid.toInt.toLong == pid) LocalVirtualMachine.getLocalVirtualMachine(pid.toInt)
else null
@@ -102,8 +101,8 @@
parent: Component,
pid: Long = default_pid,
look_and_feel: String = "",
- update_interval: Time = default_update_interval): Unit =
- {
+ update_interval: Time = default_update_interval
+ ): Unit = {
Future.thread(name = "java_monitor", daemon = true) {
try {
Isabelle_System.bash(
@@ -124,8 +123,7 @@
/* command line entry point */
- def main(args: Array[String]): Unit =
- {
+ def main(args: Array[String]): Unit = {
Command_Line.tool {
var look_and_feel = ""
var pid = 0L
--- a/src/Pure/Tools/logo.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/logo.scala Fri Apr 01 17:06:10 2022 +0200
@@ -7,20 +7,16 @@
package isabelle
-object Logo
-{
+object Logo {
/* create logo */
- def make_output_file(logo_name: String): Path =
- {
+ def make_output_file(logo_name: String): Path = {
val name = if (logo_name.isEmpty) "isabelle" else "isabelle_" + Word.lowercase(logo_name)
Path.explode(name).pdf
}
- def create_logo(logo_name: String, output_file: Path, quiet: Boolean = false): Unit =
- {
- Isabelle_System.with_tmp_file("logo", ext = "eps")(tmp_file =>
- {
+ def create_logo(logo_name: String, output_file: Path, quiet: Boolean = false): Unit = {
+ Isabelle_System.with_tmp_file("logo", ext = "eps")(tmp_file => {
val template = File.read(Path.explode("$ISABELLE_HOME/lib/logo/isabelle_any.eps"))
File.write(tmp_file, template.replace("<any>", logo_name))
@@ -35,8 +31,8 @@
/* Isabelle tool wrapper */
val isabelle_tool =
- Isabelle_Tool("logo", "create variants of the Isabelle logo (PDF)", Scala_Project.here, args =>
- {
+ Isabelle_Tool("logo", "create variants of the Isabelle logo (PDF)", Scala_Project.here,
+ args => {
var output: Option[Path] = None
var quiet = false
--- a/src/Pure/Tools/mkroot.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/mkroot.scala Fri Apr 01 17:06:10 2022 +0200
@@ -7,8 +7,7 @@
package isabelle
-object Mkroot
-{
+object Mkroot {
/** mkroot **/
def root_name(name: String): String =
@@ -25,8 +24,8 @@
init_repos: Boolean = false,
title: String = "",
author: String = "",
- progress: Progress = new Progress): Unit =
- {
+ progress: Progress = new Progress
+ ): Unit = {
Isabelle_System.make_directory(session_dir)
val name = proper_string(session_name) getOrElse session_dir.absolute_file.getName
@@ -178,8 +177,8 @@
/** Isabelle tool wrapper **/
val isabelle_tool = Isabelle_Tool("mkroot", "prepare session root directory",
- Scala_Project.here, args =>
- {
+ Scala_Project.here,
+ args => {
var author = ""
var init_repos = false
var title = ""
--- a/src/Pure/Tools/phabricator.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/phabricator.scala Fri Apr 01 17:06:10 2022 +0200
@@ -15,8 +15,7 @@
import scala.util.matching.Regex
-object Phabricator
-{
+object Phabricator {
/** defaults **/
/* required packages */
@@ -32,8 +31,7 @@
// mercurial build packages
"make", "gcc", "python", "python2-dev", "python-docutils", "python-openssl")
- def packages: List[String] =
- {
+ def packages: List[String] = {
val release = Linux.Release()
if (release.is_ubuntu_20_04) packages_ubuntu_20_04
else error("Bad Linux version: expected Ubuntu 20.04 LTS")
@@ -81,8 +79,7 @@
def global_config_script(
init: String = "",
body: String = "",
- exit: String = ""): String =
- {
+ exit: String = ""): String = {
"""#!/bin/bash
""" + (if (init.nonEmpty) "\n" + init else "") + """
{
@@ -98,16 +95,14 @@
(if (exit.nonEmpty) "\n" + exit + "\n" else "")
}
- sealed case class Config(name: String, root: Path)
- {
+ sealed case class Config(name: String, root: Path) {
def home: Path = root + Path.explode(phabricator_name())
def execute(command: String): Process_Result =
Isabelle_System.bash("bin/" + command, cwd = home.file, redirect = true).check
}
- def read_config(): List[Config] =
- {
+ def read_config(): List[Config] = {
if (global_config.is_file) {
for (entry <- Library.trim_split_lines(File.read(global_config)) if entry.nonEmpty)
yield {
@@ -120,8 +115,7 @@
else Nil
}
- def write_config(configs: List[Config]): Unit =
- {
+ def write_config(configs: List[Config]): Unit = {
File.write(global_config,
configs.map(config => config.name + ":" + config.root.implode).mkString("", "\n", "\n"))
}
@@ -138,8 +132,8 @@
val isabelle_tool1 =
Isabelle_Tool("phabricator", "invoke command-line tool within Phabricator home directory",
- Scala_Project.here, args =>
- {
+ Scala_Project.here,
+ args => {
var list = false
var name = default_name
@@ -178,8 +172,7 @@
/** setup **/
- def user_setup(name: String, description: String, ssh_setup: Boolean = false): Unit =
- {
+ def user_setup(name: String, description: String, ssh_setup: Boolean = false): Unit = {
if (!Linux.user_exists(name)) {
Linux.user_add(name, description = description, system = true, ssh_setup = ssh_setup)
}
@@ -192,8 +185,8 @@
def command_setup(name: String,
init: String = "",
body: String = "",
- exit: String = ""): Path =
- {
+ exit: String = ""
+ ): Path = {
val command = Path.explode("/usr/local/bin") + Path.basic(name)
File.write(command, global_config_script(init = init, body = body, exit = exit))
Isabelle_System.chmod("755", command)
@@ -201,11 +194,9 @@
command
}
- def mercurial_setup(mercurial_source: String, progress: Progress = new Progress): Unit =
- {
+ def mercurial_setup(mercurial_source: String, progress: Progress = new Progress): Unit = {
progress.echo("\nMercurial installation from source " + quote(mercurial_source) + " ...")
- Isabelle_System.with_tmp_dir("mercurial")(tmp_dir =>
- {
+ Isabelle_System.with_tmp_dir("mercurial")(tmp_dir => {
val archive =
if (Url.is_wellformed(mercurial_source)) {
val archive = tmp_dir + Path.basic("mercurial.tar.gz")
@@ -228,8 +219,8 @@
repo: String = "",
package_update: Boolean = false,
mercurial_source: String = "",
- progress: Progress = new Progress): Unit =
- {
+ progress: Progress = new Progress
+ ): Unit = {
/* system environment */
Linux.check_system_root()
@@ -345,8 +336,7 @@
Linux.service_restart("mysql")
- def mysql_conf(R: Regex, which: String): String =
- {
+ def mysql_conf(R: Regex, which: String): String = {
val conf = Path.explode("/etc/mysql/debian.cnf")
split_lines(File.read(conf)).collectFirst({ case R(a) => a }) match {
case Some(res) => res
@@ -528,8 +518,8 @@
val isabelle_tool2 =
Isabelle_Tool("phabricator_setup", "setup Phabricator server on Ubuntu Linux",
- Scala_Project.here, args =>
- {
+ Scala_Project.here,
+ args => {
var mercurial_source = ""
var repo = ""
var package_update = false
@@ -595,8 +585,8 @@
name: String = default_name,
config_file: Option[Path] = None,
test_user: String = "",
- progress: Progress = new Progress): Unit =
- {
+ progress: Progress = new Progress
+ ): Unit = {
Linux.check_system_root()
val config = get_config(name)
@@ -604,8 +594,7 @@
val mail_config = config_file getOrElse default_config_file
- def setup_mail: Unit =
- {
+ def setup_mail: Unit = {
progress.echo("Using mail configuration from " + mail_config)
config.execute("config set cluster.mailers --stdin < " + File.bash_path(mail_config))
@@ -636,8 +625,8 @@
val isabelle_tool3 =
Isabelle_Tool("phabricator_setup_mail", "setup mail for one Phabricator installation",
- Scala_Project.here, args =>
- {
+ Scala_Project.here,
+ args => {
var test_user = ""
var name = default_name
var config_file: Option[Path] = None
@@ -679,8 +668,7 @@
def conf_ssh_port(port: Int): String =
if (port == SSH.default_port) "#Port " + SSH.default_port else "Port " + port
- def read_ssh_port(conf: Path): Int =
- {
+ def read_ssh_port(conf: Path): Int = {
val lines = split_lines(File.read(conf))
val ports =
lines.flatMap({
@@ -695,8 +683,7 @@
}
}
- def write_ssh_port(conf: Path, port: Int): Boolean =
- {
+ def write_ssh_port(conf: Path, port: Int): Boolean = {
val old_port = read_ssh_port(conf)
if (old_port == port) false
else {
@@ -713,8 +700,8 @@
def phabricator_setup_ssh(
server_port: Int = default_server_port,
system_port: Int = default_system_port,
- progress: Progress = new Progress): Unit =
- {
+ progress: Progress = new Progress
+ ): Unit = {
Linux.check_system_root()
val configs = read_config()
@@ -799,8 +786,8 @@
val isabelle_tool4 =
Isabelle_Tool("phabricator_setup_ssh", "setup ssh service for all Phabricator installations",
- Scala_Project.here, args =>
- {
+ Scala_Project.here,
+ args => {
var server_port = default_server_port
var system_port = default_system_port
@@ -836,8 +823,7 @@
/** conduit API **/
- object API
- {
+ object API {
/* user information */
sealed case class User(
@@ -845,8 +831,8 @@
phid: String,
name: String,
real_name: String,
- roles: List[String])
- {
+ roles: List[String]
+ ) {
def is_valid: Boolean =
roles.contains("verified") &&
roles.contains("approved") &&
@@ -866,13 +852,12 @@
callsign: String,
short_name: String,
importing: Boolean,
- ssh_url: String)
- {
+ ssh_url: String
+ ) {
def is_hg: Boolean = vcs == VCS.hg
}
- object VCS extends Enumeration
- {
+ object VCS extends Enumeration {
val hg, git, svn = Value
def read(s: String): Value =
try { withName(s) }
@@ -888,8 +873,7 @@
/* result with optional error */
- sealed case class Result(result: JSON.T, error: Option[String])
- {
+ sealed case class Result(result: JSON.T, error: Option[String]) {
def ok: Boolean = error.isEmpty
def get: JSON.T = if (ok) result else Exn.error(error.get)
@@ -899,8 +883,7 @@
def get_string: String = get_value(JSON.Value.String.unapply)
}
- def make_result(json: JSON.T): Result =
- {
+ def make_result(json: JSON.T): Result = {
val result = JSON.value(json, "result").getOrElse(JSON.Object.empty)
val error_info = JSON.string(json, "error_info")
val error_code = JSON.string(json, "error_code")
@@ -914,8 +897,7 @@
new API(user, host, port)
}
- final class API private(ssh_user: String, ssh_host: String, ssh_port: Int)
- {
+ final class API private(ssh_user: String, ssh_host: String, ssh_port: Int) {
/* connection */
require(ssh_host.nonEmpty && ssh_port >= 0, "bad ssh host or port")
@@ -929,10 +911,8 @@
/* execute methods */
- def execute_raw(method: String, params: JSON.T = JSON.Object.empty): JSON.T =
- {
- Isabelle_System.with_tmp_file("params", "json")(params_file =>
- {
+ def execute_raw(method: String, params: JSON.T = JSON.Object.empty): JSON.T = {
+ Isabelle_System.with_tmp_file("params", "json")(params_file => {
File.write(params_file, JSON.Format(JSON.Object("params" -> JSON.Format(params))))
val result =
Isabelle_System.bash(
@@ -946,8 +926,10 @@
API.make_result(execute_raw(method, params = params))
def execute_search[A](
- method: String, params: JSON.Object.T, unapply: JSON.T => Option[A]): List[A] =
- {
+ method: String,
+ params: JSON.Object.T,
+ unapply: JSON.T => Option[A]
+ ): List[A] = {
val results = new mutable.ListBuffer[A]
var after = ""
@@ -974,8 +956,8 @@
def get_users(
all: Boolean = false,
phid: String = "",
- name: String = ""): List[API.User] =
- {
+ name: String = ""
+ ): List[API.User] = {
val constraints: JSON.Object.T =
(for { (key, value) <- List("phids" -> phid, "usernames" -> name) if value.nonEmpty }
yield (key, List(value))).toMap
@@ -1005,8 +987,8 @@
all: Boolean = false,
phid: String = "",
callsign: String = "",
- short_name: String = ""): List[API.Repository] =
- {
+ short_name: String = ""
+ ): List[API.Repository] = {
val constraints: JSON.Object.T =
(for {
(key, value) <- List("phids" -> phid, "callsigns" -> callsign, "shortNames" -> short_name)
@@ -1051,8 +1033,8 @@
short_name: String = "", // unique name
description: String = "",
public: Boolean = false,
- vcs: API.VCS.Value = API.VCS.hg): API.Repository =
- {
+ vcs: API.VCS.Value = API.VCS.hg
+ ): API.Repository = {
require(name.nonEmpty, "bad repository name")
val transactions =
--- a/src/Pure/Tools/print_operation.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/print_operation.scala Fri Apr 01 17:06:10 2022 +0200
@@ -7,8 +7,7 @@
package isabelle
-object Print_Operation
-{
+object Print_Operation {
def get(session: Session): List[(String, String)] =
session.get_protocol_handler(classOf[Handler]) match {
case Some(h) => h.get
@@ -18,8 +17,7 @@
/* protocol handler */
- class Handler extends Session.Protocol_Handler
- {
+ class Handler extends Session.Protocol_Handler {
private val print_operations = Synchronized[List[(String, String)]](Nil)
def get: List[(String, String)] = print_operations.value
@@ -27,10 +25,8 @@
override def init(session: Session): Unit =
session.protocol_command(Markup.PRINT_OPERATIONS)
- private def put(msg: Prover.Protocol_Output): Boolean =
- {
- val ops =
- {
+ private def put(msg: Prover.Protocol_Output): Boolean = {
+ val ops = {
import XML.Decode._
list(pair(string, string))(Symbol.decode_yxml(msg.text))
}
--- a/src/Pure/Tools/profiling_report.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/profiling_report.scala Fri Apr 01 17:06:10 2022 +0200
@@ -7,19 +7,17 @@
package isabelle
-object Profiling_Report
-{
+object Profiling_Report {
def profiling_report(
options: Options,
session: String,
theories: List[String] = Nil,
clean_name: Boolean = false,
- progress: Progress = new Progress): Unit =
- {
+ progress: Progress = new Progress
+ ): Unit = {
val store = Sessions.store(options)
- using(store.open_database_context())(db_context =>
- {
+ using(store.open_database_context())(db_context => {
val result =
db_context.input_database(session)((db, name) => Some(store.read_theories(db, name)))
result match {
@@ -48,8 +46,8 @@
val isabelle_tool =
Isabelle_Tool("profiling_report", "report Poly/ML profiling information from log files",
- Scala_Project.here, args =>
- {
+ Scala_Project.here,
+ args => {
var theories: List[String] = Nil
var clean_name = false
var options = Options.init()
--- a/src/Pure/Tools/scala_build.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/scala_build.scala Fri Apr 01 17:06:10 2022 +0200
@@ -14,20 +14,16 @@
import scala.jdk.CollectionConverters._
-object Scala_Build
-{
- class Context private[Scala_Build](java_context: isabelle.setup.Build.Context)
- {
+object Scala_Build {
+ class Context private[Scala_Build](java_context: isabelle.setup.Build.Context) {
override def toString: String = java_context.toString
- def is_module(path: Path): Boolean =
- {
+ def is_module(path: Path): Boolean = {
val module_name = java_context.module_name()
module_name.nonEmpty && File.eq(java_context.path(module_name).toFile, path.file)
}
- def module_result: Option[Path] =
- {
+ def module_result: Option[Path] = {
java_context.module_result() match {
case "" => None
case module => Some(File.path(java_context.path(module).toFile))
@@ -43,12 +39,10 @@
p <- java_context.requirement_paths(s).asScala.iterator
} yield (File.path(p.toFile))).toList
- def build(fresh: Boolean = false): String =
- {
+ def build(fresh: Boolean = false): String = {
val output0 = new ByteArrayOutputStream
val output = new PrintStream(output0)
- def get_output(): String =
- {
+ def get_output(): String = {
output.flush()
Library.trim_line(output0.toString(UTF8.charset))
}
@@ -68,8 +62,8 @@
component: Boolean = false,
no_title: Boolean = false,
do_build: Boolean = false,
- module: Option[Path] = None): Context =
- {
+ module: Option[Path] = None
+ ): Context = {
val props_name =
if (component) isabelle.setup.Build.COMPONENT_BUILD_PROPS
else isabelle.setup.Build.BUILD_PROPS
@@ -89,16 +83,14 @@
component: Boolean = false,
no_title: Boolean = false,
do_build: Boolean = false,
- module: Option[Path] = None): String =
- {
+ module: Option[Path] = None
+ ): String = {
context(dir, component = component, no_title = no_title, do_build = do_build, module = module)
.build(fresh = fresh)
}
- sealed case class Result(output: String, jar_bytes: Bytes, jar_path: Option[Path])
- {
- def write(): Unit =
- {
+ sealed case class Result(output: String, jar_bytes: Bytes, jar_path: Option[Path]) {
+ def write(): Unit = {
if (jar_path.isDefined) {
Isabelle_System.make_directory(jar_path.get.dir)
Bytes.write(jar_path.get, jar_bytes)
@@ -106,10 +98,8 @@
}
}
- def build_result(dir: Path, component: Boolean = false): Result =
- {
- Isabelle_System.with_tmp_file("result", "jar")(tmp_file =>
- {
+ def build_result(dir: Path, component: Boolean = false): Result = {
+ Isabelle_System.with_tmp_file("result", "jar")(tmp_file => {
val output =
build(dir, component = component, no_title = true, do_build = true, module = Some(tmp_file))
val jar_bytes = Bytes.read(tmp_file)
--- a/src/Pure/Tools/scala_project.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/scala_project.scala Fri Apr 01 17:06:10 2022 +0200
@@ -8,15 +8,13 @@
package isabelle
-object Scala_Project
-{
+object Scala_Project {
/** build tools **/
def java_version: String = "15"
def scala_version: String = scala.util.Properties.versionNumberString
- abstract class Build_Tool
- {
+ abstract class Build_Tool {
def project_root: Path
def init_project(dir: Path, jars: List[Path]): Unit
@@ -27,8 +25,7 @@
(dir + project_root).is_file &&
(dir + scala_src_dir).is_dir
- def package_dir(source_file: Path): Path =
- {
+ def package_dir(source_file: Path): Path = {
val dir =
package_name(source_file) match {
case Some(name) => Path.explode(space_explode('.', name).mkString("/"))
@@ -43,15 +40,13 @@
/* Gradle */
- object Gradle extends Build_Tool
- {
+ object Gradle extends Build_Tool {
override def toString: String = "Gradle"
val project_settings: Path = Path.explode("settings.gradle")
override val project_root: Path = Path.explode("build.gradle")
- private def groovy_string(s: String): String =
- {
+ private def groovy_string(s: String): String = {
s.map(c =>
c match {
case '\t' | '\b' | '\n' | '\r' | '\f' | '\\' | '\'' | '"' => "\\" + c
@@ -59,8 +54,7 @@
}).mkString("'", "", "'")
}
- override def init_project(dir: Path, jars: List[Path]): Unit =
- {
+ override def init_project(dir: Path, jars: List[Path]): Unit = {
File.write(dir + project_settings, "rootProject.name = 'Isabelle'\n")
File.write(dir + project_root,
"""plugins {
@@ -84,16 +78,13 @@
/* Maven */
- object Maven extends Build_Tool
- {
+ object Maven extends Build_Tool {
override def toString: String = "Maven"
override val project_root: Path = Path.explode("pom.xml")
- override def init_project(dir: Path, jars: List[Path]): Unit =
- {
- def dependency(jar: Path): String =
- {
+ override def init_project(dir: Path, jars: List[Path]): Unit = {
+ def dependency(jar: Path): String = {
val name = jar.expand.drop_ext.base.implode
val system_path = File.platform_path(jar.absolute)
""" <dependency>
@@ -143,8 +134,7 @@
/* plugins: modules with dynamic build */
- class Plugin(dir: Path) extends Isabelle_System.Service
- {
+ class Plugin(dir: Path) extends Isabelle_System.Service {
def context(): Scala_Build.Context = Scala_Build.context(dir)
}
@@ -153,8 +143,7 @@
/* file and directories */
- lazy val isabelle_files: (List[Path], List[Path]) =
- {
+ lazy val isabelle_files: (List[Path], List[Path]) = {
val contexts = Scala_Build.component_contexts() ::: plugins.map(_.context())
val jars1 = Path.split(Isabelle_System.getenv("ISABELLE_CLASSPATH"))
@@ -173,8 +162,7 @@
(jars, sources)
}
- lazy val isabelle_scala_files: Map[String, Path] =
- {
+ lazy val isabelle_scala_files: Map[String, Path] = {
val context = Scala_Build.context(Path.ISABELLE_HOME, component = true)
context.sources.iterator.foldLeft(Map.empty[String, Path]) {
case (map, path) =>
@@ -192,8 +180,7 @@
/* compile-time position */
- def here: Here =
- {
+ def here: Here = {
val exn = new Exception
exn.getStackTrace.toList match {
case _ :: caller :: _ =>
@@ -204,8 +191,7 @@
}
}
- class Here private[Scala_Project](name: String, line: Int)
- {
+ class Here private[Scala_Project](name: String, line: Int) {
override def toString: String = name + ":" + line
def position: Position.T =
isabelle_scala_files.get(name) match {
@@ -219,8 +205,7 @@
val default_project_dir = Path.explode("$ISABELLE_HOME_USER/scala_project")
- def package_name(source_file: Path): Option[String] =
- {
+ def package_name(source_file: Path): Option[String] = {
val lines = Library.trim_split_lines(File.read(source_file))
val Package = """\s*\bpackage\b\s*(?:object\b\s*)?((?:\w|\.)+)\b.*""".r
lines.collectFirst({ case Package(name) => name })
@@ -232,8 +217,8 @@
more_sources: List[Path] = Nil,
symlinks: Boolean = false,
force: Boolean = false,
- progress: Progress = new Progress): Unit =
- {
+ progress: Progress = new Progress
+ ): Unit = {
if (project_dir.file.exists) {
val detect = project_dir.is_dir && build_tools.exists(_.detect_project(project_dir))
@@ -272,8 +257,8 @@
val isabelle_tool =
Isabelle_Tool("scala_project", "setup IDE project for Isabelle/Java/Scala sources",
- Scala_Project.here, args =>
- {
+ Scala_Project.here,
+ args => {
var build_tool: Option[Build_Tool] = None
var project_dir = default_project_dir
var symlinks = false
--- a/src/Pure/Tools/server.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/server.scala Fri Apr 01 17:06:10 2022 +0200
@@ -24,17 +24,14 @@
import java.net.{Socket, SocketException, SocketTimeoutException, ServerSocket, InetAddress}
-object Server
-{
+object Server {
/* message argument */
- object Argument
- {
+ object Argument {
def is_name_char(c: Char): Boolean =
Symbol.is_ascii_letter(c) || Symbol.is_ascii_digit(c) || c == '_' || c == '.'
- def split(msg: String): (String, String) =
- {
+ def split(msg: String): (String, String) = {
val name = msg.takeWhile(is_name_char)
val argument = msg.substring(name.length).dropWhile(Symbol.is_ascii_blank)
(name, argument)
@@ -62,14 +59,12 @@
type Command_Body = PartialFunction[(Context, Any), Any]
- abstract class Command(val command_name: String)
- {
+ abstract class Command(val command_name: String) {
def command_body: Command_Body
override def toString: String = command_name
}
- class Commands(commands: Command*) extends Isabelle_System.Service
- {
+ class Commands(commands: Command*) extends Isabelle_System.Service {
def entries: List[Command] = commands.toList
}
@@ -96,8 +91,7 @@
case _ => JSON.Object.empty
}
- object Reply extends Enumeration
- {
+ object Reply extends Enumeration {
val OK, ERROR, FINISHED, FAILED, NOTE = Value
def message(msg: String, kind: String = ""): JSON.Object.T =
@@ -106,8 +100,7 @@
def error_message(msg: String): JSON.Object.T =
message(msg, kind = Markup.ERROR)
- def unapply(msg: String): Option[(Reply.Value, Any)] =
- {
+ def unapply(msg: String): Option[(Reply.Value, Any)] = {
if (msg == "") None
else {
val (name, argument) = Argument.split(msg)
@@ -124,8 +117,7 @@
/* handler: port, password, thread */
- abstract class Handler(port0: Int)
- {
+ abstract class Handler(port0: Int) {
val socket: ServerSocket = new ServerSocket(port0, 50, Server.localhost)
def port: Int = socket.getLocalPort
def address: String = print_address(port)
@@ -158,14 +150,12 @@
/* socket connection */
- object Connection
- {
+ object Connection {
def apply(socket: Socket): Connection =
new Connection(socket)
}
- class Connection private(socket: Socket) extends AutoCloseable
- {
+ class Connection private(socket: Socket) extends AutoCloseable {
override def toString: String = socket.toString
def close(): Unit = socket.close()
@@ -204,8 +194,7 @@
def write_byte_message(chunks: List[Bytes]): Unit =
out_lock.synchronized { Byte_Message.write_message(out, chunks) }
- def reply(r: Reply.Value, arg: Any): Unit =
- {
+ def reply(r: Reply.Value, arg: Any): Unit = {
val argument = Argument.print(arg)
write_line_message(if (argument == "") r.toString else r.toString + " " + argument)
}
@@ -222,8 +211,7 @@
/* context with output channels */
class Context private[Server](val server: Server, connection: Connection)
- extends AutoCloseable
- {
+ extends AutoCloseable {
context =>
def command_list: List[String] = command_table.keys.toList.sorted
@@ -247,8 +235,7 @@
private val _tasks = Synchronized(Set.empty[Task])
- def make_task(body: Task => JSON.Object.T): Task =
- {
+ def make_task(body: Task => JSON.Object.T): Task = {
val task = new Task(context, body)
_tasks.change(_ + task)
task
@@ -260,30 +247,27 @@
def cancel_task(id: UUID.T): Unit =
_tasks.change(tasks => { tasks.find(task => task.id == id).foreach(_.cancel()); tasks })
- def close(): Unit =
- {
- while(_tasks.change_result(tasks => { tasks.foreach(_.cancel()); (tasks.nonEmpty, tasks) }))
- { _tasks.value.foreach(_.join()) }
+ def close(): Unit = {
+ while(_tasks.change_result(tasks => { tasks.foreach(_.cancel()); (tasks.nonEmpty, tasks) })) {
+ _tasks.value.foreach(_.join())
+ }
}
}
class Connection_Progress private[Server](context: Context, more: JSON.Object.Entry*)
- extends Progress
- {
+ extends Progress {
override def echo(msg: String): Unit = context.writeln(msg, more:_*)
override def echo_warning(msg: String): Unit = context.warning(msg, more:_*)
override def echo_error_message(msg: String): Unit = context.error_message(msg, more:_*)
- override def theory(theory: Progress.Theory): Unit =
- {
+ override def theory(theory: Progress.Theory): Unit = {
val entries: List[JSON.Object.Entry] =
List("theory" -> theory.theory, "session" -> theory.session) :::
(theory.percentage match { case None => Nil case Some(p) => List("percentage" -> p) })
context.writeln(theory.message, entries ::: more.toList:_*)
}
- override def nodes_status(nodes_status: Document_Status.Nodes_Status): Unit =
- {
+ override def nodes_status(nodes_status: Document_Status.Nodes_Status): Unit = {
val json =
for ((name, node_status) <- nodes_status.present)
yield name.json + ("status" -> node_status.json)
@@ -293,8 +277,7 @@
override def toString: String = context.toString
}
- class Task private[Server](val context: Context, body: Task => JSON.Object.T)
- {
+ class Task private[Server](val context: Context, body: Task => JSON.Object.T) {
task =>
val id: UUID.T = UUID.random()
@@ -303,8 +286,7 @@
val progress: Connection_Progress = context.progress(ident)
def cancel(): Unit = progress.stop()
- private lazy val thread = Isabelle_Thread.fork(name = "server_task")
- {
+ private lazy val thread = Isabelle_Thread.fork(name = "server_task") {
Exn.capture { body(task) } match {
case Exn.Res(res) =>
context.reply(Reply.FINISHED, res + ident)
@@ -330,8 +312,7 @@
def print(port: Int, password: String): String =
print_address(port) + " (password " + quote(password) + ")"
- object Info
- {
+ object Info {
private val Pattern =
("""server "([^"]*)" = \Q""" + localhost_name + """\E:(\d+) \(password "([^"]*)"\)""").r
@@ -345,15 +326,13 @@
new Info(name, port, password)
}
- class Info private(val name: String, val port: Int, val password: String)
- {
+ class Info private(val name: String, val port: Int, val password: String) {
def address: String = print_address(port)
override def toString: String =
"server " + quote(name) + " = " + print(port, password)
- def connection(): Connection =
- {
+ def connection(): Connection = {
val connection = Connection(new Socket(localhost, port))
connection.write_line_message(password)
connection
@@ -361,14 +340,13 @@
def active: Boolean =
try {
- using(connection())(connection =>
- {
- connection.set_timeout(Time.seconds(2.0))
- connection.read_line_message() match {
- case Some(Reply(Reply.OK, _)) => true
- case _ => false
- }
- })
+ using(connection())(connection => {
+ connection.set_timeout(Time.seconds(2.0))
+ connection.read_line_message() match {
+ case Some(Reply(Reply.OK, _)) => true
+ case _ => false
+ }
+ })
}
catch {
case _: IOException => false
@@ -382,8 +360,7 @@
val default_name = "isabelle"
- object Data
- {
+ object Data {
val database = Path.explode("$ISABELLE_HOME_USER/servers.db")
val name = SQL.Column.string("name").make_primary_key
@@ -410,10 +387,9 @@
name: String = default_name,
port: Int = 0,
existing_server: Boolean = false,
- log: Logger = No_Logger): (Info, Option[Server]) =
- {
- using(SQLite.open_database(Data.database))(db =>
- {
+ log: Logger = No_Logger
+ ): (Info, Option[Server]) = {
+ using(SQLite.open_database(Data.database))(db => {
db.transaction {
Isabelle_System.chmod("600", Data.database)
db.create_table(Data.table)
@@ -431,8 +407,7 @@
val server_info = Info(name, server.port, server.password)
db.using_statement(Data.table.delete(Data.name.where_equal(name)))(_.execute())
- db.using_statement(Data.table.insert())(stmt =>
- {
+ db.using_statement(Data.table.insert())(stmt => {
stmt.string(1) = server_info.name
stmt.int(2) = server_info.port
stmt.string(3) = server_info.password
@@ -446,8 +421,7 @@
})
}
- def exit(name: String = default_name): Boolean =
- {
+ def exit(name: String = default_name): Boolean = {
using(SQLite.open_database(Data.database))(db =>
db.transaction {
find(db, name) match {
@@ -464,8 +438,8 @@
/* Isabelle tool wrapper */
val isabelle_tool =
- Isabelle_Tool("server", "manage resident Isabelle servers", Scala_Project.here, args =>
- {
+ Isabelle_Tool("server", "manage resident Isabelle servers", Scala_Project.here,
+ args => {
var console = false
var log_file: Option[Path] = None
var operation_list = false
@@ -523,16 +497,14 @@
})
}
-class Server private(port0: Int, val log: Logger) extends Server.Handler(port0)
-{
+class Server private(port0: Int, val log: Logger) extends Server.Handler(port0) {
server =>
private val _sessions = Synchronized(Map.empty[UUID.T, Headless.Session])
def err_session(id: UUID.T): Nothing = error("No session " + Library.single_quote(id.toString))
def the_session(id: UUID.T): Headless.Session = _sessions.value.getOrElse(id, err_session(id))
def add_session(entry: (UUID.T, Headless.Session)): Unit = _sessions.change(_ + entry)
- def remove_session(id: UUID.T): Headless.Session =
- {
+ def remove_session(id: UUID.T): Headless.Session = {
_sessions.change_result(sessions =>
sessions.get(id) match {
case Some(session) => (session, sessions - id)
@@ -540,8 +512,7 @@
})
}
- def shutdown(): Unit =
- {
+ def shutdown(): Unit = {
server.socket.close()
val sessions = _sessions.change_result(sessions => (sessions, Map.empty))
@@ -556,10 +527,8 @@
override def join(): Unit = { super.join(); shutdown() }
- override def handle(connection: Server.Connection): Unit =
- {
- using(new Server.Context(server, connection))(context =>
- {
+ override def handle(connection: Server.Connection): Unit = {
+ using(new Server.Context(server, connection))(context => {
connection.reply_ok(
JSON.Object(
"isabelle_id" -> Isabelle_System.isabelle_id(),
--- a/src/Pure/Tools/server_commands.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/server_commands.scala Fri Apr 01 17:06:10 2022 +0200
@@ -7,24 +7,20 @@
package isabelle
-object Server_Commands
-{
+object Server_Commands {
def default_preferences: String = Options.read_prefs()
- object Help extends Server.Command("help")
- {
+ object Help extends Server.Command("help") {
override val command_body: Server.Command_Body =
{ case (context, ()) => context.command_list }
}
- object Echo extends Server.Command("echo")
- {
+ object Echo extends Server.Command("echo") {
override val command_body: Server.Command_Body =
{ case (_, t) => t }
}
- object Cancel extends Server.Command("cancel")
- {
+ object Cancel extends Server.Command("cancel") {
sealed case class Args(task: UUID.T)
def unapply(json: JSON.T): Option[Args] =
@@ -35,14 +31,12 @@
{ case (context, Cancel(args)) => context.cancel_task(args.task) }
}
- object Shutdown extends Server.Command("shutdown")
- {
+ object Shutdown extends Server.Command("shutdown") {
override val command_body: Server.Command_Body =
{ case (context, ()) => context.server.shutdown() }
}
- object Session_Build extends Server.Command("session_build")
- {
+ object Session_Build extends Server.Command("session_build") {
sealed case class Args(
session: String,
preferences: String = default_preferences,
@@ -65,9 +59,10 @@
include_sessions = include_sessions, verbose = verbose)
}
- def command(args: Args, progress: Progress = new Progress)
- : (JSON.Object.T, Build.Results, Options, Sessions.Base_Info) =
- {
+ def command(
+ args: Args,
+ progress: Progress = new Progress
+ ) : (JSON.Object.T, Build.Results, Options, Sessions.Base_Info) = {
val options = Options.init(prefs = args.preferences, opts = args.options)
val dirs = args.dirs.map(Path.explode)
@@ -93,16 +88,15 @@
"ok" -> results.ok,
"return_code" -> results.rc,
"sessions" ->
- results.sessions.toList.sortBy(sessions_order).map(session =>
- {
- val result = results(session)
- JSON.Object(
- "session" -> session,
- "ok" -> result.ok,
- "return_code" -> result.rc,
- "timeout" -> result.timeout,
- "timing" -> result.timing.json)
- }))
+ results.sessions.toList.sortBy(sessions_order).map(session => {
+ val result = results(session)
+ JSON.Object(
+ "session" -> session,
+ "ok" -> result.ok,
+ "return_code" -> result.rc,
+ "timeout" -> result.timeout,
+ "timing" -> result.timing.json)
+ }))
if (results.ok) (results_json, results, options, base_info)
else {
@@ -116,8 +110,7 @@
context.make_task(task => Session_Build.command(args, progress = task.progress)._1) }
}
- object Session_Start extends Server.Command("session_start")
- {
+ object Session_Start extends Server.Command("session_start") {
sealed case class Args(
build: Session_Build.Args,
print_mode: List[String] = Nil)
@@ -129,9 +122,11 @@
}
yield Args(build = build, print_mode = print_mode)
- def command(args: Args, progress: Progress = new Progress, log: Logger = No_Logger)
- : (JSON.Object.T, (UUID.T, Headless.Session)) =
- {
+ def command(
+ args: Args,
+ progress: Progress = new Progress,
+ log: Logger = No_Logger
+ ) : (JSON.Object.T, (UUID.T, Headless.Session)) = {
val (_, _, options, base_info) =
try { Session_Build.command(args.build, progress = progress) }
catch { case exn: Server.Error => error(exn.message) }
@@ -152,8 +147,7 @@
override val command_body: Server.Command_Body =
{ case (context, Session_Start(args)) =>
- context.make_task(task =>
- {
+ context.make_task(task => {
val (res, entry) =
Session_Start.command(args, progress = task.progress, log = context.server.log)
context.server.add_session(entry)
@@ -162,13 +156,11 @@
}
}
- object Session_Stop extends Server.Command("session_stop")
- {
+ object Session_Stop extends Server.Command("session_stop") {
def unapply(json: JSON.T): Option[UUID.T] =
JSON.uuid(json, "session_id")
- def command(session: Headless.Session): (JSON.Object.T, Process_Result) =
- {
+ def command(session: Headless.Session): (JSON.Object.T, Process_Result) = {
val result = session.stop()
val result_json = JSON.Object("ok" -> result.ok, "return_code" -> result.rc)
@@ -186,8 +178,7 @@
}
}
- object Use_Theories extends Server.Command("use_theories")
- {
+ object Use_Theories extends Server.Command("use_theories") {
sealed case class Args(
session_id: UUID.T,
theories: List[String],
@@ -225,8 +216,8 @@
def command(args: Args,
session: Headless.Session,
id: UUID.T = UUID.random(),
- progress: Progress = new Progress): (JSON.Object.T, Headless.Use_Theories_Result) =
- {
+ progress: Progress = new Progress
+ ): (JSON.Object.T, Headless.Use_Theories_Result) = {
val result =
session.use_theories(args.theories, master_dir = args.master_dir,
check_delay = args.check_delay.getOrElse(session.default_check_delay),
@@ -240,8 +231,7 @@
def output_text(text: String): String =
Symbol.output(args.unicode_symbols, text)
- def output_message(tree: XML.Tree, pos: Position.T): JSON.Object.T =
- {
+ def output_message(tree: XML.Tree, pos: Position.T): JSON.Object.T = {
val position = "pos" -> Position.JSON(pos)
tree match {
case XML.Text(msg) => Server.Reply.message(output_text(msg)) + position
@@ -287,16 +277,14 @@
override val command_body: Server.Command_Body =
{ case (context, Use_Theories(args)) =>
- context.make_task(task =>
- {
+ context.make_task(task => {
val session = context.server.the_session(args.session_id)
Use_Theories.command(args, session, id = task.id, progress = task.progress)._1
})
}
}
- object Purge_Theories extends Server.Command("purge_theories")
- {
+ object Purge_Theories extends Server.Command("purge_theories") {
sealed case class Args(
session_id: UUID.T,
theories: List[String] = Nil,
@@ -312,9 +300,10 @@
}
yield { Args(session_id, theories = theories, master_dir = master_dir, all = all) }
- def command(args: Args, session: Headless.Session)
- : (JSON.Object.T, (List[Document.Node.Name], List[Document.Node.Name])) =
- {
+ def command(
+ args: Args,
+ session: Headless.Session
+ ) : (JSON.Object.T, (List[Document.Node.Name], List[Document.Node.Name])) = {
val (purged, retained) =
session.purge_theories(
theories = args.theories, master_dir = args.master_dir, all = args.all)
--- a/src/Pure/Tools/simplifier_trace.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/simplifier_trace.scala Fri Apr 01 17:06:10 2022 +0200
@@ -11,8 +11,7 @@
import scala.collection.immutable.SortedMap
-object Simplifier_Trace
-{
+object Simplifier_Trace {
/* trace items from the prover */
val TEXT = "text"
@@ -27,12 +26,11 @@
val MEMORY = "memory"
val Memory = new Properties.Boolean(MEMORY)
- object Item
- {
+ object Item {
case class Data(
serial: Long, markup: String, text: String,
- parent: Long, props: Properties.T, content: XML.Body)
- {
+ parent: Long, props: Properties.T, content: XML.Body
+ ) {
def memory: Boolean = Memory.unapply(props) getOrElse true
}
@@ -55,10 +53,8 @@
case class Answer private[Simplifier_Trace](val name: String, val string: String)
- object Answer
- {
- object step
- {
+ object Answer {
+ object step {
val skip = Answer("skip", "Skip")
val continue = Answer("continue", "Continue")
val continue_trace = Answer("continue_trace", "Continue (with full trace)")
@@ -68,8 +64,7 @@
val all = List(continue, continue_trace, continue_passive, continue_disable, skip)
}
- object hint_fail
- {
+ object hint_fail {
val exit = Answer("exit", "Exit")
val redo = Answer("redo", "Redo")
@@ -98,8 +93,8 @@
case class Context(
last_serial: Long = 0L,
- questions: SortedMap[Long, Question] = SortedMap.empty)
- {
+ questions: SortedMap[Long, Question] = SortedMap.empty
+ ) {
def +(q: Question): Context =
copy(questions = questions + ((q.data.serial, q)))
@@ -114,21 +109,22 @@
case class Index(text: String, content: XML.Body)
- object Index
- {
+ object Index {
def of_data(data: Item.Data): Index =
Index(data.text, data.content)
}
- def handle_results(session: Session, id: Document_ID.Command, results: Command.Results): Context =
- {
+ def handle_results(
+ session: Session,
+ id: Document_ID.Command,
+ results: Command.Results
+ ): Context = {
val slot = Future.promise[Context]
the_manager(session).send(Handle_Results(session, id, results, slot))
slot.join
}
- def generate_trace(session: Session, results: Command.Results): Trace =
- {
+ def generate_trace(session: Session, results: Command.Results): Trace = {
val slot = Future.promise[Trace]
the_manager(session).send(Generate_Trace(results, slot))
slot.join
@@ -140,8 +136,7 @@
def send_reply(session: Session, serial: Long, answer: Answer): Unit =
the_manager(session).send(Reply(session, serial, answer))
- def make_manager: Consumer_Thread[Any] =
- {
+ def make_manager: Consumer_Thread[Any] = {
var contexts = Map.empty[Document_ID.Command, Context]
var memory_children = Map.empty[Long, Set[Long]]
@@ -153,8 +148,7 @@
(id, context.questions(serial))
}
- def do_cancel(serial: Long, id: Document_ID.Command): Unit =
- {
+ def do_cancel(serial: Long, id: Document_ID.Command): Unit = {
// To save memory, we could try to remove empty contexts at this point.
// However, if a new serial gets attached to the same command_id after we deleted
// its context, its last_serial counter will start at 0 again, and we'll think the
@@ -162,22 +156,19 @@
contexts += (id -> (contexts(id) - serial))
}
- def do_reply(session: Session, serial: Long, answer: Answer): Unit =
- {
+ def do_reply(session: Session, serial: Long, answer: Answer): Unit = {
session.protocol_command(
"Simplifier_Trace.reply", Value.Long(serial), answer.name)
}
Consumer_Thread.fork[Any]("Simplifier_Trace.manager", daemon = true)(
- consume = (arg: Any) =>
- {
+ consume = (arg: Any) => {
arg match {
case Handle_Results(session, id, results, slot) =>
var new_context = contexts.getOrElse(id, Context())
var new_serial = new_context.last_serial
- for ((serial, result) <- results.iterator if serial > new_context.last_serial)
- {
+ for ((serial, result) <- results.iterator if serial > new_context.last_serial) {
result match {
case Item(markup, data) =>
memory_children +=
@@ -264,8 +255,7 @@
case Reply(session, serial, answer) =>
find_question(serial) match {
case Some((id, Question(data, _))) =>
- if (data.markup == Markup.SIMP_TRACE_STEP && data.memory)
- {
+ if (data.markup == Markup.SIMP_TRACE_STEP && data.memory) {
val index = Index.of_data(data)
memory += (index -> answer)
}
@@ -294,19 +284,16 @@
/* protocol handler */
- class Handler extends Session.Protocol_Handler
- {
+ class Handler extends Session.Protocol_Handler {
private var the_session: Session = null
- override def init(session: Session): Unit =
- {
+ override def init(session: Session): Unit = {
try { the_manager(session) }
catch { case ERROR(_) => managers.change(map => map + (session -> make_manager)) }
the_session = session
}
- override def exit(): Unit =
- {
+ override def exit(): Unit = {
val session = the_session
if (session != null) {
val manager = the_manager(session)
--- a/src/Pure/Tools/spell_checker.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/spell_checker.scala Fri Apr 01 17:06:10 2022 +0200
@@ -16,21 +16,21 @@
import scala.collection.immutable.SortedMap
-object Spell_Checker
-{
+object Spell_Checker {
/* words within text */
- def marked_words(base: Text.Offset, text: String, mark: Text.Info[String] => Boolean)
- : List[Text.Info[String]] =
- {
+ def marked_words(
+ base: Text.Offset,
+ text: String,
+ mark: Text.Info[String] => Boolean
+ ) : List[Text.Info[String]] = {
val result = new mutable.ListBuffer[Text.Info[String]]
var offset = 0
def apostrophe(c: Int): Boolean =
c == '\'' && (offset + 1 == text.length || text(offset + 1) != '\'')
- @tailrec def scan(pred: Int => Boolean): Unit =
- {
+ @tailrec def scan(pred: Int => Boolean): Unit = {
if (offset < text.length) {
val c = text.codePointAt(offset)
if (pred(c)) {
@@ -53,8 +53,7 @@
result.toList
}
- def current_word(rendering: Rendering, range: Text.Range): Option[Text.Info[String]] =
- {
+ def current_word(rendering: Rendering, range: Text.Range): Option[Text.Info[String]] = {
for {
spell_range <- rendering.spell_checker_point(range)
text <- rendering.get_text(spell_range)
@@ -65,20 +64,17 @@
/* dictionaries */
- class Dictionary private[Spell_Checker](val path: Path)
- {
+ class Dictionary private[Spell_Checker](val path: Path) {
val lang: String = path.drop_ext.file_name
val user_path: Path = Path.explode("$ISABELLE_HOME_USER/dictionaries") + Path.basic(lang)
override def toString: String = lang
}
- private object Decl
- {
+ private object Decl {
def apply(name: String, include: Boolean): String =
if (include) name else "-" + name
- def unapply(decl: String): Option[(String, Boolean)] =
- {
+ def unapply(decl: String): Option[(String, Boolean)] = {
val decl1 = decl.trim
if (decl1 == "" || decl1.startsWith("#")) None
else
@@ -104,8 +100,7 @@
}
-class Spell_Checker private(dictionary: Spell_Checker.Dictionary)
-{
+class Spell_Checker private(dictionary: Spell_Checker.Dictionary) {
override def toString: String = dictionary.toString
@@ -126,8 +121,7 @@
case None => false
}
- private def load(): Unit =
- {
+ private def load(): Unit = {
val main_dictionary = split_lines(File.read_gzip(dictionary.path))
val permanent_updates =
@@ -155,8 +149,7 @@
}
load()
- private def save(): Unit =
- {
+ private def save(): Unit = {
val permanent_decls =
(for {
(word, upd) <- updates.iterator
@@ -179,8 +172,7 @@
}
}
- def update(word: String, include: Boolean, permanent: Boolean): Unit =
- {
+ def update(word: String, include: Boolean, permanent: Boolean): Unit = {
updates += (word -> Spell_Checker.Update(include, permanent))
if (include) {
@@ -190,8 +182,7 @@
else { save(); load() }
}
- def reset(): Unit =
- {
+ def reset(): Unit = {
updates = SortedMap.empty
load()
}
@@ -220,8 +211,7 @@
/* completion: suggestions for unknown words */
- private def suggestions(word: String): Option[List[String]] =
- {
+ private def suggestions(word: String): Option[List[String]] = {
val res =
Untyped.method(dict.getClass.getSuperclass, "searchSuggestions", classOf[String]).
invoke(dict, word).asInstanceOf[JList[AnyRef]].toArray.toList.map(_.toString)
@@ -247,8 +237,7 @@
result.getOrElse(Nil).map(recover_case)
}
- def completion(rendering: Rendering, caret: Text.Offset): Option[Completion.Result] =
- {
+ def completion(rendering: Rendering, caret: Text.Offset): Option[Completion.Result] = {
val caret_range = rendering.before_caret_range(caret)
for {
word <- Spell_Checker.current_word(rendering, caret_range)
@@ -261,8 +250,7 @@
}
}
-class Spell_Checker_Variable
-{
+class Spell_Checker_Variable {
private val no_spell_checker: (String, Option[Spell_Checker]) = ("", None)
private var current_spell_checker = no_spell_checker
--- a/src/Pure/Tools/task_statistics.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/task_statistics.scala Fri Apr 01 17:06:10 2022 +0200
@@ -15,20 +15,19 @@
import org.jfree.chart.renderer.xy.{XYBarRenderer, StandardXYBarPainter}
-object Task_Statistics
-{
+object Task_Statistics {
def apply(session_name: String, task_statistics: List[Properties.T]): Task_Statistics =
new Task_Statistics(session_name, task_statistics)
}
final class Task_Statistics private(
- val session_name: String, val task_statistics: List[Properties.T])
-{
+ val session_name: String,
+ val task_statistics: List[Properties.T]
+) {
private val Task_Name = new Properties.String("task_name")
private val Run = new Properties.Int("run")
- def chart(bins: Int = 100): JFreeChart =
- {
+ def chart(bins: Int = 100): JFreeChart = {
val values = new Array[Double](task_statistics.length)
for ((Run(x), i) <- task_statistics.iterator.zipWithIndex)
values(i) = java.lang.Math.log10((x max 1).toDouble / 1000000)
--- a/src/Pure/Tools/update.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/update.scala Fri Apr 01 17:06:10 2022 +0200
@@ -7,15 +7,14 @@
package isabelle
-object Update
-{
+object Update {
def update(options: Options, logic: String,
progress: Progress = new Progress,
log: Logger = No_Logger,
dirs: List[Path] = Nil,
select_dirs: List[Path] = Nil,
- selection: Sessions.Selection = Sessions.Selection.empty): Unit =
- {
+ selection: Sessions.Selection = Sessions.Selection.empty
+ ): Unit = {
val context =
Dump.Context(options, progress = progress, dirs = dirs, select_dirs = select_dirs,
selection = selection, skip_base = true)
@@ -38,24 +37,23 @@
case t => List(t)
}
- context.sessions(logic, log = log).foreach(_.process((args: Dump.Args) =>
- {
- progress.echo("Processing theory " + args.print_node + " ...")
+ context.sessions(logic, log = log).foreach(_.process((args: Dump.Args) => {
+ progress.echo("Processing theory " + args.print_node + " ...")
- val snapshot = args.snapshot
- for (node_name <- snapshot.node_files) {
- val node = snapshot.get_node(node_name)
- val xml =
- snapshot.state.xml_markup(snapshot.version, node_name,
- elements = Markup.Elements(Markup.UPDATE, Markup.LANGUAGE))
+ val snapshot = args.snapshot
+ for (node_name <- snapshot.node_files) {
+ val node = snapshot.get_node(node_name)
+ val xml =
+ snapshot.state.xml_markup(snapshot.version, node_name,
+ elements = Markup.Elements(Markup.UPDATE, Markup.LANGUAGE))
- val source1 = Symbol.encode(XML.content(update_xml(xml)))
- if (source1 != Symbol.encode(node.source)) {
- progress.echo("Updating " + node_name.path)
- File.write(node_name.path, source1)
- }
+ val source1 = Symbol.encode(XML.content(update_xml(xml)))
+ if (source1 != Symbol.encode(node.source)) {
+ progress.echo("Updating " + node_name.path)
+ File.write(node_name.path, source1)
}
- }))
+ }
+ }))
context.check_errors
}
@@ -65,8 +63,7 @@
val isabelle_tool =
Isabelle_Tool("update", "update theory sources based on PIDE markup", Scala_Project.here,
- args =>
- {
+ args => {
var base_sessions: List[String] = Nil
var select_dirs: List[Path] = Nil
var requirements = false
--- a/src/Pure/Tools/update_cartouches.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/update_cartouches.scala Fri Apr 01 17:06:10 2022 +0200
@@ -10,14 +10,12 @@
import scala.util.matching.Regex
-object Update_Cartouches
-{
+object Update_Cartouches {
/* update cartouches */
val Text_Antiq: Regex = """(?s)@\{\s*text\b\s*(.+)\}""".r
- def update_text(content: String): String =
- {
+ def update_text(content: String): String = {
(try { Some(Antiquote.read(content)) } catch { case ERROR(_) => None }) match {
case Some(ants) =>
val ants1: List[Antiquote.Antiquote] =
@@ -35,8 +33,7 @@
}
}
- def update_cartouches(replace_text: Boolean, path: Path): Unit =
- {
+ def update_cartouches(replace_text: Boolean, path: Path): Unit = {
val text0 = File.read(path)
// outer syntax cartouches
@@ -77,8 +74,8 @@
val isabelle_tool =
Isabelle_Tool("update_cartouches", "update theory syntax to use cartouches",
- Scala_Project.here, args =>
- {
+ Scala_Project.here,
+ args => {
var replace_text = false
val getopts = Getopts("""
--- a/src/Pure/Tools/update_comments.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/update_comments.scala Fri Apr 01 17:06:10 2022 +0200
@@ -10,15 +10,12 @@
import scala.annotation.tailrec
-object Update_Comments
-{
- def update_comments(path: Path): Unit =
- {
+object Update_Comments {
+ def update_comments(path: Path): Unit = {
def make_comment(tok: Token): String =
Symbol.comment + Symbol.space + Symbol.cartouche(Symbol.trim_blanks(tok.content))
- @tailrec def update(toks: List[Token], result: List[String]): String =
- {
+ @tailrec def update(toks: List[Token], result: List[String]): String = {
toks match {
case tok :: rest
if tok.source == "--" || tok.source == Symbol.comment =>
@@ -48,8 +45,8 @@
val isabelle_tool =
Isabelle_Tool("update_comments", "update formal comments in outer syntax",
- Scala_Project.here, args =>
- {
+ Scala_Project.here,
+ args => {
val getopts = Getopts("""
Usage: isabelle update_comments [FILES|DIRS...]
--- a/src/Pure/Tools/update_header.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/update_header.scala Fri Apr 01 17:06:10 2022 +0200
@@ -7,10 +7,8 @@
package isabelle
-object Update_Header
-{
- def update_header(section: String, path: Path): Unit =
- {
+object Update_Header {
+ def update_header(section: String, path: Path): Unit = {
val text0 = File.read(path)
val text1 =
(for (tok <- Token.explode(Keyword.Keywords.empty, text0).iterator)
@@ -30,8 +28,8 @@
val isabelle_tool =
Isabelle_Tool("update_header", "replace obsolete theory header command",
- Scala_Project.here, args =>
- {
+ Scala_Project.here,
+ args => {
var section = "section"
val getopts = Getopts("""
--- a/src/Pure/Tools/update_then.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/update_then.scala Fri Apr 01 17:06:10 2022 +0200
@@ -7,10 +7,8 @@
package isabelle
-object Update_Then
-{
- def update_then(path: Path): Unit =
- {
+object Update_Then {
+ def update_then(path: Path): Unit = {
val text0 = File.read(path)
val text1 =
(for (tok <- Token.explode(Keyword.Keywords.empty, text0).iterator)
@@ -32,8 +30,7 @@
val isabelle_tool =
Isabelle_Tool("update_then", "expand old Isar command conflations 'hence' and 'thus'",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
val getopts = Getopts("""
Usage: isabelle update_then [FILES|DIRS...]
--- a/src/Pure/Tools/update_theorems.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/Tools/update_theorems.scala Fri Apr 01 17:06:10 2022 +0200
@@ -7,10 +7,8 @@
package isabelle
-object Update_Theorems
-{
- def update_theorems(path: Path): Unit =
- {
+object Update_Theorems {
+ def update_theorems(path: Path): Unit = {
val text0 = File.read(path)
val text1 =
(for (tok <- Token.explode(Keyword.Keywords.empty, text0).iterator)
@@ -32,8 +30,7 @@
/* Isabelle tool wrapper */
val isabelle_tool = Isabelle_Tool("update_theorems", "update toplevel theorem keywords",
- Scala_Project.here, args =>
- {
+ Scala_Project.here, args => {
val getopts = Getopts("""
Usage: isabelle update_theorems [FILES|DIRS...]
--- a/src/Pure/library.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/library.scala Fri Apr 01 17:06:10 2022 +0200
@@ -12,12 +12,10 @@
import scala.util.matching.Regex
-object Library
-{
+object Library {
/* resource management */
- def using[A <: AutoCloseable, B](a: A)(f: A => B): B =
- {
+ def using[A <: AutoCloseable, B](a: A)(f: A => B): B = {
try { f(a) }
finally { if (a != null) a.close() }
}
@@ -26,15 +24,13 @@
/* integers */
private val small_int = 10000
- private lazy val small_int_table =
- {
+ private lazy val small_int_table = {
val array = new Array[String](small_int)
for (i <- 0 until small_int) array(i) = i.toString
array
}
- def is_small_int(s: String): Boolean =
- {
+ def is_small_int(s: String): Boolean = {
val len = s.length
1 <= len && len <= 4 &&
s.forall(c => '0' <= c && c <= '9') &&
@@ -52,8 +48,7 @@
/* separated chunks */
- def separate[A](s: A, list: List[A]): List[A] =
- {
+ def separate[A](s: A, list: List[A]): List[A] = {
val result = new mutable.ListBuffer[A]
var first = true
for (x <- list) {
@@ -72,8 +67,7 @@
def separated_chunks(sep: Char => Boolean, source: CharSequence): Iterator[CharSequence] =
new Iterator[CharSequence] {
private val end = source.length
- private def next_chunk(i: Int): Option[(CharSequence, Int)] =
- {
+ private def next_chunk(i: Int): Option[(CharSequence, Int)] = {
if (i < end) {
var j = i
var cont = true
@@ -115,8 +109,7 @@
def indent_lines(n: Int, str: String): String =
prefix_lines(Symbol.spaces(n), str)
- def first_line(source: CharSequence): String =
- {
+ def first_line(source: CharSequence): String = {
val lines = separated_chunks(_ == '\n', source)
if (lines.hasNext) lines.next().toString
else ""
@@ -134,8 +127,7 @@
/* strings */
- def make_string(f: StringBuilder => Unit, capacity: Int = 16): String =
- {
+ def make_string(f: StringBuilder => Unit, capacity: Int = 16): String = {
val s = new StringBuilder(capacity)
f(s)
s.toString
@@ -174,8 +166,7 @@
/* CharSequence */
- class Reverse(text: CharSequence, start: Int, end: Int) extends CharSequence
- {
+ class Reverse(text: CharSequence, start: Int, end: Int) extends CharSequence {
require(0 <= start && start <= end && end <= text.length, "bad reverse range")
def this(text: CharSequence) = this(text, 0, text.length)
@@ -187,8 +178,7 @@
if (0 <= i && i <= j && j <= length) new Reverse(text, end - j, end - i)
else throw new IndexOutOfBoundsException
- override def toString: String =
- {
+ override def toString: String = {
val buf = new StringBuilder(length)
for (i <- 0 until length)
buf.append(charAt(i))
@@ -196,8 +186,7 @@
}
}
- class Line_Termination(text: CharSequence) extends CharSequence
- {
+ class Line_Termination(text: CharSequence) extends CharSequence {
def length: Int = text.length + 1
def charAt(i: Int): Char = if (i == text.length) '\n' else text.charAt(i)
def subSequence(i: Int, j: Int): CharSequence =
@@ -227,8 +216,7 @@
def take_prefix[A](pred: A => Boolean, xs: List[A]): (List[A], List[A]) =
(xs.takeWhile(pred), xs.dropWhile(pred))
- def take_suffix[A](pred: A => Boolean, xs: List[A]): (List[A], List[A]) =
- {
+ def take_suffix[A](pred: A => Boolean, xs: List[A]): (List[A], List[A]) = {
val rev_xs = xs.reverse
(rev_xs.dropWhile(pred).reverse, rev_xs.takeWhile(pred).reverse)
}
@@ -246,15 +234,13 @@
else if (xs.isEmpty) ys
else ys.foldRight(xs)(Library.insert(_)(_))
- def distinct[A](xs: List[A], eq: (A, A) => Boolean = (x: A, y: A) => x == y): List[A] =
- {
+ def distinct[A](xs: List[A], eq: (A, A) => Boolean = (x: A, y: A) => x == y): List[A] = {
val result = new mutable.ListBuffer[A]
xs.foreach(x => if (!result.exists(y => eq(x, y))) result += x)
result.toList
}
- def duplicates[A](lst: List[A], eq: (A, A) => Boolean = (x: A, y: A) => x == y): List[A] =
- {
+ def duplicates[A](lst: List[A], eq: (A, A) => Boolean = (x: A, y: A) => x == y): List[A] = {
val result = new mutable.ListBuffer[A]
@tailrec def dups(rest: List[A]): Unit =
rest match {
@@ -297,16 +283,10 @@
/* reflection */
- def is_subclass[A, B](a: Class[A], b: Class[B]): Boolean =
- {
+ def is_subclass[A, B](a: Class[A], b: Class[B]): Boolean = {
import scala.language.existentials
- @tailrec def subclass(c: Class[_]): Boolean =
- {
- c == b ||
- {
- val d = c.getSuperclass
- d != null && subclass(d)
- }
+ @tailrec def subclass(c: Class[_]): Boolean = {
+ c == b || { val d = c.getSuperclass; d != null && subclass(d) }
}
subclass(a)
}
--- a/src/Pure/pure_thy.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/pure_thy.scala Fri Apr 01 17:06:10 2022 +0200
@@ -7,8 +7,7 @@
package isabelle
-object Pure_Thy
-{
+object Pure_Thy {
/* Pure logic */
val DUMMY: String = "dummy"
--- a/src/Pure/term.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/term.scala Fri Apr 01 17:06:10 2022 +0200
@@ -9,12 +9,10 @@
package isabelle
-object Term
-{
+object Term {
/* types and terms */
- sealed case class Indexname(name: String, index: Int = 0)
- {
+ sealed case class Indexname(name: String, index: Int = 0) {
private lazy val hash: Int = (name, index).hashCode()
override def hashCode(): Int = hash
@@ -37,8 +35,7 @@
type Sort = List[Class]
sealed abstract class Typ
- case class Type(name: String, args: List[Typ] = Nil) extends Typ
- {
+ case class Type(name: String, args: List[Typ] = Nil) extends Typ {
private lazy val hash: Int = ("Type", name, args).hashCode()
override def hashCode(): Int = hash
@@ -46,16 +43,14 @@
if (this == dummyT) "_"
else "Type(" + name + (if (args.isEmpty) "" else "," + args) + ")"
}
- case class TFree(name: String, sort: Sort = Nil) extends Typ
- {
+ case class TFree(name: String, sort: Sort = Nil) extends Typ {
private lazy val hash: Int = ("TFree", name, sort).hashCode()
override def hashCode(): Int = hash
override def toString: String =
"TFree(" + name + (if (sort.isEmpty) "" else "," + sort) + ")"
}
- case class TVar(name: Indexname, sort: Sort = Nil) extends Typ
- {
+ case class TVar(name: Indexname, sort: Sort = Nil) extends Typ {
private lazy val hash: Int = ("TVar", name, sort).hashCode()
override def hashCode(): Int = hash
@@ -64,16 +59,14 @@
}
val dummyT: Type = Type("dummy")
- sealed abstract class Term
- {
+ sealed abstract class Term {
def head: Term =
this match {
case App(fun, _) => fun.head
case _ => this
}
}
- case class Const(name: String, typargs: List[Typ] = Nil) extends Term
- {
+ case class Const(name: String, typargs: List[Typ] = Nil) extends Term {
private lazy val hash: Int = ("Const", name, typargs).hashCode()
override def hashCode(): Int = hash
@@ -81,16 +74,14 @@
if (this == dummy) "_"
else "Const(" + name + (if (typargs.isEmpty) "" else "," + typargs) + ")"
}
- case class Free(name: String, typ: Typ = dummyT) extends Term
- {
+ case class Free(name: String, typ: Typ = dummyT) extends Term {
private lazy val hash: Int = ("Free", name, typ).hashCode()
override def hashCode(): Int = hash
override def toString: String =
"Free(" + name + (if (typ == dummyT) "" else "," + typ) + ")"
}
- case class Var(name: Indexname, typ: Typ = dummyT) extends Term
- {
+ case class Var(name: Indexname, typ: Typ = dummyT) extends Term {
private lazy val hash: Int = ("Var", name, typ).hashCode()
override def hashCode(): Int = hash
@@ -98,13 +89,11 @@
"Var(" + name + (if (typ == dummyT) "" else "," + typ) + ")"
}
case class Bound(index: Int) extends Term
- case class Abs(name: String, typ: Typ, body: Term) extends Term
- {
+ case class Abs(name: String, typ: Typ, body: Term) extends Term {
private lazy val hash: Int = ("Abs", name, typ, body).hashCode()
override def hashCode(): Int = hash
}
- case class App(fun: Term, arg: Term) extends Term
- {
+ case class App(fun: Term, arg: Term) extends Term {
private lazy val hash: Int = ("App", fun, arg).hashCode()
override def hashCode(): Int = hash
@@ -121,48 +110,39 @@
sealed abstract class Proof
case object MinProof extends Proof
case class PBound(index: Int) extends Proof
- case class Abst(name: String, typ: Typ, body: Proof) extends Proof
- {
+ case class Abst(name: String, typ: Typ, body: Proof) extends Proof {
private lazy val hash: Int = ("Abst", name, typ, body).hashCode()
override def hashCode(): Int = hash
}
- case class AbsP(name: String, hyp: Term, body: Proof) extends Proof
- {
+ case class AbsP(name: String, hyp: Term, body: Proof) extends Proof {
private lazy val hash: Int = ("AbsP", name, hyp, body).hashCode()
override def hashCode(): Int = hash
}
- case class Appt(fun: Proof, arg: Term) extends Proof
- {
+ case class Appt(fun: Proof, arg: Term) extends Proof {
private lazy val hash: Int = ("Appt", fun, arg).hashCode()
override def hashCode(): Int = hash
}
- case class AppP(fun: Proof, arg: Proof) extends Proof
- {
+ case class AppP(fun: Proof, arg: Proof) extends Proof {
private lazy val hash: Int = ("AppP", fun, arg).hashCode()
override def hashCode(): Int = hash
}
- case class Hyp(hyp: Term) extends Proof
- {
+ case class Hyp(hyp: Term) extends Proof {
private lazy val hash: Int = ("Hyp", hyp).hashCode()
override def hashCode(): Int = hash
}
- case class PAxm(name: String, types: List[Typ]) extends Proof
- {
+ case class PAxm(name: String, types: List[Typ]) extends Proof {
private lazy val hash: Int = ("PAxm", name, types).hashCode()
override def hashCode(): Int = hash
}
- case class PClass(typ: Typ, cls: Class) extends Proof
- {
+ case class PClass(typ: Typ, cls: Class) extends Proof {
private lazy val hash: Int = ("PClass", typ, cls).hashCode()
override def hashCode(): Int = hash
}
- case class Oracle(name: String, prop: Term, types: List[Typ]) extends Proof
- {
+ case class Oracle(name: String, prop: Term, types: List[Typ]) extends Proof {
private lazy val hash: Int = ("Oracle", name, prop, types).hashCode()
override def hashCode(): Int = hash
}
- case class PThm(serial: Long, theory_name: String, name: String, types: List[Typ]) extends Proof
- {
+ case class PThm(serial: Long, theory_name: String, name: String, types: List[Typ]) extends Proof {
private lazy val hash: Int = ("PThm", serial, theory_name, name, types).hashCode()
override def hashCode(): Int = hash
}
@@ -170,16 +150,14 @@
/* type classes within the logic */
- object Class_Const
- {
+ object Class_Const {
val suffix = "_class"
def apply(c: Class): String = c + suffix
def unapply(s: String): Option[Class] =
if (s.endsWith(suffix)) Some(s.substring(0, s.length - suffix.length)) else None
}
- object OFCLASS
- {
+ object OFCLASS {
def apply(ty: Typ, s: Sort): List[Term] = s.map(c => apply(ty, c))
def apply(ty: Typ, c: Class): Term =
@@ -197,8 +175,7 @@
/** cache **/
- object Cache
- {
+ object Cache {
def make(
xz: XZ.Cache = XZ.Cache.make(),
max_string: Int = isabelle.Cache.default_max_string,
@@ -209,8 +186,7 @@
}
class Cache(xz: XZ.Cache, max_string: Int, initial_size: Int)
- extends XML.Cache(xz, max_string, initial_size)
- {
+ extends XML.Cache(xz, max_string, initial_size) {
protected def cache_indexname(x: Indexname): Indexname =
lookup(x) getOrElse store(Indexname(cache_string(x.name), x.index))
@@ -218,8 +194,7 @@
if (x.isEmpty) Nil
else lookup(x) getOrElse store(x.map(cache_string))
- protected def cache_typ(x: Typ): Typ =
- {
+ protected def cache_typ(x: Typ): Typ = {
if (x == dummyT) dummyT
else
lookup(x) match {
@@ -233,8 +208,7 @@
}
}
- protected def cache_typs(x: List[Typ]): List[Typ] =
- {
+ protected def cache_typs(x: List[Typ]): List[Typ] = {
if (x.isEmpty) Nil
else {
lookup(x) match {
@@ -244,8 +218,7 @@
}
}
- protected def cache_term(x: Term): Term =
- {
+ protected def cache_term(x: Term): Term = {
lookup(x) match {
case Some(y) => y
case None =>
@@ -261,8 +234,7 @@
}
}
- protected def cache_proof(x: Proof): Proof =
- {
+ protected def cache_proof(x: Proof): Proof = {
if (x == MinProof) MinProof
else {
lookup(x) match {
--- a/src/Pure/term_xml.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/term_xml.scala Fri Apr 01 17:06:10 2022 +0200
@@ -7,12 +7,10 @@
package isabelle
-object Term_XML
-{
+object Term_XML {
import Term._
- object Encode
- {
+ object Encode {
import XML.Encode._
val indexname: P[Indexname] =
@@ -39,8 +37,7 @@
{ case App(a, b) => (Nil, pair(term, term)(a, b)) }))
}
- object Decode
- {
+ object Decode {
import XML.Decode._
val indexname: P[Indexname] =
@@ -66,8 +63,7 @@
{ case (List(a), b) => val (c, d) = pair(typ, term)(b); Abs(a, c, d) },
{ case (Nil, a) => val (b, c) = pair(term, term)(a); App(b, c) }))
- def term_env(env: Map[String, Typ]): T[Term] =
- {
+ def term_env(env: Map[String, Typ]): T[Term] = {
def env_type(x: String, t: Typ): Typ =
if (t == dummyT && env.isDefinedAt(x)) env(x) else t
@@ -82,8 +78,7 @@
term
}
- def proof_env(env: Map[String, Typ]): T[Proof] =
- {
+ def proof_env(env: Map[String, Typ]): T[Proof] = {
val term = term_env(env)
def proof: T[Proof] =
variant[Proof](List(
--- a/src/Pure/thm_name.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Pure/thm_name.scala Fri Apr 01 17:06:10 2022 +0200
@@ -10,10 +10,8 @@
import scala.math.Ordering
-object Thm_Name
-{
- object Ordering extends scala.math.Ordering[Thm_Name]
- {
+object Thm_Name {
+ object Ordering extends scala.math.Ordering[Thm_Name] {
def compare(thm_name1: Thm_Name, thm_name2: Thm_Name): Int =
thm_name1.name compare thm_name2.name match {
case 0 => thm_name1.index compare thm_name2.index
@@ -32,8 +30,7 @@
}
}
-sealed case class Thm_Name(name: String, index: Int)
-{
+sealed case class Thm_Name(name: String, index: Int) {
def print: String =
if (name == "" || index == 0) name
else name + "(" + index + ")"
--- a/src/Tools/Graphview/graph_file.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/Graphview/graph_file.scala Fri Apr 01 17:06:10 2022 +0200
@@ -13,16 +13,13 @@
import java.awt.{Color, Graphics2D}
-object Graph_File
-{
- def write(file: JFile, graphview: Graphview): Unit =
- {
+object Graph_File {
+ def write(file: JFile, graphview: Graphview): Unit = {
val box = graphview.bounding_box()
val w = box.width.ceil.toInt
val h = box.height.ceil.toInt
- def paint(gfx: Graphics2D): Unit =
- {
+ def paint(gfx: Graphics2D): Unit = {
gfx.setColor(graphview.background_color)
gfx.fillRect(0, 0, w, h)
gfx.translate(- box.x, - box.y)
@@ -35,8 +32,7 @@
else error("Bad type of file: " + quote(name) + " (.png or .pdf expected)")
}
- def write(options: Options, file: JFile, graph: Graph_Display.Graph): Unit =
- {
+ def write(options: Options, file: JFile, graph: Graph_Display.Graph): Unit = {
val the_options = options
val graphview =
new Graphview(graph.transitive_reduction_acyclic) { def options = the_options }
@@ -46,8 +42,7 @@
}
def make_pdf(options: Options, graph: Graph_Display.Graph): Bytes =
- Isabelle_System.with_tmp_file("graph", ext = "pdf")(graph_file =>
- {
+ Isabelle_System.with_tmp_file("graph", ext = "pdf")(graph_file => {
write(options, graph_file.file, graph)
Bytes.read(graph_file)
})
--- a/src/Tools/Graphview/graph_panel.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/Graphview/graph_panel.scala Fri Apr 01 17:06:10 2022 +0200
@@ -20,8 +20,7 @@
import scala.swing.event.{Event, Key, MousePressed, MouseDragged, MouseClicked, MouseEvent}
-class Graph_Panel(val graphview: Graphview) extends BorderPanel
-{
+class Graph_Panel(val graphview: Graphview) extends BorderPanel {
graph_panel =>
@@ -30,10 +29,8 @@
/* painter */
- private val painter = new Panel
- {
- override def paint(gfx: Graphics2D): Unit =
- {
+ private val painter = new Panel {
+ override def paint(gfx: Graphics2D): Unit = {
super.paintComponent(gfx)
gfx.setColor(graphview.background_color)
@@ -44,24 +41,21 @@
}
}
- def set_preferred_size(): Unit =
- {
+ def set_preferred_size(): Unit = {
val box = graphview.bounding_box()
val s = Transform.scale_discrete
painter.preferredSize = new Dimension((box.width * s).ceil.toInt, (box.height * s).ceil.toInt)
painter.revalidate()
}
- def refresh(): Unit =
- {
+ def refresh(): Unit = {
if (painter != null) {
set_preferred_size()
painter.repaint()
}
}
- def scroll_to_node(node: Graph_Display.Node): Unit =
- {
+ def scroll_to_node(node: Graph_Display.Node): Unit = {
val gap = graphview.metrics.gap
val info = graphview.layout.get_node(node)
@@ -101,17 +95,17 @@
listenTo(mouse.moves)
listenTo(mouse.clicks)
reactions +=
- {
- case MousePressed(_, p, _, _, _) =>
- Mouse_Interaction.pressed(p)
- painter.repaint()
- case MouseDragged(_, to, _) =>
- Mouse_Interaction.dragged(to)
- painter.repaint()
- case e @ MouseClicked(_, p, m, n, _) =>
- Mouse_Interaction.clicked(p, m, n, SwingUtilities.isRightMouseButton(e.peer))
- painter.repaint()
- }
+ {
+ case MousePressed(_, p, _, _, _) =>
+ Mouse_Interaction.pressed(p)
+ painter.repaint()
+ case MouseDragged(_, to, _) =>
+ Mouse_Interaction.dragged(to)
+ painter.repaint()
+ case e @ MouseClicked(_, p, m, n, _) =>
+ Mouse_Interaction.clicked(p, m, n, SwingUtilities.isRightMouseButton(e.peer))
+ painter.repaint()
+ }
contents = painter
}
@@ -120,44 +114,37 @@
/* transform */
- def rescale(s: Double): Unit =
- {
+ def rescale(s: Double): Unit = {
Transform.scale = s
if (zoom != null) zoom.set_item((Transform.scale_discrete * 100).floor.toInt)
refresh()
}
- def fit_to_window(): Unit =
- {
+ def fit_to_window(): Unit = {
Transform.fit_to_window()
refresh()
}
- private object Transform
- {
+ private object Transform {
private var _scale: Double = 1.0
def scale: Double = _scale
- def scale_=(s: Double): Unit =
- {
+ def scale_=(s: Double): Unit = {
_scale = (s min 10.0) max 0.1
}
- def scale_discrete: Double =
- {
+ def scale_discrete: Double = {
val font_height = GUI.line_metrics(graphview.metrics.font).getHeight.toInt
(scale * font_height).floor / font_height
}
- def apply(): AffineTransform =
- {
+ def apply(): AffineTransform = {
val box = graphview.bounding_box()
val t = AffineTransform.getScaleInstance(scale_discrete, scale_discrete)
t.translate(- box.x, - box.y)
t
}
- def fit_to_window(): Unit =
- {
+ def fit_to_window(): Unit = {
if (graphview.visible_graph.is_empty)
rescale(1.0)
else {
@@ -166,8 +153,7 @@
}
}
- def pane_to_graph_coordinates(at: Point2D): Point2D =
- {
+ def pane_to_graph_coordinates(at: Point2D): Point2D = {
val s = Transform.scale_discrete
val p = Transform().inverseTransform(graph_pane.peer.getViewport.getViewPosition, null)
@@ -182,13 +168,11 @@
graphview.model.Colors.events += { case _ => painter.repaint() }
graphview.model.Mutators.events += { case _ => painter.repaint() }
- private object Mouse_Interaction
- {
+ private object Mouse_Interaction {
private var draginfo: (Point, List[Graph_Display.Node], List[Layout.Dummy]) =
(new Point(0, 0), Nil, Nil)
- def pressed(at: Point): Unit =
- {
+ def pressed(at: Point): Unit = {
val c = Transform.pane_to_graph_coordinates(at)
val l =
graphview.find_node(c) match {
@@ -205,8 +189,7 @@
draginfo = (at, l, d)
}
- def dragged(to: Point): Unit =
- {
+ def dragged(to: Point): Unit = {
val (from, p, d) = draginfo
val s = Transform.scale_discrete
@@ -229,8 +212,7 @@
draginfo = (to, p, d)
}
- def clicked(at: Point, m: Key.Modifiers, clicks: Int, right_click: Boolean): Unit =
- {
+ def clicked(at: Point, m: Key.Modifiers, clicks: Int, right_click: Boolean): Unit = {
val c = Transform.pane_to_graph_coordinates(at)
if (clicks < 2) {
--- a/src/Tools/Graphview/graphview.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/Graphview/graphview.scala Fri Apr 01 17:06:10 2022 +0200
@@ -15,8 +15,7 @@
import javax.swing.JComponent
-abstract class Graphview(full_graph: Graph_Display.Graph)
-{
+abstract class Graphview(full_graph: Graph_Display.Graph) {
graphview =>
@@ -47,8 +46,7 @@
case (dummy: Layout.Dummy, (info, _)) if Shapes.shape(info).contains(at) => dummy
})
- def bounding_box(): Rectangle2D.Double =
- {
+ def bounding_box(): Rectangle2D.Double = {
var x0 = 0.0
var y0 = 0.0
var x1 = 0.0
@@ -67,8 +65,7 @@
new Rectangle2D.Double(x0, y0, x1 - x0, y1 - y0)
}
- def update_layout(): Unit =
- {
+ def update_layout(): Unit = {
val metrics = Metrics(make_font())
val visible_graph = model.make_visible_graph()
@@ -103,8 +100,7 @@
/* font rendering information */
- def make_font(): Font =
- {
+ def make_font(): Font = {
val family = options.string("graphview_font_family")
val size = options.int("graphview_font_size")
new Font(family, Font.PLAIN, size)
@@ -119,8 +115,7 @@
var show_dummies = false
var editor_style = false
- object Colors
- {
+ object Colors {
private val filter_colors = List(
new Color(0xD9, 0xF2, 0xE2), // blue
new Color(0xFF, 0xE7, 0xD8), // orange
@@ -132,15 +127,13 @@
)
private var curr : Int = -1
- def next(): Color =
- {
+ def next(): Color = {
curr = (curr + 1) % filter_colors.length
filter_colors(curr)
}
}
- def paint(gfx: Graphics2D): Unit =
- {
+ def paint(gfx: Graphics2D): Unit = {
gfx.setRenderingHints(Metrics.rendering_hints)
for (node <- graphview.current_node)
@@ -155,8 +148,7 @@
var current_node: Option[Graph_Display.Node] = None
- object Selection
- {
+ object Selection {
private val state = Synchronized[List[Graph_Display.Node]](Nil)
def get(): List[Graph_Display.Node] = state.value
--- a/src/Tools/Graphview/layout.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/Graphview/layout.scala Fri Apr 01 17:06:10 2022 +0200
@@ -17,14 +17,11 @@
import isabelle._
-object Layout
-{
+object Layout {
/* graph structure */
- object Vertex
- {
- object Ordering extends scala.math.Ordering[Vertex]
- {
+ object Vertex {
+ object Ordering extends scala.math.Ordering[Vertex] {
def compare(v1: Vertex, v2: Vertex): Int =
(v1, v2) match {
case (Node(a), Node(b)) => Graph_Display.Node.Ordering.compare(a, b)
@@ -56,8 +53,12 @@
case class Dummy(node1: Graph_Display.Node, node2: Graph_Display.Node, index: Int) extends Vertex
sealed case class Info(
- x: Double, y: Double, width2: Double, height2: Double, lines: List[String])
- {
+ x: Double,
+ y: Double,
+ width2: Double,
+ height2: Double,
+ lines: List[String]
+ ) {
def left: Double = x - width2
def right: Double = x + width2
def width: Double = 2 * width2
@@ -89,10 +90,12 @@
private type Levels = Map[Vertex, Int]
private val empty_levels: Levels = Map.empty
- def make(options: Options, metrics: Metrics,
+ def make(
+ options: Options,
+ metrics: Metrics,
node_text: (Graph_Display.Node, XML.Body) => String,
- input_graph: Graph_Display.Graph): Layout =
- {
+ input_graph: Graph_Display.Graph
+ ): Layout = {
if (input_graph.is_empty) empty
else {
/* initial graph */
@@ -187,11 +190,12 @@
private type Level = List[Vertex]
private def minimize_crossings(
- options: Options, graph: Graph, levels: List[Level]): List[Level] =
- {
+ options: Options,
+ graph: Graph,
+ levels: List[Level]
+ ): List[Level] = {
def resort(parent: Level, child: Level, top_down: Boolean): Level =
- child.map(v =>
- {
+ child.map(v => {
val ps = if (top_down) graph.imm_preds(v) else graph.imm_succs(v)
val weight =
ps.foldLeft(0.0) { case (w, p) => w + (0 max parent.indexOf(p)) } / (ps.size max 1)
@@ -222,8 +226,7 @@
}._1
}
- private def level_list(levels: Levels): List[Level] =
- {
+ private def level_list(levels: Levels): List[Level] = {
val max_lev = levels.foldLeft(-1) { case (m, (_, l)) => m max l }
val buckets = new Array[Level](max_lev + 1)
for (l <- 0 to max_lev) { buckets(l) = Nil }
@@ -251,8 +254,7 @@
/*This is an auxiliary class which is used by the layout algorithm when
calculating coordinates with the "pendulum method". A "region" is a
group of vertices which "stick together".*/
- private class Region(val content: List[Vertex])
- {
+ private class Region(val content: List[Vertex]) {
def distance(metrics: Metrics, graph: Graph, that: Region): Double =
vertex_left(graph, that.content.head) -
vertex_right(graph, this.content.last) -
@@ -272,8 +274,11 @@
}
private def pendulum(
- options: Options, metrics: Metrics, levels: List[Level], levels_graph: Graph): Graph =
- {
+ options: Options,
+ metrics: Metrics,
+ levels: List[Level],
+ levels_graph: Graph
+ ): Graph = {
def combine_regions(graph: Graph, top_down: Boolean, level: List[Region]): List[Region] =
level match {
case r1 :: rest =>
@@ -293,8 +298,7 @@
case _ => level
}
- def deflect(level: List[Region], top_down: Boolean, graph: Graph): (Graph, Boolean) =
- {
+ def deflect(level: List[Region], top_down: Boolean, graph: Graph): (Graph, Boolean) = {
level.indices.foldLeft((graph, false)) {
case ((graph, moved), i) =>
val r = level(i)
@@ -336,8 +340,7 @@
/** rubberband method **/
- private def force_weight(graph: Graph, v: Vertex): Double =
- {
+ private def force_weight(graph: Graph, v: Vertex): Double = {
val preds = graph.imm_preds(v)
val succs = graph.imm_succs(v)
val n = preds.size + succs.size
@@ -349,8 +352,11 @@
}
private def rubberband(
- options: Options, metrics: Metrics, levels: List[Level], graph: Graph): Graph =
- {
+ options: Options,
+ metrics: Metrics,
+ levels: List[Level],
+ graph: Graph
+ ): Graph = {
val gap = metrics.gap
(1 to (2 * options.int("graphview_iterations_rubberband"))).foldLeft(graph) {
@@ -373,12 +379,11 @@
final class Layout private(
val metrics: Metrics,
val input_graph: Graph_Display.Graph,
- val output_graph: Layout.Graph)
-{
+ val output_graph: Layout.Graph
+) {
/* vertex coordinates */
- def translate_vertex(v: Layout.Vertex, dx: Double, dy: Double): Layout =
- {
+ def translate_vertex(v: Layout.Vertex, dx: Double, dy: Double): Layout = {
if ((dx == 0.0 && dy == 0.0) || !output_graph.defined(v)) this
else {
val output_graph1 =
@@ -406,8 +411,7 @@
new Iterator[Layout.Info] {
private var index = 0
def hasNext: Boolean = output_graph.defined(Layout.Dummy(edge._1, edge._2, index))
- def next(): Layout.Info =
- {
+ def next(): Layout.Info = {
val info = output_graph.get_node(Layout.Dummy(edge._1, edge._2, index))
index += 1
info
--- a/src/Tools/Graphview/main_panel.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/Graphview/main_panel.scala Fri Apr 01 17:06:10 2022 +0200
@@ -13,8 +13,7 @@
import scala.swing.{SplitPane, Orientation}
-class Main_Panel(graphview: Graphview) extends SplitPane(Orientation.Vertical)
-{
+class Main_Panel(graphview: Graphview) extends SplitPane(Orientation.Vertical) {
oneTouchExpandable = true
val graph_panel = new Graph_Panel(graphview)
@@ -23,8 +22,7 @@
leftComponent = tree_panel
rightComponent = graph_panel
- def update_layout(): Unit =
- {
+ def update_layout(): Unit = {
graphview.update_layout()
tree_panel.refresh()
graph_panel.refresh()
--- a/src/Tools/Graphview/metrics.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/Graphview/metrics.scala Fri Apr 01 17:06:10 2022 +0200
@@ -15,10 +15,8 @@
import java.awt.geom.Rectangle2D
-object Metrics
-{
- val rendering_hints: RenderingHints =
- {
+object Metrics {
+ val rendering_hints: RenderingHints = {
val hints = new HashMap[RenderingHints.Key, AnyRef]
hints.put(RenderingHints.KEY_ANTIALIASING, RenderingHints.VALUE_ANTIALIAS_ON)
hints.put(RenderingHints.KEY_FRACTIONALMETRICS, RenderingHints.VALUE_FRACTIONALMETRICS_ON)
@@ -33,8 +31,7 @@
val default: Metrics = apply(new Font("Helvetica", Font.PLAIN, 12))
}
-class Metrics private(val font: Font)
-{
+class Metrics private(val font: Font) {
def string_bounds(s: String): Rectangle2D = font.getStringBounds(s, Metrics.font_render_context)
private val mix = string_bounds("mix")
val space_width: Double = string_bounds(" ").getWidth
@@ -56,8 +53,7 @@
def level_height2(num_lines: Int, num_edges: Int): Double =
(node_height2(num_lines) + gap) max (node_height2(1) * (num_edges max 5))
- object Pretty_Metric extends Pretty.Metric
- {
+ object Pretty_Metric extends Pretty.Metric {
val unit: Double = space_width max 1.0
def apply(s: String): Double = if (s == "\n") 1.0 else string_bounds(s).getWidth / unit
}
--- a/src/Tools/Graphview/model.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/Graphview/model.scala Fri Apr 01 17:06:10 2022 +0200
@@ -13,28 +13,24 @@
import java.awt.Color
-class Mutator_Container(val available: List[Mutator])
-{
+class Mutator_Container(val available: List[Mutator]) {
val events = new Mutator_Event.Bus
private var _mutators : List[Mutator.Info] = Nil
def apply(): List[Mutator.Info] = _mutators
- def apply(mutators: List[Mutator.Info]): Unit =
- {
+ def apply(mutators: List[Mutator.Info]): Unit = {
_mutators = mutators
events.event(Mutator_Event.New_List(mutators))
}
- def add(mutator: Mutator.Info): Unit =
- {
+ def add(mutator: Mutator.Info): Unit = {
_mutators = _mutators ::: List(mutator)
events.event(Mutator_Event.Add(mutator))
}
}
-class Model(val full_graph: Graph_Display.Graph)
-{
+class Model(val full_graph: Graph_Display.Graph) {
val Mutators =
new Mutator_Container(
List(
@@ -61,8 +57,7 @@
private var _colors = Map.empty[Graph_Display.Node, Color]
def colors = _colors
- private def build_colors(): Unit =
- {
+ private def build_colors(): Unit = {
_colors =
Colors().foldLeft(Map.empty[Graph_Display.Node, Color]) {
case (colors, m) =>
--- a/src/Tools/Graphview/mutator.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/Graphview/mutator.scala Fri Apr 01 17:06:10 2022 +0200
@@ -14,8 +14,7 @@
import scala.collection.immutable.SortedSet
-object Mutator
-{
+object Mutator {
sealed case class Info(enabled: Boolean, color: Color, mutator: Mutator)
def make(graphview: Graphview, m: Mutator): Info =
@@ -24,16 +23,16 @@
class Graph_Filter(
val name: String,
val description: String,
- pred: Graph_Display.Graph => Graph_Display.Graph) extends Filter
- {
+ pred: Graph_Display.Graph => Graph_Display.Graph
+ ) extends Filter {
def filter(graph: Graph_Display.Graph): Graph_Display.Graph = pred(graph)
}
class Graph_Mutator(
val name: String,
val description: String,
- pred: (Graph_Display.Graph, Graph_Display.Graph) => Graph_Display.Graph) extends Mutator
- {
+ pred: (Graph_Display.Graph, Graph_Display.Graph) => Graph_Display.Graph
+ ) extends Mutator {
def mutate(full_graph: Graph_Display.Graph, graph: Graph_Display.Graph): Graph_Display.Graph =
pred(full_graph, graph)
}
@@ -112,9 +111,11 @@
"Hides specified edge.",
(g, e) => e != edge)
- private def add_node_group(from: Graph_Display.Graph, to: Graph_Display.Graph,
- keys: List[Graph_Display.Node]) =
- {
+ private def add_node_group(
+ from: Graph_Display.Graph,
+ to: Graph_Display.Graph,
+ keys: List[Graph_Display.Node]
+ ) = {
// Add Nodes
val with_nodes =
keys.foldLeft(to) { case (graph, key) => graph.default_node(key, from.get_node(key)) }
@@ -161,8 +162,7 @@
})
}
-trait Mutator
-{
+trait Mutator {
val name: String
val description: String
def mutate(full_graph: Graph_Display.Graph, graph: Graph_Display.Graph): Graph_Display.Graph
@@ -170,8 +170,7 @@
override def toString: String = name
}
-trait Filter extends Mutator
-{
+trait Filter extends Mutator {
def mutate(full_graph: Graph_Display.Graph, graph: Graph_Display.Graph): Graph_Display.Graph = filter(graph)
def filter(graph: Graph_Display.Graph): Graph_Display.Graph
}
--- a/src/Tools/Graphview/mutator_dialog.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/Graphview/mutator_dialog.scala Fri Apr 01 17:06:10 2022 +0200
@@ -24,27 +24,24 @@
container: Mutator_Container,
caption: String,
reverse_caption: String = "Inverse",
- show_color_chooser: Boolean = true)
- extends Dialog
-{
+ show_color_chooser: Boolean = true
+) extends Dialog {
title = caption
private var _panels: List[Mutator_Panel] = Nil
private def panels = _panels
- private def panels_=(panels: List[Mutator_Panel]): Unit =
- {
+ private def panels_=(panels: List[Mutator_Panel]): Unit = {
_panels = panels
paint_panels()
}
container.events +=
- {
- case Mutator_Event.Add(m) => add_panel(new Mutator_Panel(m))
- case Mutator_Event.New_List(ms) => panels = get_panels(ms)
- }
+ {
+ case Mutator_Event.Add(m) => add_panel(new Mutator_Panel(m))
+ case Mutator_Event.New_List(ms) => panels = get_panels(ms)
+ }
- override def open(): Unit =
- {
+ override def open(): Unit = {
if (!visible) panels = get_panels(container())
super.open()
}
@@ -60,8 +57,7 @@
private def get_mutators(panels: List[Mutator_Panel]): List[Mutator.Info] =
panels.map(panel => panel.get_mutator)
- private def movePanelUp(m: Mutator_Panel) =
- {
+ private def movePanelUp(m: Mutator_Panel) = {
def moveUp(l: List[Mutator_Panel]): List[Mutator_Panel] =
l match {
case x :: y :: xs => if (y == m) y :: x :: xs else x :: moveUp(y :: xs)
@@ -71,8 +67,7 @@
panels = moveUp(panels)
}
- private def movePanelDown(m: Mutator_Panel) =
- {
+ private def movePanelDown(m: Mutator_Panel) = {
def moveDown(l: List[Mutator_Panel]): List[Mutator_Panel] =
l match {
case x :: y :: xs => if (x == m) y :: x :: xs else x :: moveDown(y :: xs)
@@ -82,18 +77,15 @@
panels = moveDown(panels)
}
- private def removePanel(m: Mutator_Panel): Unit =
- {
+ private def removePanel(m: Mutator_Panel): Unit = {
panels = panels.filter(_ != m).toList
}
- private def add_panel(m: Mutator_Panel): Unit =
- {
+ private def add_panel(m: Mutator_Panel): Unit = {
panels = panels ::: List(m)
}
- def paint_panels(): Unit =
- {
+ def paint_panels(): Unit = {
Focus_Traveral_Policy.clear()
filter_panel.contents.clear()
panels.map(x => {
@@ -160,8 +152,7 @@
}
private class Mutator_Panel(initials: Mutator.Info)
- extends BoxPanel(Orientation.Horizontal)
- {
+ extends BoxPanel(Orientation.Horizontal) {
private val inputs: List[(String, Input)] = get_inputs()
var focusList = List.empty[java.awt.Component]
private val enabledBox = new Check_Box_Input("Enabled", initials.enabled)
@@ -243,8 +234,7 @@
focusList = focusList.reverse
- def get_mutator: Mutator.Info =
- {
+ def get_mutator: Mutator.Info = {
val model = graphview.model
val m =
initials.mutator match {
@@ -317,53 +307,48 @@
}
}
- private trait Input
- {
+ private trait Input {
def string: String
def bool: Boolean
}
private class Text_Field_Input(txt: String, check: String => Boolean = (_: String) => true)
- extends TextField(txt) with Input
- {
+ extends TextField(txt) with Input {
preferredSize = new Dimension(125, 18)
private val default_foreground = foreground
reactions +=
- {
- case ValueChanged(_) =>
- foreground = if (check(text)) default_foreground else graphview.error_color
- }
+ {
+ case ValueChanged(_) =>
+ foreground = if (check(text)) default_foreground else graphview.error_color
+ }
def string = text
def bool = true
}
- private class Check_Box_Input(txt: String, s: Boolean) extends CheckBox(txt) with Input
- {
+ private class Check_Box_Input(txt: String, s: Boolean)
+ extends CheckBox(txt) with Input {
selected = s
def string = ""
def bool = selected
}
- private object Focus_Traveral_Policy extends FocusTraversalPolicy
- {
+ private object Focus_Traveral_Policy extends FocusTraversalPolicy {
private var items = Vector.empty[java.awt.Component]
def add(c: java.awt.Component): Unit = { items = items :+ c }
def addAll(cs: IterableOnce[java.awt.Component]): Unit = { items = items ++ cs }
def clear(): Unit = { items = Vector.empty }
- def getComponentAfter(root: java.awt.Container, c: java.awt.Component): java.awt.Component =
- {
+ def getComponentAfter(root: java.awt.Container, c: java.awt.Component): java.awt.Component = {
val i = items.indexOf(c)
if (i < 0) getDefaultComponent(root)
else items((i + 1) % items.length)
}
- def getComponentBefore(root: java.awt.Container, c: java.awt.Component): java.awt.Component =
- {
+ def getComponentBefore(root: java.awt.Container, c: java.awt.Component): java.awt.Component = {
val i = items.indexOf(c)
if (i < 0) getDefaultComponent(root)
else items((i - 1) % items.length)
--- a/src/Tools/Graphview/mutator_event.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/Graphview/mutator_event.scala Fri Apr 01 17:06:10 2022 +0200
@@ -11,16 +11,14 @@
import isabelle._
-object Mutator_Event
-{
+object Mutator_Event {
sealed abstract class Message
case class Add(m: Mutator.Info) extends Message
case class New_List(m: List[Mutator.Info]) extends Message
type Receiver = PartialFunction[Message, Unit]
- class Bus
- {
+ class Bus {
private val receivers = Synchronized[List[Receiver]](Nil)
def += (r: Receiver): Unit = receivers.change(Library.insert(r))
--- a/src/Tools/Graphview/popups.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/Graphview/popups.scala Fri Apr 01 17:06:10 2022 +0200
@@ -14,13 +14,12 @@
import scala.swing.{Action, Menu, MenuItem, Separator}
-object Popups
-{
+object Popups {
def apply(
graph_panel: Graph_Panel,
mouse_node: Option[Graph_Display.Node],
- selected_nodes: List[Graph_Display.Node]): JPopupMenu =
- {
+ selected_nodes: List[Graph_Display.Node]
+ ): JPopupMenu = {
val graphview = graph_panel.graphview
val add_mutator = graphview.model.Mutators.add _
--- a/src/Tools/Graphview/shapes.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/Graphview/shapes.scala Fri Apr 01 17:06:10 2022 +0200
@@ -15,16 +15,14 @@
RoundRectangle2D, PathIterator}
-object Shapes
-{
+object Shapes {
private val default_stroke =
new BasicStroke(1, BasicStroke.CAP_BUTT, BasicStroke.JOIN_ROUND)
def shape(info: Layout.Info): Rectangle2D.Double =
new Rectangle2D.Double(info.x - info.width2, info.y - info.height2, info.width, info.height)
- def highlight_node(gfx: Graphics2D, graphview: Graphview, node: Graph_Display.Node): Unit =
- {
+ def highlight_node(gfx: Graphics2D, graphview: Graphview, node: Graph_Display.Node): Unit = {
val metrics = graphview.metrics
val extra = metrics.char_width
val info = graphview.layout.get_node(node)
@@ -37,8 +35,7 @@
info.height + 2 * extra))
}
- def paint_node(gfx: Graphics2D, graphview: Graphview, node: Graph_Display.Node): Unit =
- {
+ def paint_node(gfx: Graphics2D, graphview: Graphview, node: Graph_Display.Node): Unit = {
val metrics = graphview.metrics
val info = graphview.layout.get_node(node)
val c = graphview.node_color(node)
@@ -66,21 +63,17 @@
}
}
- def paint_dummy(gfx: Graphics2D, graphview: Graphview, info: Layout.Info): Unit =
- {
+ def paint_dummy(gfx: Graphics2D, graphview: Graphview, info: Layout.Info): Unit = {
gfx.setStroke(default_stroke)
gfx.setColor(graphview.dummy_color)
gfx.draw(shape(info))
}
- object Straight_Edge
- {
- def paint(gfx: Graphics2D, graphview: Graphview, edge: Graph_Display.Edge): Unit =
- {
+ object Straight_Edge {
+ def paint(gfx: Graphics2D, graphview: Graphview, edge: Graph_Display.Edge): Unit = {
val p = graphview.layout.get_node(edge._1)
val q = graphview.layout.get_node(edge._2)
- val ds =
- {
+ val ds = {
val a = p.y min q.y
val b = p.y max q.y
graphview.layout.dummies_iterator(edge).filter(d => a < d.y && d.y < b).toList
@@ -101,16 +94,13 @@
}
}
- object Cardinal_Spline_Edge
- {
+ object Cardinal_Spline_Edge {
private val slack = 0.1
- def paint(gfx: Graphics2D, graphview: Graphview, edge: Graph_Display.Edge): Unit =
- {
+ def paint(gfx: Graphics2D, graphview: Graphview, edge: Graph_Display.Edge): Unit = {
val p = graphview.layout.get_node(edge._1)
val q = graphview.layout.get_node(edge._2)
- val ds =
- {
+ val ds = {
val a = p.y min q.y
val b = p.y max q.y
graphview.layout.dummies_iterator(edge).filter(d => a < d.y && d.y < b).toList
@@ -154,14 +144,11 @@
}
}
- object Arrow_Head
- {
+ object Arrow_Head {
type Point = (Double, Double)
- private def position(path: GeneralPath, shape: Shape): Option[AffineTransform] =
- {
- def intersecting_line(path: GeneralPath, shape: Shape): Option[(Point, Point)] =
- {
+ private def position(path: GeneralPath, shape: Shape): Option[AffineTransform] = {
+ def intersecting_line(path: GeneralPath, shape: Shape): Option[(Point, Point)] = {
val i = path.getPathIterator(null, 1.0)
val seg = Array[Double](0.0, 0.0, 0.0, 0.0, 0.0, 0.0)
var p1 = (0.0, 0.0)
@@ -182,8 +169,7 @@
None
}
- def binary_search(line: (Point, Point), shape: Shape): Option[AffineTransform] =
- {
+ def binary_search(line: (Point, Point), shape: Shape): Option[AffineTransform] = {
val ((fx, fy), (tx, ty)) = line
if (shape.contains(fx, fy) == shape.contains(tx, ty))
None
@@ -210,8 +196,7 @@
}
}
- def paint(gfx: Graphics2D, path: GeneralPath, shape: Shape): Unit =
- {
+ def paint(gfx: Graphics2D, path: GeneralPath, shape: Shape): Unit = {
position(path, shape) match {
case None =>
case Some(at) =>
--- a/src/Tools/Graphview/tree_panel.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/Graphview/tree_panel.scala Fri Apr 01 17:06:10 2022 +0200
@@ -19,12 +19,11 @@
import scala.swing.{Component, ScrollPane, BorderPanel, Label, TextField, Button, CheckBox, Action}
-class Tree_Panel(val graphview: Graphview, graph_panel: Graph_Panel) extends BorderPanel
-{
+class Tree_Panel(val graphview: Graphview, graph_panel: Graph_Panel)
+extends BorderPanel {
/* main actions */
- private def selection_action(): Unit =
- {
+ private def selection_action(): Unit = {
if (tree != null) {
graphview.current_node = None
graphview.Selection.clear()
@@ -45,8 +44,7 @@
}
}
- private def point_action(path: TreePath): Unit =
- {
+ private def point_action(path: TreePath): Unit = {
if (tree_pane != null && path != null) {
val action_node =
path.getLastPathComponent match {
@@ -149,8 +147,7 @@
/* main layout */
- def refresh(): Unit =
- {
+ def refresh(): Unit = {
val new_nodes = graphview.visible_graph.topological_order
if (new_nodes != nodes) {
tree.clearSelection
--- a/src/Tools/VSCode/src/build_vscode_extension.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/VSCode/src/build_vscode_extension.scala Fri Apr 01 17:06:10 2022 +0200
@@ -10,8 +10,7 @@
import isabelle._
-object Build_VSCode
-{
+object Build_VSCode {
/* build grammar */
def default_logic: String = Isabelle_System.getenv("ISABELLE_LOGIC")
@@ -19,8 +18,8 @@
def build_grammar(options: Options, build_dir: Path,
logic: String = default_logic,
dirs: List[Path] = Nil,
- progress: Progress = new Progress): Unit =
- {
+ progress: Progress = new Progress
+ ): Unit = {
val keywords =
Sessions.base_info(options, logic, dirs = dirs).check.base.overall_syntax.keywords
@@ -141,8 +140,8 @@
target_dir: Path = Path.current,
logic: String = default_logic,
dirs: List[Path] = Nil,
- progress: Progress = new Progress): Unit =
- {
+ progress: Progress = new Progress
+ ): Unit = {
Isabelle_System.require_command("node")
Isabelle_System.require_command("yarn")
Isabelle_System.require_command("vsce")
@@ -158,12 +157,10 @@
/* build */
val vsix_name =
- Isabelle_System.with_tmp_dir("build")(build_dir =>
- {
+ Isabelle_System.with_tmp_dir("build")(build_dir => {
val manifest_text = File.read(VSCode_Main.extension_dir + VSCode_Main.MANIFEST)
val manifest_entries = split_lines(manifest_text).filter(_.nonEmpty)
- val manifest_shasum: String =
- {
+ val manifest_shasum: String = {
val a = SHA1.digest(manifest_text).shasum("<MANIFEST>")
val bs =
for (entry <- manifest_entries)
@@ -217,8 +214,8 @@
val isabelle_tool =
Isabelle_Tool("build_vscode_extension", "build Isabelle/VSCode extension module",
- Scala_Project.here, args =>
- {
+ Scala_Project.here,
+ args => {
var target_dir = Path.current
var dirs: List[Path] = Nil
var logic = default_logic
--- a/src/Tools/VSCode/src/build_vscodium.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/VSCode/src/build_vscodium.scala Fri Apr 01 17:06:10 2022 +0200
@@ -14,8 +14,7 @@
import java.util.Base64
-object Build_VSCodium
-{
+object Build_VSCodium {
/* global parameters */
lazy val version: String = Isabelle_System.getenv_strict("ISABELLE_VSCODE_VERSION")
@@ -27,8 +26,7 @@
/* Isabelle symbols (static subset only) */
- def make_symbols(): File.Content =
- {
+ def make_symbols(): File.Content = {
val symbols = Symbol.Symbols.load(static = true)
val symbols_js =
JSON.Format.apply_lines(
@@ -42,8 +40,7 @@
File.Content(Path.explode("symbols.json"), symbols_js)
}
- def make_isabelle_encoding(header: String): File.Content =
- {
+ def make_isabelle_encoding(header: String): File.Content = {
val symbols = Symbol.Symbols.load(static = true)
val symbols_js =
JSON.Format.apply_lines(
@@ -64,19 +61,17 @@
platform: Platform.Family.Value,
download_template: String,
build_name: String,
- env: List[String])
- {
+ env: List[String]
+ ) {
def is_linux: Boolean = platform == Platform.Family.linux
def download_name: String = "VSCodium-" + download_template.replace("{VERSION}", version)
def download_zip: Boolean = download_name.endsWith(".zip")
- def download(dir: Path, progress: Progress = new Progress): Unit =
- {
+ def download(dir: Path, progress: Progress = new Progress): Unit = {
if (download_zip) Isabelle_System.require_command("unzip", test = "-h")
- Isabelle_System.with_tmp_file("download")(download_file =>
- {
+ Isabelle_System.with_tmp_file("download")(download_file => {
Isabelle_System.download_file(vscodium_download + "/" + version + "/" + download_name,
download_file, progress = progress)
@@ -90,8 +85,7 @@
})
}
- def get_vscodium_repository(build_dir: Path, progress: Progress = new Progress): Unit =
- {
+ def get_vscodium_repository(build_dir: Path, progress: Progress = new Progress): Unit = {
progress.echo("Getting VSCodium repository ...")
Isabelle_System.bash(
List(
@@ -104,8 +98,7 @@
Isabelle_System.bash(environment + "\n" + "./get_repo.sh", cwd = build_dir.file).check
}
- def platform_dir(dir: Path): Path =
- {
+ def platform_dir(dir: Path): Path = {
val platform_name =
if (platform == Platform.Family.windows) Platform.Family.native(platform)
else Platform.Family.standard(platform)
@@ -118,8 +111,7 @@
(("MS_TAG=" + Bash.string(version)) :: "SHOULD_BUILD=yes" :: "VSCODE_ARCH=x64" :: env)
.map(s => "export " + s + "\n").mkString
- def patch_sources(base_dir: Path): String =
- {
+ def patch_sources(base_dir: Path): String = {
val dir = base_dir + Path.explode("vscode")
Isabelle_System.with_copy_dir(dir, dir.orig) {
// macos icns
@@ -152,8 +144,7 @@
}
}
- def patch_resources(base_dir: Path): String =
- {
+ def patch_resources(base_dir: Path): String = {
val dir = base_dir + resources
val patch =
Isabelle_System.with_copy_dir(dir, dir.orig) {
@@ -189,8 +180,7 @@
patch
}
- def init_resources(base_dir: Path): Path =
- {
+ def init_resources(base_dir: Path): Path = {
val dir = base_dir + resources
if (platform == Platform.Family.macos) {
Isabelle_System.symlink(Path.explode("VSCodium.app/Contents/Resources"), dir)
@@ -198,10 +188,8 @@
dir
}
- def setup_node(target_dir: Path, progress: Progress): Unit =
- {
- Isabelle_System.with_tmp_dir("download")(download_dir =>
- {
+ def setup_node(target_dir: Path, progress: Progress): Unit = {
+ Isabelle_System.with_tmp_dir("download")(download_dir => {
download(download_dir, progress = progress)
val dir1 = init_resources(download_dir)
val dir2 = init_resources(target_dir)
@@ -213,8 +201,7 @@
})
}
- def setup_electron(dir: Path): Unit =
- {
+ def setup_electron(dir: Path): Unit = {
val electron = Path.explode("electron")
platform match {
case Platform.Family.linux | Platform.Family.linux_arm =>
@@ -228,8 +215,7 @@
}
}
- def setup_executables(dir: Path): Unit =
- {
+ def setup_executables(dir: Path): Unit = {
Isabelle_System.rm_tree(dir + Path.explode("bin"))
if (platform == Platform.Family.windows) {
@@ -248,8 +234,7 @@
// see https://github.com/microsoft/vscode/blob/main/build/gulpfile.vscode.js
// function computeChecksum(filename)
- private def file_checksum(path: Path): String =
- {
+ private def file_checksum(path: Path): String = {
val digest = MessageDigest.getInstance("MD5")
digest.update(Bytes.read(path).array)
Bytes(Base64.getEncoder.encode(digest.digest()))
@@ -282,8 +267,7 @@
/* check system */
- def check_system(platforms: List[Platform.Family.Value]): Unit =
- {
+ def check_system(platforms: List[Platform.Family.Value]): Unit = {
if (Platform.family != Platform.Family.linux) error("Not a Linux/x86_64 system")
Isabelle_System.require_command("git")
@@ -303,13 +287,11 @@
/* original repository clones and patches */
- def vscodium_patch(verbose: Boolean = false, progress: Progress = new Progress): String =
- {
+ def vscodium_patch(verbose: Boolean = false, progress: Progress = new Progress): String = {
val platform_info = linux_platform_info
check_system(List(platform_info.platform))
- Isabelle_System.with_tmp_dir("build")(build_dir =>
- {
+ Isabelle_System.with_tmp_dir("build")(build_dir => {
platform_info.get_vscodium_repository(build_dir, progress = progress)
val vscode_dir = build_dir + Path.explode("vscode")
progress.echo("Prepare ...")
@@ -337,8 +319,8 @@
target_dir: Path = Path.current,
platforms: List[Platform.Family.Value] = default_platforms,
verbose: Boolean = false,
- progress: Progress = new Progress): Unit =
- {
+ progress: Progress = new Progress
+ ): Unit = {
check_system(platforms)
@@ -366,8 +348,7 @@
for (platform <- platforms) yield {
val platform_info = the_platform_info(platform)
- Isabelle_System.with_tmp_dir("build")(build_dir =>
- {
+ Isabelle_System.with_tmp_dir("build")(build_dir => {
progress.echo("\n* Building " + platform + ":")
platform_info.get_vscodium_repository(build_dir, progress = progress)
@@ -440,8 +421,8 @@
val isabelle_tool1 =
Isabelle_Tool("build_vscodium", "build component for VSCodium",
- Scala_Project.here, args =>
- {
+ Scala_Project.here,
+ args => {
var target_dir = Path.current
var platforms = default_platforms
var verbose = false
@@ -474,8 +455,8 @@
val isabelle_tool2 =
Isabelle_Tool("vscode_patch", "patch VSCode source tree",
- Scala_Project.here, args =>
- {
+ Scala_Project.here,
+ args => {
var base_dir = Path.current
val getopts = Getopts("""
--- a/src/Tools/VSCode/src/channel.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/VSCode/src/channel.scala Fri Apr 01 17:06:10 2022 +0200
@@ -14,8 +14,12 @@
import scala.collection.mutable
-class Channel(in: InputStream, out: OutputStream, log: Logger = No_Logger, verbose: Boolean = false)
-{
+class Channel(
+ in: InputStream,
+ out: OutputStream,
+ log: Logger = No_Logger,
+ verbose: Boolean = false
+) {
/* read message */
private val Content_Length = """^\s*Content-Length:\s*(\d+)\s*$""".r
@@ -26,23 +30,20 @@
case None => ""
}
- private def read_header(): List[String] =
- {
+ private def read_header(): List[String] = {
val header = new mutable.ListBuffer[String]
var line = ""
while ({ line = read_line(); line != "" }) header += line
header.toList
}
- private def read_content(n: Int): String =
- {
+ private def read_content(n: Int): String = {
val bytes = Bytes.read_stream(in, limit = n)
if (bytes.length == n) bytes.text
else error("Bad message content (unexpected EOF after " + bytes.length + " of " + n + " bytes)")
}
- def read(): Option[JSON.T] =
- {
+ def read(): Option[JSON.T] = {
read_header() match {
case Nil => None
case Content_Length(s) :: _ =>
@@ -61,8 +62,7 @@
/* write message */
- def write(json: JSON.T): Unit =
- {
+ def write(json: JSON.T): Unit = {
val msg = JSON.Format(json)
val content = UTF8.bytes(msg)
val n = content.length
@@ -90,8 +90,7 @@
def log_warning(msg: String): Unit = display_message(LSP.MessageType.Warning, msg, false)
def log_writeln(msg: String): Unit = display_message(LSP.MessageType.Info, msg, false)
- object Error_Logger extends Logger
- {
+ object Error_Logger extends Logger {
def apply(msg: => String): Unit = log_error_message(msg)
}
--- a/src/Tools/VSCode/src/dynamic_output.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/VSCode/src/dynamic_output.scala Fri Apr 01 17:06:10 2022 +0200
@@ -10,13 +10,13 @@
import isabelle._
-object Dynamic_Output
-{
- sealed case class State(do_update: Boolean = true, output: List[XML.Tree] = Nil)
- {
+object Dynamic_Output {
+ sealed case class State(do_update: Boolean = true, output: List[XML.Tree] = Nil) {
def handle_update(
- resources: VSCode_Resources, channel: Channel, restriction: Option[Set[Command]]): State =
- {
+ resources: VSCode_Resources,
+ channel: Channel,
+ restriction: Option[Set[Command]]
+ ): State = {
val st1 =
resources.get_caret() match {
case None => copy(output = Nil)
@@ -57,8 +57,7 @@
}
-class Dynamic_Output private(server: Language_Server)
-{
+class Dynamic_Output private(server: Language_Server) {
private val state = Synchronized(Dynamic_Output.State())
private def handle_update(restriction: Option[Set[Command]]): Unit =
@@ -76,15 +75,13 @@
handle_update(None)
}
- def init(): Unit =
- {
+ def init(): Unit = {
server.session.commands_changed += main
server.session.caret_focus += main
handle_update(None)
}
- def exit(): Unit =
- {
+ def exit(): Unit = {
server.session.commands_changed -= main
server.session.caret_focus -= main
}
--- a/src/Tools/VSCode/src/language_server.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/VSCode/src/language_server.scala Fri Apr 01 17:06:10 2022 +0200
@@ -19,8 +19,7 @@
import scala.collection.mutable
-object Language_Server
-{
+object Language_Server {
type Editor = isabelle.Editor[Unit]
@@ -29,8 +28,8 @@
private lazy val default_logic = Isabelle_System.getenv("ISABELLE_LOGIC")
val isabelle_tool =
- Isabelle_Tool("vscode_server", "VSCode Language Server for PIDE", Scala_Project.here, args =>
- {
+ Isabelle_Tool("vscode_server", "VSCode Language Server for PIDE", Scala_Project.here,
+ args => {
try {
var logic_ancestor: Option[String] = None
var log_file: Option[Path] = None
@@ -109,8 +108,8 @@
session_requirements: Boolean = false,
session_no_build: Boolean = false,
modes: List[String] = Nil,
- log: Logger = No_Logger)
-{
+ log: Logger = No_Logger
+) {
server =>
@@ -136,8 +135,9 @@
File_Watcher(sync_documents, options.seconds("vscode_load_delay"))
private val delay_input: Delay =
- Delay.last(options.seconds("vscode_input_delay"), channel.Error_Logger)
- { resources.flush_input(session, channel) }
+ Delay.last(options.seconds("vscode_input_delay"), channel.Error_Logger) {
+ resources.flush_input(session, channel)
+ }
private val delay_load: Delay =
Delay.last(options.seconds("vscode_load_delay"), channel.Error_Logger) {
@@ -147,8 +147,7 @@
if (invoke_load) delay_load.invoke()
}
- private def close_document(file: JFile): Unit =
- {
+ private def close_document(file: JFile): Unit = {
if (resources.close_model(file)) {
file_watcher.register_parent(file)
sync_documents(Set(file))
@@ -157,19 +156,19 @@
}
}
- private def sync_documents(changed: Set[JFile]): Unit =
- {
+ private def sync_documents(changed: Set[JFile]): Unit = {
resources.sync_models(changed)
delay_input.invoke()
delay_output.invoke()
}
private def change_document(
- file: JFile, version: Long, changes: List[LSP.TextDocumentChange]): Unit =
- {
+ file: JFile,
+ version: Long,
+ changes: List[LSP.TextDocumentChange]
+ ): Unit = {
val norm_changes = new mutable.ListBuffer[LSP.TextDocumentChange]
- @tailrec def norm(chs: List[LSP.TextDocumentChange]): Unit =
- {
+ @tailrec def norm(chs: List[LSP.TextDocumentChange]): Unit = {
if (chs.nonEmpty) {
val (full_texts, rest1) = chs.span(_.range.isEmpty)
val (edits, rest2) = rest1.span(_.range.nonEmpty)
@@ -190,11 +189,11 @@
/* caret handling */
private val delay_caret_update: Delay =
- Delay.last(options.seconds("vscode_input_delay"), channel.Error_Logger)
- { session.caret_focus.post(Session.Caret_Focus) }
+ Delay.last(options.seconds("vscode_input_delay"), channel.Error_Logger) {
+ session.caret_focus.post(Session.Caret_Focus)
+ }
- private def update_caret(caret: Option[(JFile, Line.Position)]): Unit =
- {
+ private def update_caret(caret: Option[(JFile, Line.Position)]): Unit = {
resources.update_caret(caret)
delay_caret_update.invoke()
delay_input.invoke()
@@ -206,13 +205,11 @@
private lazy val preview_panel = new Preview_Panel(resources)
private lazy val delay_preview: Delay =
- Delay.last(options.seconds("vscode_output_delay"), channel.Error_Logger)
- {
+ Delay.last(options.seconds("vscode_output_delay"), channel.Error_Logger) {
if (preview_panel.flush(channel)) delay_preview.invoke()
}
- private def request_preview(file: JFile, column: Int): Unit =
- {
+ private def request_preview(file: JFile, column: Int): Unit = {
preview_panel.request(file, column)
delay_preview.invoke()
}
@@ -221,19 +218,16 @@
/* output to client */
private val delay_output: Delay =
- Delay.last(options.seconds("vscode_output_delay"), channel.Error_Logger)
- {
+ Delay.last(options.seconds("vscode_output_delay"), channel.Error_Logger) {
if (resources.flush_output(channel)) delay_output.invoke()
}
- def update_output(changed_nodes: Iterable[JFile]): Unit =
- {
+ def update_output(changed_nodes: Iterable[JFile]): Unit = {
resources.update_output(changed_nodes)
delay_output.invoke()
}
- def update_output_visible(): Unit =
- {
+ def update_output_visible(): Unit = {
resources.update_output_visible()
delay_output.invoke()
}
@@ -252,16 +246,13 @@
/* init and exit */
- def init(id: LSP.Id): Unit =
- {
- def reply_ok(msg: String): Unit =
- {
+ def init(id: LSP.Id): Unit = {
+ def reply_ok(msg: String): Unit = {
channel.write(LSP.Initialize.reply(id, ""))
channel.writeln(msg)
}
- def reply_error(msg: String): Unit =
- {
+ def reply_error(msg: String): Unit = {
channel.write(LSP.Initialize.reply(id, msg))
channel.error_message(msg)
}
@@ -289,8 +280,8 @@
if (!build().ok) { progress.echo(fail_msg); error(fail_msg) }
}
- val resources = new VSCode_Resources(options, base_info, log)
- {
+ val resources =
+ new VSCode_Resources(options, base_info, log) {
override def commit(change: Session.Change): Unit =
if (change.deps_changed || undefined_blobs(change.version.nodes).nonEmpty)
delay_load.invoke()
@@ -320,8 +311,7 @@
}
}
- def shutdown(id: LSP.Id): Unit =
- {
+ def shutdown(id: LSP.Id): Unit = {
def reply(err: String): Unit = channel.write(LSP.Shutdown.reply(id, err))
session_.change({
@@ -348,8 +338,7 @@
})
}
- def exit(): Unit =
- {
+ def exit(): Unit = {
log("\n")
sys.exit(if (session_.value.isEmpty) Process_Result.RC.ok else Process_Result.RC.failure)
}
@@ -357,8 +346,7 @@
/* completion */
- def completion(id: LSP.Id, node_pos: Line.Node_Position): Unit =
- {
+ def completion(id: LSP.Id, node_pos: Line.Node_Position): Unit = {
val result =
(for ((rendering, offset) <- rendering_offset(node_pos))
yield rendering.completion(node_pos, offset)) getOrElse Nil
@@ -368,8 +356,7 @@
/* spell-checker dictionary */
- def update_dictionary(include: Boolean, permanent: Boolean): Unit =
- {
+ def update_dictionary(include: Boolean, permanent: Boolean): Unit = {
for {
spell_checker <- resources.spell_checker.get
caret <- resources.get_caret()
@@ -382,10 +369,8 @@
}
}
- def reset_dictionary(): Unit =
- {
- for (spell_checker <- resources.spell_checker.get)
- {
+ def reset_dictionary(): Unit = {
+ for (spell_checker <- resources.spell_checker.get) {
spell_checker.reset()
update_output_visible()
}
@@ -394,8 +379,7 @@
/* hover */
- def hover(id: LSP.Id, node_pos: Line.Node_Position): Unit =
- {
+ def hover(id: LSP.Id, node_pos: Line.Node_Position): Unit = {
val result =
for {
(rendering, offset) <- rendering_offset(node_pos)
@@ -411,8 +395,7 @@
/* goto definition */
- def goto_definition(id: LSP.Id, node_pos: Line.Node_Position): Unit =
- {
+ def goto_definition(id: LSP.Id, node_pos: Line.Node_Position): Unit = {
val result =
(for ((rendering, offset) <- rendering_offset(node_pos))
yield rendering.hyperlinks(Text.Range(offset, offset + 1))) getOrElse Nil
@@ -422,8 +405,7 @@
/* document highlights */
- def document_highlights(id: LSP.Id, node_pos: Line.Node_Position): Unit =
- {
+ def document_highlights(id: LSP.Id, node_pos: Line.Node_Position): Unit = {
val result =
(for ((rendering, offset) <- rendering_offset(node_pos))
yield {
@@ -437,12 +419,10 @@
/* main loop */
- def start(): Unit =
- {
+ def start(): Unit = {
log("Server started " + Date.now())
- def handle(json: JSON.T): Unit =
- {
+ def handle(json: JSON.T): Unit = {
try {
json match {
case LSP.Initialize(id) => init(id)
@@ -477,8 +457,7 @@
catch { case exn: Throwable => channel.log_error_message(Exn.message(exn)) }
}
- @tailrec def loop(): Unit =
- {
+ @tailrec def loop(): Unit = {
channel.read() match {
case Some(json) =>
json match {
@@ -495,8 +474,7 @@
/* abstract editor operations */
- object editor extends Language_Server.Editor
- {
+ object editor extends Language_Server.Editor {
/* session */
override def session: Session = server.session
@@ -511,16 +489,14 @@
override def current_node_snapshot(context: Unit): Option[Document.Snapshot] =
resources.get_caret().map(_.model.snapshot())
- override def node_snapshot(name: Document.Node.Name): Document.Snapshot =
- {
+ override def node_snapshot(name: Document.Node.Name): Document.Snapshot = {
resources.get_model(name) match {
case Some(model) => model.snapshot()
case None => session.snapshot(name)
}
}
- def current_command(snapshot: Document.Snapshot): Option[Command] =
- {
+ def current_command(snapshot: Document.Snapshot): Option[Command] = {
resources.get_caret() match {
case Some(caret) => snapshot.current_command(caret.node_name, caret.offset)
case None => None
@@ -545,9 +521,11 @@
/* hyperlinks */
override def hyperlink_command(
- focus: Boolean, snapshot: Document.Snapshot, id: Document_ID.Generic,
- offset: Symbol.Offset = 0): Option[Hyperlink] =
- {
+ focus: Boolean,
+ snapshot: Document.Snapshot,
+ id: Document_ID.Generic,
+ offset: Symbol.Offset = 0
+ ): Option[Hyperlink] = {
if (snapshot.is_outdated) None
else
snapshot.find_command_position(id, offset).map(node_pos =>
--- a/src/Tools/VSCode/src/lsp.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/VSCode/src/lsp.scala Fri Apr 01 17:06:10 2022 +0200
@@ -13,16 +13,13 @@
import java.io.{File => JFile}
-object LSP
-{
+object LSP {
/* abstract message */
- object Message
- {
+ object Message {
val empty: JSON.Object.T = JSON.Object("jsonrpc" -> "2.0")
- def log(prefix: String, json: JSON.T, logger: Logger, verbose: Boolean): Unit =
- {
+ def log(prefix: String, json: JSON.T, logger: Logger, verbose: Boolean): Unit = {
val header =
json match {
case JSON.Object(obj) => obj -- (obj.keySet - "method" - "id")
@@ -36,8 +33,7 @@
/* notification */
- object Notification
- {
+ object Notification {
def apply(method: String, params: JSON.T): JSON.T =
Message.empty + ("method" -> method) + ("params" -> params)
@@ -48,8 +44,7 @@
} yield (method, params)
}
- class Notification0(name: String)
- {
+ class Notification0(name: String) {
def unapply(json: JSON.T): Option[Unit] =
json match {
case Notification(method, _) if method == name => Some(())
@@ -60,13 +55,11 @@
/* request message */
- object Id
- {
+ object Id {
def empty: Id = Id("")
}
- sealed case class Id(id: Any)
- {
+ sealed case class Id(id: Any) {
require(
id.isInstanceOf[Int] ||
id.isInstanceOf[Long] ||
@@ -76,8 +69,7 @@
override def toString: String = id.toString
}
- object RequestMessage
- {
+ object RequestMessage {
def apply(id: Id, method: String, params: JSON.T): JSON.T =
Message.empty + ("id" -> id.id) + ("method" -> method) + ("params" -> params)
@@ -89,8 +81,7 @@
} yield (Id(id), method, params)
}
- class Request0(name: String)
- {
+ class Request0(name: String) {
def unapply(json: JSON.T): Option[Id] =
json match {
case RequestMessage(id, method, _) if method == name => Some(id)
@@ -98,8 +89,7 @@
}
}
- class RequestTextDocumentPosition(name: String)
- {
+ class RequestTextDocumentPosition(name: String) {
def unapply(json: JSON.T): Option[(Id, Line.Node_Position)] =
json match {
case RequestMessage(id, method, Some(TextDocumentPosition(node_pos))) if method == name =>
@@ -111,8 +101,7 @@
/* response message */
- object ResponseMessage
- {
+ object ResponseMessage {
def apply(id: Id, result: Option[JSON.T] = None, error: Option[ResponseError] = None): JSON.T =
Message.empty + ("id" -> id.id) ++
JSON.optional("result" -> result) ++
@@ -126,14 +115,12 @@
JSON.string(json, "id") == Some("") && JSON.value(json, "result").isDefined
}
- sealed case class ResponseError(code: Int, message: String, data: Option[JSON.T] = None)
- {
+ sealed case class ResponseError(code: Int, message: String, data: Option[JSON.T] = None) {
def json: JSON.T =
JSON.Object("code" -> code, "message" -> message) ++ JSON.optional("data" -> data)
}
- object ErrorCodes
- {
+ object ErrorCodes {
val ParseError = -32700
val InvalidRequest = -32600
val MethodNotFound = -32601
@@ -146,15 +133,13 @@
/* init and exit */
- object Initialize extends Request0("initialize")
- {
+ object Initialize extends Request0("initialize") {
def reply(id: Id, error: String): JSON.T =
ResponseMessage.strict(
id, Some(JSON.Object("capabilities" -> ServerCapabilities.json)), error)
}
- object ServerCapabilities
- {
+ object ServerCapabilities {
val json: JSON.T =
JSON.Object(
"textDocumentSync" -> 2,
@@ -170,8 +155,7 @@
object Initialized extends Notification0("initialized")
- object Shutdown extends Request0("shutdown")
- {
+ object Shutdown extends Request0("shutdown") {
def reply(id: Id, error: String): JSON.T =
ResponseMessage.strict(id, Some("OK"), error)
}
@@ -181,8 +165,7 @@
/* document positions */
- object Position
- {
+ object Position {
def apply(pos: Line.Position): JSON.T =
JSON.Object("line" -> pos.line, "character" -> pos.column)
@@ -193,8 +176,7 @@
} yield Line.Position(line, column)
}
- object Range
- {
+ object Range {
def compact(range: Line.Range): List[Int] =
List(range.start.line, range.start.column, range.stop.line, range.stop.column)
@@ -208,8 +190,7 @@
}
}
- object Location
- {
+ object Location {
def apply(loc: Line.Node_Range): JSON.T =
JSON.Object("uri" -> Url.print_file_name(loc.name), "range" -> Range(loc.range))
@@ -222,8 +203,7 @@
} yield Line.Node_Range(Url.absolute_file_name(uri), range)
}
- object TextDocumentPosition
- {
+ object TextDocumentPosition {
def unapply(json: JSON.T): Option[Line.Node_Position] =
for {
doc <- JSON.value(json, "textDocument")
@@ -237,13 +217,11 @@
/* marked strings */
- sealed case class MarkedString(text: String, language: String = "plaintext")
- {
+ sealed case class MarkedString(text: String, language: String = "plaintext") {
def json: JSON.T = JSON.Object("language" -> language, "value" -> text)
}
- object MarkedStrings
- {
+ object MarkedStrings {
def json(msgs: List[MarkedString]): Option[JSON.T] =
msgs match {
case Nil => None
@@ -255,16 +233,14 @@
/* diagnostic messages */
- object MessageType
- {
+ object MessageType {
val Error = 1
val Warning = 2
val Info = 3
val Log = 4
}
- object DisplayMessage
- {
+ object DisplayMessage {
def apply(message_type: Int, message: String, show: Boolean): JSON.T =
Notification(if (show) "window/showMessage" else "window/logMessage",
JSON.Object("type" -> message_type, "message" -> message))
@@ -273,8 +249,7 @@
/* commands */
- sealed case class Command(title: String, command: String, arguments: List[JSON.T] = Nil)
- {
+ sealed case class Command(title: String, command: String, arguments: List[JSON.T] = Nil) {
def json: JSON.T =
JSON.Object("title" -> title, "command" -> command, "arguments" -> arguments)
}
@@ -282,8 +257,7 @@
/* document edits */
- object DidOpenTextDocument
- {
+ object DidOpenTextDocument {
def unapply(json: JSON.T): Option[(JFile, String, Long, String)] =
json match {
case Notification("textDocument/didOpen", Some(params)) =>
@@ -302,8 +276,7 @@
sealed case class TextDocumentChange(range: Option[Line.Range], text: String)
- object DidChangeTextDocument
- {
+ object DidChangeTextDocument {
def unapply_change(json: JSON.T): Option[TextDocumentChange] =
for { text <- JSON.string(json, "text") }
yield TextDocumentChange(JSON.value(json, "range", Range.unapply), text)
@@ -322,8 +295,7 @@
}
}
- class TextDocumentNotification(name: String)
- {
+ class TextDocumentNotification(name: String) {
def unapply(json: JSON.T): Option[JFile] =
json match {
case Notification(method, Some(params)) if method == name =>
@@ -342,21 +314,18 @@
/* workspace edits */
- sealed case class TextEdit(range: Line.Range, new_text: String)
- {
+ sealed case class TextEdit(range: Line.Range, new_text: String) {
def json: JSON.T = JSON.Object("range" -> Range(range), "newText" -> new_text)
}
- sealed case class TextDocumentEdit(file: JFile, version: Long, edits: List[TextEdit])
- {
+ sealed case class TextDocumentEdit(file: JFile, version: Long, edits: List[TextEdit]) {
def json: JSON.T =
JSON.Object(
"textDocument" -> JSON.Object("uri" -> Url.print_file(file), "version" -> version),
"edits" -> edits.map(_.json))
}
- object WorkspaceEdit
- {
+ object WorkspaceEdit {
def apply(edits: List[TextDocumentEdit]): JSON.T =
RequestMessage(Id.empty, "workspace/applyEdit",
JSON.Object("edit" -> JSON.Object("documentChanges" -> edits.map(_.json))))
@@ -372,8 +341,8 @@
documentation: Option[String] = None,
text: Option[String] = None,
range: Option[Line.Range] = None,
- command: Option[Command] = None)
- {
+ command: Option[Command] = None
+ ) {
def json: JSON.T =
JSON.Object("label" -> label) ++
JSON.optional("kind" -> kind) ++
@@ -385,8 +354,7 @@
JSON.optional("command" -> command.map(_.json))
}
- object Completion extends RequestTextDocumentPosition("textDocument/completion")
- {
+ object Completion extends RequestTextDocumentPosition("textDocument/completion") {
def reply(id: Id, result: List[CompletionItem]): JSON.T =
ResponseMessage(id, Some(result.map(_.json)))
}
@@ -394,38 +362,31 @@
/* spell checker */
- object Include_Word extends Notification0("PIDE/include_word")
- {
+ object Include_Word extends Notification0("PIDE/include_word") {
val command = Command("Include word", "isabelle.include-word")
}
- object Include_Word_Permanently extends Notification0("PIDE/include_word_permanently")
- {
+ object Include_Word_Permanently extends Notification0("PIDE/include_word_permanently") {
val command = Command("Include word permanently", "isabelle.include-word-permanently")
}
- object Exclude_Word extends Notification0("PIDE/exclude_word")
- {
+ object Exclude_Word extends Notification0("PIDE/exclude_word") {
val command = Command("Exclude word", "isabelle.exclude-word")
}
- object Exclude_Word_Permanently extends Notification0("PIDE/exclude_word_permanently")
- {
+ object Exclude_Word_Permanently extends Notification0("PIDE/exclude_word_permanently") {
val command = Command("Exclude word permanently", "isabelle.exclude-word-permanently")
}
- object Reset_Words extends Notification0("PIDE/reset_words")
- {
+ object Reset_Words extends Notification0("PIDE/reset_words") {
val command = Command("Reset non-permanent words", "isabelle.reset-words")
}
/* hover request */
- object Hover extends RequestTextDocumentPosition("textDocument/hover")
- {
- def reply(id: Id, result: Option[(Line.Range, List[MarkedString])]): JSON.T =
- {
+ object Hover extends RequestTextDocumentPosition("textDocument/hover") {
+ def reply(id: Id, result: Option[(Line.Range, List[MarkedString])]): JSON.T = {
val res =
result match {
case Some((range, contents)) =>
@@ -441,8 +402,7 @@
/* goto definition request */
- object GotoDefinition extends RequestTextDocumentPosition("textDocument/definition")
- {
+ object GotoDefinition extends RequestTextDocumentPosition("textDocument/definition") {
def reply(id: Id, result: List[Line.Node_Range]): JSON.T =
ResponseMessage(id, Some(result.map(Location.apply)))
}
@@ -450,15 +410,13 @@
/* document highlights request */
- object DocumentHighlight
- {
+ object DocumentHighlight {
def text(range: Line.Range): DocumentHighlight = DocumentHighlight(range, Some(1))
def read(range: Line.Range): DocumentHighlight = DocumentHighlight(range, Some(2))
def write(range: Line.Range): DocumentHighlight = DocumentHighlight(range, Some(3))
}
- sealed case class DocumentHighlight(range: Line.Range, kind: Option[Int] = None)
- {
+ sealed case class DocumentHighlight(range: Line.Range, kind: Option[Int] = None) {
def json: JSON.T =
kind match {
case None => JSON.Object("range" -> Range(range))
@@ -466,8 +424,7 @@
}
}
- object DocumentHighlights extends RequestTextDocumentPosition("textDocument/documentHighlight")
- {
+ object DocumentHighlights extends RequestTextDocumentPosition("textDocument/documentHighlight") {
def reply(id: Id, result: List[DocumentHighlight]): JSON.T =
ResponseMessage(id, Some(result.map(_.json)))
}
@@ -475,9 +432,13 @@
/* diagnostics */
- sealed case class Diagnostic(range: Line.Range, message: String,
- severity: Option[Int] = None, code: Option[Int] = None, source: Option[String] = None)
- {
+ sealed case class Diagnostic(
+ range: Line.Range,
+ message: String,
+ severity: Option[Int] = None,
+ code: Option[Int] = None,
+ source: Option[String] = None
+ ) {
def json: JSON.T =
Message.empty + ("range" -> Range(range)) + ("message" -> message) ++
JSON.optional("severity" -> severity) ++
@@ -485,16 +446,14 @@
JSON.optional("source" -> source)
}
- object DiagnosticSeverity
- {
+ object DiagnosticSeverity {
val Error = 1
val Warning = 2
val Information = 3
val Hint = 4
}
- object PublishDiagnostics
- {
+ object PublishDiagnostics {
def apply(file: JFile, diagnostics: List[Diagnostic]): JSON.T =
Notification("textDocument/publishDiagnostics",
JSON.Object("uri" -> Url.print_file(file), "diagnostics" -> diagnostics.map(_.json)))
@@ -503,15 +462,13 @@
/* decorations */
- sealed case class Decoration_Options(range: Line.Range, hover_message: List[MarkedString])
- {
+ sealed case class Decoration_Options(range: Line.Range, hover_message: List[MarkedString]) {
def json: JSON.T =
JSON.Object("range" -> Range.compact(range)) ++
JSON.optional("hover_message" -> MarkedStrings.json(hover_message))
}
- sealed case class Decoration(decorations: List[(String, List[Decoration_Options])])
- {
+ sealed case class Decoration(decorations: List[(String, List[Decoration_Options])]) {
def json(file: JFile): JSON.T =
Notification("PIDE/decoration",
JSON.Object(
@@ -526,8 +483,7 @@
/* caret update: bidirectional */
- object Caret_Update
- {
+ object Caret_Update {
def apply(node_pos: Line.Node_Position, focus: Boolean): JSON.T =
Notification("PIDE/caret_update",
JSON.Object(
@@ -553,8 +509,7 @@
/* dynamic output */
- object Dynamic_Output
- {
+ object Dynamic_Output {
def apply(content: String): JSON.T =
Notification("PIDE/dynamic_output", JSON.Object("content" -> content))
}
@@ -562,14 +517,13 @@
/* state output */
- object State_Output
- {
+ object State_Output {
def apply(id: Counter.ID, content: String, auto_update: Boolean): JSON.T =
- Notification("PIDE/state_output", JSON.Object("id" -> id, "content" -> content, "auto_update" -> auto_update))
+ Notification("PIDE/state_output",
+ JSON.Object("id" -> id, "content" -> content, "auto_update" -> auto_update))
}
- class State_Id_Notification(name: String)
- {
+ class State_Id_Notification(name: String) {
def unapply(json: JSON.T): Option[Counter.ID] =
json match {
case Notification(method, Some(params)) if method == name => JSON.long(params, "id")
@@ -582,8 +536,7 @@
object State_Locate extends State_Id_Notification("PIDE/state_locate")
object State_Update extends State_Id_Notification("PIDE/state_update")
- object State_Auto_Update
- {
+ object State_Auto_Update {
def unapply(json: JSON.T): Option[(Counter.ID, Boolean)] =
json match {
case Notification("PIDE/state_auto_update", Some(params)) =>
@@ -598,8 +551,7 @@
/* preview */
- object Preview_Request
- {
+ object Preview_Request {
def unapply(json: JSON.T): Option[(JFile, Int)] =
json match {
case Notification("PIDE/preview_request", Some(params)) =>
@@ -612,8 +564,7 @@
}
}
- object Preview_Response
- {
+ object Preview_Response {
def apply(file: JFile, column: Int, label: String, content: String): JSON.T =
Notification("PIDE/preview_response",
JSON.Object(
--- a/src/Tools/VSCode/src/preview_panel.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/VSCode/src/preview_panel.scala Fri Apr 01 17:06:10 2022 +0200
@@ -12,17 +12,14 @@
import java.io.{File => JFile}
-class Preview_Panel(resources: VSCode_Resources)
-{
+class Preview_Panel(resources: VSCode_Resources) {
private val pending = Synchronized(Map.empty[JFile, Int])
def request(file: JFile, column: Int): Unit =
pending.change(map => map + (file -> column))
- def flush(channel: Channel): Boolean =
- {
- pending.change_result(map =>
- {
+ def flush(channel: Channel): Boolean = {
+ pending.change_result(map => {
val map1 =
map.iterator.foldLeft(map) {
case (m, (file, column)) =>
--- a/src/Tools/VSCode/src/state_panel.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/VSCode/src/state_panel.scala Fri Apr 01 17:06:10 2022 +0200
@@ -10,20 +10,17 @@
import isabelle._
-object State_Panel
-{
+object State_Panel {
private val make_id = Counter.make()
private val instances = Synchronized(Map.empty[Counter.ID, State_Panel])
- def init(server: Language_Server): Unit =
- {
+ def init(server: Language_Server): Unit = {
val instance = new State_Panel(server)
instances.change(_ + (instance.id -> instance))
instance.init()
}
- def exit(id: Counter.ID): Unit =
- {
+ def exit(id: Counter.ID): Unit = {
instances.change(map =>
map.get(id) match {
case None => map
@@ -45,8 +42,7 @@
}
-class State_Panel private(val server: Language_Server)
-{
+class State_Panel private(val server: Language_Server) {
/* output */
val id: Counter.ID = State_Panel.make_id()
@@ -80,8 +76,7 @@
def locate(): Unit = print_state.locate_query()
- def update(): Unit =
- {
+ def update(): Unit = {
server.editor.current_node_snapshot(()) match {
case Some(snapshot) =>
(server.editor.current_command((), snapshot), print_state.get_location) match {
@@ -97,8 +92,7 @@
private val auto_update_enabled = Synchronized(true)
- def auto_update(set: Option[Boolean] = None): Unit =
- {
+ def auto_update(set: Option[Boolean] = None): Unit = {
val enabled =
auto_update_enabled.guarded_access(a =>
set match {
@@ -121,16 +115,14 @@
auto_update()
}
- def init(): Unit =
- {
+ def init(): Unit = {
server.session.commands_changed += main
server.session.caret_focus += main
server.editor.send_wait_dispatcher { print_state.activate() }
server.editor.send_dispatcher { auto_update() }
}
- def exit(): Unit =
- {
+ def exit(): Unit = {
output_active.change(_ => false)
server.session.commands_changed -= main
server.session.caret_focus -= main
--- a/src/Tools/VSCode/src/vscode_main.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/VSCode/src/vscode_main.scala Fri Apr 01 17:06:10 2022 +0200
@@ -12,8 +12,7 @@
import java.util.zip.ZipFile
-object VSCode_Main
-{
+object VSCode_Main {
/* vscodium command-line interface */
def server_log_path: Path =
@@ -32,8 +31,8 @@
server_log: Boolean = false,
verbose: Boolean = false,
background: Boolean = false,
- progress: Progress = new Progress): Process_Result =
- {
+ progress: Progress = new Progress
+): Process_Result = {
def platform_path(s: String): String = File.platform_path(Path.explode(s))
val args_json =
@@ -82,8 +81,7 @@
val MANIFEST: Path = Path.explode("MANIFEST")
- private def shasum_vsix(vsix_path: Path): String =
- {
+ private def shasum_vsix(vsix_path: Path): String = {
val name = "extension/MANIFEST.shasum"
def err(): Nothing = error("Cannot retrieve " + quote(name) + " from " + vsix_path)
if (vsix_path.is_file) {
@@ -100,14 +98,12 @@
else err()
}
- private def shasum_dir(dir: Path): Option[String] =
- {
+ private def shasum_dir(dir: Path): Option[String] = {
val path = dir + MANIFEST.shasum
if (path.is_file) Some(File.read(path)) else None
}
- def locate_extension(): Option[Path] =
- {
+ def locate_extension(): Option[Path] = {
val out = run_vscodium(List("--locate-extension", extension_name)).check.out
if (out.nonEmpty) Some(Path.explode(File.standard_path(out))) else None
}
@@ -122,8 +118,8 @@
def install_extension(
vsix_path: Path = default_vsix_path,
- progress: Progress = new Progress): Unit =
- {
+ progress: Progress = new Progress
+ ): Unit = {
val new_shasum = shasum_vsix(vsix_path)
val old_shasum = locate_extension().flatMap(shasum_dir)
val current = old_shasum.isDefined && old_shasum.get == new_shasum
@@ -159,8 +155,7 @@
}
"""
- def init_settings(): Unit =
- {
+ def init_settings(): Unit = {
if (!settings_path.is_file) {
Isabelle_System.make_directory(settings_path.dir)
File.write(settings_path, default_settings)
@@ -171,8 +166,8 @@
/* Isabelle tool wrapper */
val isabelle_tool =
- Isabelle_Tool("vscode", "Isabelle/VSCode interface wrapper", Scala_Project.here, args =>
- {
+ Isabelle_Tool("vscode", "Isabelle/VSCode interface wrapper", Scala_Project.here,
+ args => {
var logic_ancestor = ""
var console = false
var edit_extension = false
--- a/src/Tools/VSCode/src/vscode_model.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/VSCode/src/vscode_model.scala Fri Apr 01 17:06:10 2022 +0200
@@ -12,12 +12,10 @@
import java.io.{File => JFile}
-object VSCode_Model
-{
+object VSCode_Model {
/* decorations */
- object Decoration
- {
+ object Decoration {
def empty(typ: String): Decoration = Decoration(typ, Nil)
def ranges(typ: String, ranges: List[Text.Range]): Decoration =
@@ -28,13 +26,11 @@
/* content */
- object Content
- {
+ object Content {
val empty: Content = Content(Line.Document.empty)
}
- sealed case class Content(doc: Line.Document)
- {
+ sealed case class Content(doc: Line.Document) {
override def toString: String = doc.toString
def text_length: Text.Offset = doc.text_length
def text_range: Text.Range = doc.text_range
@@ -57,9 +53,11 @@
}).toList
}
- def init(session: Session, editor: Language_Server.Editor, node_name: Document.Node.Name)
- : VSCode_Model =
- {
+ def init(
+ session: Session,
+ editor: Language_Server.Editor,
+ node_name: Document.Node.Name
+ ): VSCode_Model = {
VSCode_Model(session, editor, node_name, Content.empty,
node_required = File_Format.registry.is_theory(node_name))
}
@@ -76,8 +74,8 @@
last_perspective: Document.Node.Perspective_Text = Document.Node.no_perspective_text,
pending_edits: List[Text.Edit] = Nil,
published_diagnostics: List[Text.Info[Command.Results]] = Nil,
- published_decorations: List[VSCode_Model.Decoration] = Nil) extends Document.Model
-{
+ published_decorations: List[VSCode_Model.Decoration] = Nil
+) extends Document.Model {
model =>
@@ -104,9 +102,10 @@
/* perspective */
- def node_perspective(doc_blobs: Document.Blobs, caret: Option[Line.Position])
- : (Boolean, Document.Node.Perspective_Text) =
- {
+ def node_perspective(
+ doc_blobs: Document.Blobs,
+ caret: Option[Line.Position]
+ ): (Boolean, Document.Node.Perspective_Text) = {
if (is_theory) {
val snapshot = model.snapshot()
@@ -158,8 +157,7 @@
/* edits */
- def change_text(text: String, range: Option[Line.Range] = None): Option[VSCode_Model] =
- {
+ def change_text(text: String, range: Option[Line.Range] = None): Option[VSCode_Model] = {
val insert = Line.normalize(text)
range match {
case None =>
@@ -184,9 +182,8 @@
unicode_symbols: Boolean,
doc_blobs: Document.Blobs,
file: JFile,
- caret: Option[Line.Position])
- : Option[((List[LSP.TextDocumentEdit], List[Document.Edit_Text]), VSCode_Model)] =
- {
+ caret: Option[Line.Position]
+ ): Option[((List[LSP.TextDocumentEdit], List[Document.Edit_Text]), VSCode_Model)] = {
val workspace_edits =
if (unicode_symbols && version.isDefined) {
val edits = content.recode_symbols
@@ -197,8 +194,7 @@
val (reparse, perspective) = node_perspective(doc_blobs, caret)
if (reparse || pending_edits.nonEmpty || last_perspective != perspective ||
- workspace_edits.nonEmpty)
- {
+ workspace_edits.nonEmpty) {
val prover_edits = node_edits(node_header, pending_edits, perspective)
val edits = (workspace_edits, prover_edits)
Some((edits, copy(pending_edits = Nil, last_perspective = perspective)))
@@ -209,9 +205,9 @@
/* publish annotations */
- def publish(rendering: VSCode_Rendering):
- (Option[List[Text.Info[Command.Results]]], Option[List[VSCode_Model.Decoration]], VSCode_Model) =
- {
+ def publish(
+ rendering: VSCode_Rendering
+ ): (Option[List[Text.Info[Command.Results]]], Option[List[VSCode_Model.Decoration]], VSCode_Model) = {
val diagnostics = rendering.diagnostics
val decorations =
if (node_visible) rendering.decorations
--- a/src/Tools/VSCode/src/vscode_rendering.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/VSCode/src/vscode_rendering.scala Fri Apr 01 17:06:10 2022 +0200
@@ -15,13 +15,14 @@
import scala.annotation.tailrec
-object VSCode_Rendering
-{
+object VSCode_Rendering {
/* decorations */
- private def color_decorations(prefix: String, types: Set[Rendering.Color.Value],
- colors: List[Text.Info[Rendering.Color.Value]]): List[VSCode_Model.Decoration] =
- {
+ private def color_decorations(
+ prefix: String,
+ types: Set[Rendering.Color.Value],
+ colors: List[Text.Info[Rendering.Color.Value]]
+ ): List[VSCode_Model.Decoration] = {
val color_ranges =
colors.foldLeft(Map.empty[Rendering.Color.Value, List[Text.Range]]) {
case (m, Text.Info(range, c)) => m + (c -> (range :: m.getOrElse(c, Nil)))
@@ -66,8 +67,7 @@
}
class VSCode_Rendering(snapshot: Document.Snapshot, val model: VSCode_Model)
- extends Rendering(snapshot, model.resources.options, model.session)
-{
+extends Rendering(snapshot, model.resources.options, model.session) {
rendering =>
def resources: VSCode_Resources = model.resources
@@ -86,8 +86,7 @@
/* completion */
- def completion(node_pos: Line.Node_Position, caret: Text.Offset): List[LSP.CompletionItem] =
- {
+ def completion(node_pos: Line.Node_Position, caret: Text.Offset): List[LSP.CompletionItem] = {
val doc = model.content.doc
val line = node_pos.pos.line
val unicode = node_pos.name.endsWith(".thy")
@@ -148,8 +147,7 @@
case _ => None
}).filterNot(info => info.info.is_empty)
- def diagnostics_output(results: List[Text.Info[Command.Results]]): List[LSP.Diagnostic] =
- {
+ def diagnostics_output(results: List[Text.Info[Command.Results]]): List[LSP.Diagnostic] = {
(for {
Text.Info(text_range, res) <- results.iterator
range = model.content.doc.range(text_range)
@@ -164,8 +162,7 @@
/* text color */
- def text_color(range: Text.Range): List[Text.Info[Rendering.Color.Value]] =
- {
+ def text_color(range: Text.Range): List[Text.Info[Rendering.Color.Value]] = {
snapshot.select(range, Rendering.text_color_elements, _ =>
{
case Text.Info(_, XML.Elem(Markup(name, props), _)) =>
@@ -180,11 +177,13 @@
private sealed case class Color_Info(
color: Rendering.Color.Value, offset: Text.Offset, end_offset: Text.Offset, end_line: Int)
- def text_overview_color: List[Text.Info[Rendering.Color.Value]] =
- {
- @tailrec def loop(offset: Text.Offset, line: Int, lines: List[Line], colors: List[Color_Info])
- : List[Text.Info[Rendering.Color.Value]] =
- {
+ def text_overview_color: List[Text.Info[Rendering.Color.Value]] = {
+ @tailrec def loop(
+ offset: Text.Offset,
+ line: Int,
+ lines: List[Line],
+ colors: List[Color_Info]
+ ): List[Text.Info[Rendering.Color.Value]] = {
if (lines.nonEmpty) {
val end_offset = offset + lines.head.text.length
val colors1 =
@@ -235,8 +234,7 @@
dotted(model.content.text_range)))).flatten :::
List(VSCode_Spell_Checker.decoration(rendering))
- def decoration_output(decoration: List[VSCode_Model.Decoration]): LSP.Decoration =
- {
+ def decoration_output(decoration: List[VSCode_Model.Decoration]): LSP.Decoration = {
val entries =
for (deco <- decoration)
yield {
@@ -260,9 +258,11 @@
/* hyperlinks */
- def hyperlink_source_file(source_name: String, line1: Int, range: Symbol.Range)
- : Option[Line.Node_Range] =
- {
+ def hyperlink_source_file(
+ source_name: String,
+ line1: Int,
+ range: Symbol.Range
+ ): Option[Line.Node_Range] = {
for {
platform_path <- resources.source_file(source_name)
file <-
@@ -285,8 +285,7 @@
}
}
- def hyperlink_command(id: Document_ID.Generic, range: Symbol.Range): Option[Line.Node_Range] =
- {
+ def hyperlink_command(id: Document_ID.Generic, range: Symbol.Range): Option[Line.Node_Range] = {
if (snapshot.is_outdated) None
else
for {
@@ -309,8 +308,7 @@
case _ => None
}
- def hyperlinks(range: Text.Range): List[Line.Node_Range] =
- {
+ def hyperlinks(range: Text.Range): List[Line.Node_Range] = {
snapshot.cumulate[List[Line.Node_Range]](
range, Nil, VSCode_Rendering.hyperlink_elements, _ =>
{
--- a/src/Tools/VSCode/src/vscode_resources.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/VSCode/src/vscode_resources.scala Fri Apr 01 17:06:10 2022 +0200
@@ -14,8 +14,7 @@
import scala.util.parsing.input.Reader
-object VSCode_Resources
-{
+object VSCode_Resources {
/* internal state */
sealed case class State(
@@ -23,8 +22,8 @@
caret: Option[(JFile, Line.Position)] = None,
overlays: Document.Overlays = Document.Overlays.empty,
pending_input: Set[JFile] = Set.empty,
- pending_output: Set[JFile] = Set.empty)
- {
+ pending_output: Set[JFile] = Set.empty
+ ) {
def update_models(changed: Iterable[(JFile, VSCode_Model)]): State =
copy(
models = models ++ changed,
@@ -63,18 +62,16 @@
/* caret */
- sealed case class Caret(file: JFile, model: VSCode_Model, offset: Text.Offset)
- {
+ sealed case class Caret(file: JFile, model: VSCode_Model, offset: Text.Offset) {
def node_name: Document.Node.Name = model.node_name
}
}
class VSCode_Resources(
- val options: Options,
- session_base_info: Sessions.Base_Info,
- log: Logger = No_Logger)
- extends Resources(session_base_info.sessions_structure, session_base_info.check.base, log = log)
-{
+ val options: Options,
+ session_base_info: Sessions.Base_Info,
+ log: Logger = No_Logger)
+extends Resources(session_base_info.sessions_structure, session_base_info.check.base, log = log) {
resources =>
private val state = Synchronized(VSCode_Resources.State())
@@ -103,8 +100,7 @@
}
}
- override def append(dir: String, source_path: Path): String =
- {
+ override def append(dir: String, source_path: Path): String = {
val path = source_path.expand
if (dir == "" || path.is_absolute) File.platform_path(path)
else if (path.is_current) dir
@@ -121,8 +117,7 @@
/* file content */
- def read_file_content(name: Document.Node.Name): Option[String] =
- {
+ def read_file_content(name: Document.Node.Name): Option[String] = {
make_theory_content(name) orElse {
try { Some(Line.normalize(File.read(node_file(name)))) }
catch { case ERROR(_) => None }
@@ -135,8 +130,7 @@
case None => read_file_content(name)
}
- override def with_thy_reader[A](name: Document.Node.Name, f: Reader[Char] => A): A =
- {
+ override def with_thy_reader[A](name: Document.Node.Name, f: Reader[Char] => A): A = {
val file = node_file(name)
get_model(file) match {
case Some(model) => f(Scan.char_reader(model.content.text))
@@ -160,15 +154,14 @@
file: JFile,
version: Long,
text: String,
- range: Option[Line.Range] = None): Unit =
- {
- state.change(st =>
- {
- val model = st.models.getOrElse(file, VSCode_Model.init(session, editor, node_name(file)))
- val model1 =
- (model.change_text(text, range) getOrElse model).set_version(version).external(false)
- st.update_models(Some(file -> model1))
- })
+ range: Option[Line.Range] = None
+ ): Unit = {
+ state.change(st => {
+ val model = st.models.getOrElse(file, VSCode_Model.init(session, editor, node_name(file)))
+ val model1 =
+ (model.change_text(text, range) getOrElse model).set_version(version).external(false)
+ st.update_models(Some(file -> model1))
+ })
}
def close_model(file: JFile): Boolean =
@@ -179,17 +172,16 @@
})
def sync_models(changed_files: Set[JFile]): Unit =
- state.change(st =>
- {
- val changed_models =
- (for {
- (file, model) <- st.models.iterator
- if changed_files(file) && model.external_file
- text <- read_file_content(model.node_name)
- model1 <- model.change_text(text)
- } yield (file, model1)).toList
- st.update_models(changed_models)
- })
+ state.change(st => {
+ val changed_models =
+ (for {
+ (file, model) <- st.models.iterator
+ if changed_files(file) && model.external_file
+ text <- read_file_content(model.node_name)
+ model1 <- model.change_text(text)
+ } yield (file, model1)).toList
+ st.update_models(changed_models)
+ })
/* overlays */
@@ -209,85 +201,82 @@
def resolve_dependencies(
session: Session,
editor: Language_Server.Editor,
- file_watcher: File_Watcher): (Boolean, Boolean) =
- {
- state.change_result(st =>
- {
- /* theory files */
+ file_watcher: File_Watcher
+ ): (Boolean, Boolean) = {
+ state.change_result(st => {
+ /* theory files */
- val thys =
- (for ((_, model) <- st.models.iterator if model.is_theory)
- yield (model.node_name, Position.none)).toList
+ val thys =
+ (for ((_, model) <- st.models.iterator if model.is_theory)
+ yield (model.node_name, Position.none)).toList
- val thy_files1 = resources.dependencies(thys).theories
+ val thy_files1 = resources.dependencies(thys).theories
- val thy_files2 =
- (for {
- (_, model) <- st.models.iterator
- thy_name <- resources.make_theory_name(model.node_name)
- } yield thy_name).toList
+ val thy_files2 =
+ (for {
+ (_, model) <- st.models.iterator
+ thy_name <- resources.make_theory_name(model.node_name)
+ } yield thy_name).toList
- /* auxiliary files */
+ /* auxiliary files */
- val stable_tip_version =
- if (st.models.forall(entry => entry._2.is_stable))
- session.get_state().stable_tip_version
- else None
+ val stable_tip_version =
+ if (st.models.forall(entry => entry._2.is_stable))
+ session.get_state().stable_tip_version
+ else None
- val aux_files =
- stable_tip_version match {
- case Some(version) => undefined_blobs(version.nodes)
- case None => Nil
- }
+ val aux_files =
+ stable_tip_version match {
+ case Some(version) => undefined_blobs(version.nodes)
+ case None => Nil
+ }
- /* loaded models */
+ /* loaded models */
- val loaded_models =
- (for {
- node_name <- thy_files1.iterator ++ thy_files2.iterator ++ aux_files.iterator
- file = node_file(node_name)
- if !st.models.isDefinedAt(file)
- text <- { file_watcher.register_parent(file); read_file_content(node_name) }
- }
- yield {
- val model = VSCode_Model.init(session, editor, node_name)
- val model1 = (model.change_text(text) getOrElse model).external(true)
- (file, model1)
- }).toList
+ val loaded_models =
+ (for {
+ node_name <- thy_files1.iterator ++ thy_files2.iterator ++ aux_files.iterator
+ file = node_file(node_name)
+ if !st.models.isDefinedAt(file)
+ text <- { file_watcher.register_parent(file); read_file_content(node_name) }
+ }
+ yield {
+ val model = VSCode_Model.init(session, editor, node_name)
+ val model1 = (model.change_text(text) getOrElse model).external(true)
+ (file, model1)
+ }).toList
- val invoke_input = loaded_models.nonEmpty
- val invoke_load = stable_tip_version.isEmpty
+ val invoke_input = loaded_models.nonEmpty
+ val invoke_load = stable_tip_version.isEmpty
- ((invoke_input, invoke_load), st.update_models(loaded_models))
- })
+ ((invoke_input, invoke_load), st.update_models(loaded_models))
+ })
}
/* pending input */
- def flush_input(session: Session, channel: Channel): Unit =
- {
- state.change(st =>
- {
- val changed_models =
- (for {
- file <- st.pending_input.iterator
- model <- st.models.get(file)
- (edits, model1) <-
- model.flush_edits(false, st.document_blobs, file, st.get_caret(file))
- } yield (edits, (file, model1))).toList
+ def flush_input(session: Session, channel: Channel): Unit = {
+ state.change(st => {
+ val changed_models =
+ (for {
+ file <- st.pending_input.iterator
+ model <- st.models.get(file)
+ (edits, model1) <-
+ model.flush_edits(false, st.document_blobs, file, st.get_caret(file))
+ } yield (edits, (file, model1))).toList
- for { ((workspace_edits, _), _) <- changed_models if workspace_edits.nonEmpty }
- channel.write(LSP.WorkspaceEdit(workspace_edits))
+ for { ((workspace_edits, _), _) <- changed_models if workspace_edits.nonEmpty }
+ channel.write(LSP.WorkspaceEdit(workspace_edits))
- session.update(st.document_blobs, changed_models.flatMap(res => res._1._2))
+ session.update(st.document_blobs, changed_models.flatMap(res => res._1._2))
- st.copy(
- models = st.models ++ changed_models.iterator.map(_._2),
- pending_input = Set.empty)
- })
+ st.copy(
+ models = st.models ++ changed_models.iterator.map(_._2),
+ pending_input = Set.empty)
+ })
}
@@ -300,38 +289,35 @@
state.change(st => st.copy(pending_output = st.pending_output ++
(for ((file, model) <- st.models.iterator if model.node_visible) yield file)))
- def flush_output(channel: Channel): Boolean =
- {
- state.change_result(st =>
- {
- val (postponed, flushed) =
- (for {
- file <- st.pending_output.iterator
- model <- st.models.get(file)
- } yield (file, model, model.rendering())).toList.partition(_._3.snapshot.is_outdated)
+ def flush_output(channel: Channel): Boolean = {
+ state.change_result(st => {
+ val (postponed, flushed) =
+ (for {
+ file <- st.pending_output.iterator
+ model <- st.models.get(file)
+ } yield (file, model, model.rendering())).toList.partition(_._3.snapshot.is_outdated)
- val changed_iterator =
- for {
- (file, model, rendering) <- flushed.iterator
- (changed_diags, changed_decos, model1) = model.publish(rendering)
- if changed_diags.isDefined || changed_decos.isDefined
+ val changed_iterator =
+ for {
+ (file, model, rendering) <- flushed.iterator
+ (changed_diags, changed_decos, model1) = model.publish(rendering)
+ if changed_diags.isDefined || changed_decos.isDefined
+ }
+ yield {
+ for (diags <- changed_diags)
+ channel.write(LSP.PublishDiagnostics(file, rendering.diagnostics_output(diags)))
+ if (pide_extensions) {
+ for (decos <- changed_decos)
+ channel.write(rendering.decoration_output(decos).json(file))
}
- yield {
- for (diags <- changed_diags)
- channel.write(LSP.PublishDiagnostics(file, rendering.diagnostics_output(diags)))
- if (pide_extensions) {
- for (decos <- changed_decos)
- channel.write(rendering.decoration_output(decos).json(file))
- }
- (file, model1)
- }
+ (file, model1)
+ }
- (postponed.nonEmpty,
- st.copy(
- models = st.models ++ changed_iterator,
- pending_output = postponed.map(_._1).toSet))
- }
- )
+ (postponed.nonEmpty,
+ st.copy(
+ models = st.models ++ changed_iterator,
+ pending_output = postponed.map(_._1).toSet))
+ })
}
@@ -354,8 +340,7 @@
def update_caret(caret: Option[(JFile, Line.Position)]): Unit =
state.change(_.update_caret(caret))
- def get_caret(): Option[VSCode_Resources.Caret] =
- {
+ def get_caret(): Option[VSCode_Resources.Caret] = {
val st = state.value
for {
(file, pos) <- st.caret
--- a/src/Tools/VSCode/src/vscode_spell_checker.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/VSCode/src/vscode_spell_checker.scala Fri Apr 01 17:06:10 2022 +0200
@@ -10,10 +10,8 @@
import isabelle._
-object VSCode_Spell_Checker
-{
- def decoration(rendering: VSCode_Rendering): VSCode_Model.Decoration =
- {
+object VSCode_Spell_Checker {
+ def decoration(rendering: VSCode_Rendering): VSCode_Model.Decoration = {
val model = rendering.model
val ranges =
(for {
@@ -28,8 +26,7 @@
def completion(rendering: VSCode_Rendering, caret: Text.Offset): Option[Completion.Result] =
rendering.resources.spell_checker.get.flatMap(_.completion(rendering, caret))
- def menu_items(rendering: VSCode_Rendering, caret: Text.Offset): List[LSP.CompletionItem] =
- {
+ def menu_items(rendering: VSCode_Rendering, caret: Text.Offset): List[LSP.CompletionItem] = {
val result =
for {
spell_checker <- rendering.resources.spell_checker.get
--- a/src/Tools/jEdit/jedit_main/isabelle_sidekick.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/jedit_main/isabelle_sidekick.scala Fri Apr 01 17:06:10 2022 +0200
@@ -19,31 +19,27 @@
import sidekick.{SideKickParser, SideKickParsedData, IAsset}
-object Isabelle_Sidekick
-{
+object Isabelle_Sidekick {
def int_to_pos(offset: Text.Offset): Position =
new Position {
def getOffset: Text.Offset = offset
override def toString: String = offset.toString
}
- def root_data(buffer: Buffer): SideKickParsedData =
- {
+ def root_data(buffer: Buffer): SideKickParsedData = {
val data = new SideKickParsedData(buffer.getName)
data.getAsset(data.root).setEnd(int_to_pos(buffer.getLength))
data
}
- class Keyword_Asset(keyword: String, text: String, range: Text.Range) extends IAsset
- {
+ class Keyword_Asset(keyword: String, text: String, range: Text.Range) extends IAsset {
private val css = GUI.imitate_font_css(GUI.label_font())
protected var _name: String = text
protected var _start: Position = int_to_pos(range.start)
protected var _end: Position = int_to_pos(range.stop)
override def getIcon: Icon = null
- override def getShortString: String =
- {
+ override def getShortString: String = {
val n = keyword.length
val s =
_name.indexOf(keyword) match {
@@ -67,9 +63,11 @@
class Asset(name: String, range: Text.Range) extends Keyword_Asset("", name, range)
- def swing_markup_tree(tree: Markup_Tree, parent: DefaultMutableTreeNode,
- swing_node: Text.Info[List[XML.Elem]] => DefaultMutableTreeNode): Unit =
- {
+ def swing_markup_tree(
+ tree: Markup_Tree,
+ parent: DefaultMutableTreeNode,
+ swing_node: Text.Info[List[XML.Elem]] => DefaultMutableTreeNode
+ ): Unit = {
for ((_, entry) <- tree.branches) {
val node = swing_node(Text.Info(entry.range, entry.markup))
swing_markup_tree(entry.subtree, node, swing_node)
@@ -79,8 +77,7 @@
}
-class Isabelle_Sidekick(name: String) extends SideKickParser(name)
-{
+class Isabelle_Sidekick(name: String) extends SideKickParser(name) {
override def supportsCompletion = false
@@ -91,8 +88,7 @@
def parser(buffer: Buffer, syntax: Outer_Syntax, data: SideKickParsedData): Boolean = false
- def parse(buffer: Buffer, error_source: errorlist.DefaultErrorSource): SideKickParsedData =
- {
+ def parse(buffer: Buffer, error_source: errorlist.DefaultErrorSource): SideKickParsedData = {
stopped = false
// FIXME lock buffer (!??)
@@ -113,18 +109,16 @@
class Isabelle_Sidekick_Structure(
- name: String,
- node_name: Buffer => Option[Document.Node.Name],
- parse: (Outer_Syntax, Document.Node.Name, CharSequence) => List[Document_Structure.Document])
- extends Isabelle_Sidekick(name)
-{
- override def parser(buffer: Buffer, syntax: Outer_Syntax, data: SideKickParsedData): Boolean =
- {
+ name: String,
+ node_name: Buffer => Option[Document.Node.Name],
+ parse: (Outer_Syntax, Document.Node.Name, CharSequence) => List[Document_Structure.Document]
+) extends Isabelle_Sidekick(name) {
+ override def parser(buffer: Buffer, syntax: Outer_Syntax, data: SideKickParsedData): Boolean = {
def make_tree(
parent: DefaultMutableTreeNode,
offset: Text.Offset,
- documents: List[Document_Structure.Document]): Unit =
- {
+ documents: List[Document_Structure.Document]
+ ): Unit = {
documents.foldLeft(offset) {
case (i, document) =>
document match {
@@ -178,10 +172,8 @@
(_, _, text) => Document_Structure.parse_ml_sections(true, text))
-class Isabelle_Sidekick_Markup extends Isabelle_Sidekick("isabelle-markup")
-{
- override def parser(buffer: Buffer, syntax: Outer_Syntax, data: SideKickParsedData): Boolean =
- {
+class Isabelle_Sidekick_Markup extends Isabelle_Sidekick("isabelle-markup") {
+ override def parser(buffer: Buffer, syntax: Outer_Syntax, data: SideKickParsedData): Boolean = {
val opt_snapshot =
Document_Model.get(buffer) match {
case Some(model) if model.is_theory => Some(model.snapshot())
@@ -195,19 +187,18 @@
snapshot.version, command, Command.Markup_Index.markup,
command.range, Markup.Elements.full)
Isabelle_Sidekick.swing_markup_tree(markup, data.root,
- (info: Text.Info[List[XML.Elem]]) =>
- {
- val range = info.range + command_start
- val content = command.source(info.range).replace('\n', ' ')
- val info_text = Pretty.formatted(Pretty.fbreaks(info.info), margin = 40.0).mkString
+ (info: Text.Info[List[XML.Elem]]) => {
+ val range = info.range + command_start
+ val content = command.source(info.range).replace('\n', ' ')
+ val info_text = Pretty.formatted(Pretty.fbreaks(info.info), margin = 40.0).mkString
- new DefaultMutableTreeNode(
- new Isabelle_Sidekick.Asset(command.toString, range) {
- override def getShortString: String = content
- override def getLongString: String = info_text
- override def toString: String = quote(content) + " " + range.toString
- })
- })
+ new DefaultMutableTreeNode(
+ new Isabelle_Sidekick.Asset(command.toString, range) {
+ override def getShortString: String = content
+ override def getLongString: String = info_text
+ override def toString: String = quote(content) + " " + range.toString
+ })
+ })
}
true
case None => false
@@ -216,16 +207,14 @@
}
-class Isabelle_Sidekick_News extends Isabelle_Sidekick("isabelle-news")
-{
+class Isabelle_Sidekick_News extends Isabelle_Sidekick("isabelle-news") {
private val Heading1 = """^New in (.*)\w*$""".r
private val Heading2 = """^\*\*\*\w*(.*)\w*\*\*\*\w*$""".r
private def make_node(s: String, start: Text.Offset, stop: Text.Offset): DefaultMutableTreeNode =
new DefaultMutableTreeNode(new Isabelle_Sidekick.Asset(s, Text.Range(start, stop)))
- override def parser(buffer: Buffer, syntax: Outer_Syntax, data: SideKickParsedData): Boolean =
- {
+ override def parser(buffer: Buffer, syntax: Outer_Syntax, data: SideKickParsedData): Boolean = {
var offset = 0
var end_offset = 0
@@ -270,8 +259,7 @@
}
}
-class Isabelle_Sidekick_Bibtex extends SideKickParser("bibtex")
-{
+class Isabelle_Sidekick_Bibtex extends SideKickParser("bibtex") {
override def supportsCompletion = false
private class Asset(label: String, label_html: String, range: Text.Range, source: String)
@@ -280,8 +268,7 @@
override def getLongString: String = source
}
- def parse(buffer: Buffer, error_source: errorlist.DefaultErrorSource): SideKickParsedData =
- {
+ def parse(buffer: Buffer, error_source: errorlist.DefaultErrorSource): SideKickParsedData = {
val data = Isabelle_Sidekick.root_data(buffer)
try {
--- a/src/Tools/jEdit/jedit_main/scala_console.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/jedit_main/scala_console.scala Fri Apr 01 17:06:10 2022 +0200
@@ -15,8 +15,7 @@
import java.io.{OutputStream, Writer, PrintWriter}
-class Scala_Console extends Shell("Scala")
-{
+class Scala_Console extends Shell("Scala") {
/* global state -- owned by GUI thread */
@volatile private var interpreters = Map.empty[Console, Interpreter]
@@ -25,12 +24,10 @@
@volatile private var global_out: Output = null
@volatile private var global_err: Output = null
- private val console_stream = new OutputStream
- {
+ private val console_stream = new OutputStream {
val buf = new StringBuilder(100)
- override def flush(): Unit =
- {
+ override def flush(): Unit = {
val s = buf.synchronized { val s = buf.toString; buf.clear(); s }
val str = UTF8.decode_permissive(s)
GUI_Thread.later {
@@ -42,29 +39,25 @@
override def close(): Unit = flush()
- def write(byte: Int): Unit =
- {
+ def write(byte: Int): Unit = {
val c = byte.toChar
buf.synchronized { buf.append(c) }
if (c == '\n') flush()
}
}
- private val console_writer = new Writer
- {
+ private val console_writer = new Writer {
def flush(): Unit = console_stream.flush()
def close(): Unit = console_stream.flush()
- def write(cbuf: Array[Char], off: Int, len: Int): Unit =
- {
+ def write(cbuf: Array[Char], off: Int, len: Int): Unit = {
if (len > 0) {
UTF8.bytes(new String(cbuf.slice(off, off + len))).foreach(console_stream.write(_))
}
}
}
- private def with_console[A](console: Console, out: Output, err: Output)(e: => A): A =
- {
+ private def with_console[A](console: Console, out: Output, err: Output)(e: => A): A = {
global_console = console
global_out = out
global_err = if (err == null) out else err
@@ -81,8 +74,7 @@
}
}
- private def report_error(str: String): Unit =
- {
+ private def report_error(str: String): Unit = {
if (global_console == null || global_err == null) isabelle.Output.writeln(str)
else GUI_Thread.later { global_err.print(global_console.getErrorColor, str) }
}
@@ -95,8 +87,7 @@
private case class Execute(console: Console, out: Output, err: Output, command: String)
extends Request
- private class Interpreter
- {
+ private class Interpreter {
private val running = Synchronized[Option[Thread]](None)
def interrupt(): Unit = running.change(opt => { opt.foreach(_.interrupt()); opt })
@@ -106,8 +97,7 @@
print_writer = new PrintWriter(console_writer, true),
class_loader = new JARClassLoader)
- val thread: Consumer_Thread[Request] = Consumer_Thread.fork("Scala_Console")
- {
+ val thread: Consumer_Thread[Request] = Consumer_Thread.fork("Scala_Console") {
case Start(console) =>
interp.bind("view", "org.gjt.sp.jedit.View", console.getView)
interp.bind("console", "console.Console", console)
@@ -136,15 +126,13 @@
/* jEdit console methods */
- override def openConsole(console: Console): Unit =
- {
+ override def openConsole(console: Console): Unit = {
val interp = new Interpreter
interp.thread.send(Start(console))
interpreters += (console -> interp)
}
- override def closeConsole(console: Console): Unit =
- {
+ override def closeConsole(console: Console): Unit = {
interpreters.get(console) match {
case Some(interp) =>
interp.interrupt()
@@ -154,8 +142,7 @@
}
}
- override def printInfoMessage(out: Output): Unit =
- {
+ override def printInfoMessage(out: Output): Unit = {
out.print(null,
"This shell evaluates Isabelle/Scala expressions.\n\n" +
"The contents of package isabelle and isabelle.jedit are imported.\n" +
@@ -165,15 +152,17 @@
" PIDE -- Isabelle/PIDE plugin (e.g. PIDE.session, PIDE.snapshot, PIDE.rendering)\n")
}
- override def printPrompt(console: Console, out: Output): Unit =
- {
+ override def printPrompt(console: Console, out: Output): Unit = {
out.writeAttrs(ConsolePane.colorAttributes(console.getInfoColor), "scala>")
out.writeAttrs(ConsolePane.colorAttributes(console.getPlainColor), " ")
}
override def execute(
- console: Console, input: String, out: Output, err: Output, command: String): Unit =
- {
+ console: Console,
+ input: String,
+ out: Output,
+ err: Output, command: String
+ ): Unit = {
interpreters(console).thread.send(Execute(console, out, err, command))
}
--- a/src/Tools/jEdit/src/active.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/active.scala Fri Apr 01 17:06:10 2022 +0200
@@ -11,10 +11,8 @@
import org.gjt.sp.jedit.{ServiceManager, View}
-object Active
-{
- abstract class Handler
- {
+object Active {
+ abstract class Handler {
def handle(
view: View, text: String, elem: XML.Elem,
doc_view: Document_View, snapshot: Document.Snapshot): Boolean
@@ -24,8 +22,7 @@
ServiceManager.getServiceNames(classOf[Handler]).toList
.map(ServiceManager.getService(classOf[Handler], _))
- def action(view: View, text: String, elem: XML.Elem): Unit =
- {
+ def action(view: View, text: String, elem: XML.Elem): Unit = {
GUI_Thread.require {}
Document_View.get(view.getTextArea) match {
@@ -40,12 +37,11 @@
}
}
- class Misc_Handler extends Active.Handler
- {
+ class Misc_Handler extends Active.Handler {
override def handle(
view: View, text: String, elem: XML.Elem,
- doc_view: Document_View, snapshot: Document.Snapshot): Boolean =
- {
+ doc_view: Document_View, snapshot: Document.Snapshot
+ ): Boolean = {
val text_area = doc_view.text_area
val model = doc_view.model
val buffer = model.buffer
--- a/src/Tools/jEdit/src/base_plugin.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/base_plugin.scala Fri Apr 01 17:06:10 2022 +0200
@@ -13,10 +13,8 @@
import org.gjt.sp.util.SyntaxUtilities
-class Base_Plugin extends EBPlugin
-{
- override def start(): Unit =
- {
+class Base_Plugin extends EBPlugin {
+ override def start(): Unit = {
Isabelle_System.init()
GUI.use_isabelle_fonts()
@@ -26,8 +24,7 @@
Syntax_Style.set_extender(Syntax_Style.Base_Extender)
}
- override def stop(): Unit =
- {
+ override def stop(): Unit = {
Syntax_Style.set_extender(new SyntaxUtilities.StyleExtender)
}
--- a/src/Tools/jEdit/src/completion_popup.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/completion_popup.scala Fri Apr 01 17:06:10 2022 +0200
@@ -24,12 +24,10 @@
import org.gjt.sp.util.StandardUtilities
-object Completion_Popup
-{
+object Completion_Popup {
/** items with HTML rendering **/
- private class Item(val item: Completion.Item)
- {
+ private class Item(val item: Completion.Item) {
private val html =
item.description match {
case a :: bs =>
@@ -44,12 +42,10 @@
/** jEdit text area **/
- object Text_Area
- {
+ object Text_Area {
private val key = new Object
- def apply(text_area: TextArea): Option[Completion_Popup.Text_Area] =
- {
+ def apply(text_area: TextArea): Option[Completion_Popup.Text_Area] = {
GUI_Thread.require {}
text_area.getClientProperty(key) match {
case text_area_completion: Completion_Popup.Text_Area => Some(text_area_completion)
@@ -74,8 +70,7 @@
case None => false
}
- def exit(text_area: JEditTextArea): Unit =
- {
+ def exit(text_area: JEditTextArea): Unit = {
GUI_Thread.require {}
apply(text_area) match {
case None =>
@@ -85,8 +80,7 @@
}
}
- def init(text_area: JEditTextArea): Completion_Popup.Text_Area =
- {
+ def init(text_area: JEditTextArea): Completion_Popup.Text_Area = {
exit(text_area)
val text_area_completion = new Text_Area(text_area)
text_area.putClientProperty(key, text_area_completion)
@@ -94,8 +88,7 @@
text_area_completion
}
- def dismissed(text_area: TextArea): Boolean =
- {
+ def dismissed(text_area: TextArea): Boolean = {
GUI_Thread.require {}
apply(text_area) match {
case Some(text_area_completion) => text_area_completion.dismissed()
@@ -104,8 +97,7 @@
}
}
- class Text_Area private(text_area: JEditTextArea)
- {
+ class Text_Area private(text_area: JEditTextArea) {
// owned by GUI thread
private var completion_popup: Option[Completion_Popup] = None
@@ -118,8 +110,7 @@
/* rendering */
- def rendering(rendering: JEdit_Rendering, line_range: Text.Range): Option[Text.Info[Color]] =
- {
+ def rendering(rendering: JEdit_Rendering, line_range: Text.Range): Option[Text.Info[Color]] = {
active_range match {
case Some(range) => range.try_restrict(line_range)
case None =>
@@ -151,8 +142,8 @@
def syntax_completion(
history: Completion.History,
explicit: Boolean,
- opt_rendering: Option[JEdit_Rendering]): Option[Completion.Result] =
- {
+ opt_rendering: Option[JEdit_Rendering]
+ ): Option[Completion.Result] = {
val buffer = text_area.getBuffer
val unicode = Isabelle_Encoding.is_active(buffer)
@@ -183,8 +174,7 @@
/* completion action: text area */
- private def insert(item: Completion.Item): Unit =
- {
+ private def insert(item: Completion.Item): Unit = {
GUI_Thread.require {}
val buffer = text_area.getBuffer
@@ -229,8 +219,8 @@
immediate: Boolean = false,
explicit: Boolean = false,
delayed: Boolean = false,
- word_only: Boolean = false): Boolean =
- {
+ word_only: Boolean = false
+ ): Boolean = {
val view = text_area.getView
val layered = view.getLayeredPane
val buffer = text_area.getBuffer
@@ -239,8 +229,7 @@
val history = PIDE.plugin.completion_history.value
val unicode = Isabelle_Encoding.is_active(buffer)
- def open_popup(result: Completion.Result): Unit =
- {
+ def open_popup(result: Completion.Result): Unit = {
val font =
painter.getFont.deriveFont(
Font_Info.main_size(PIDE.options.real("jedit_popup_font_scale")))
@@ -255,15 +244,12 @@
val items = result.items.map(new Item(_))
val completion =
- new Completion_Popup(Some(range), layered, loc2, font, items)
- {
- override def complete(item: Completion.Item): Unit =
- {
+ new Completion_Popup(Some(range), layered, loc2, font, items) {
+ override def complete(item: Completion.Item): Unit = {
PIDE.plugin.completion_history.update(item)
insert(item)
}
- override def propagate(evt: KeyEvent): Unit =
- {
+ override def propagate(evt: KeyEvent): Unit = {
if (view.getKeyEventInterceptor == null)
JEdit_Lib.propagate_key(view, evt)
else if (view.getKeyEventInterceptor == inner_key_listener) {
@@ -277,8 +263,7 @@
}
if (evt.getID == KeyEvent.KEY_TYPED) input(evt)
}
- override def shutdown(focus: Boolean): Unit =
- {
+ override def shutdown(focus: Boolean): Unit = {
if (view.getKeyEventInterceptor == inner_key_listener)
view.setKeyEventInterceptor(null)
if (focus) text_area.requestFocus()
@@ -298,8 +283,7 @@
val caret = text_area.getCaretPosition
val opt_rendering = Document_View.get(text_area).map(_.get_rendering())
val result0 = syntax_completion(history, explicit, opt_rendering)
- val (no_completion, semantic_completion) =
- {
+ val (no_completion, semantic_completion) = {
opt_rendering match {
case Some(rendering) =>
rendering.semantic_completion_result(history, unicode, result0.map(_.range),
@@ -309,8 +293,7 @@
}
if (no_completion) false
else {
- val result =
- {
+ val result = {
val result1 =
if (word_only) None
else Completion.Result.merge(history, semantic_completion, result0)
@@ -345,8 +328,7 @@
/* input key events */
- def input(evt: KeyEvent): Unit =
- {
+ def input(evt: KeyEvent): Unit = {
GUI_Thread.require {}
if (!evt.isConsumed) {
@@ -383,8 +365,7 @@
/* dismiss popup */
- def dismissed(): Boolean =
- {
+ def dismissed(): Boolean = {
GUI_Thread.require {}
completion_popup match {
@@ -406,8 +387,7 @@
private def activate(): Unit =
text_area.addKeyListener(outer_key_listener)
- private def deactivate(): Unit =
- {
+ private def deactivate(): Unit = {
dismissed()
text_area.removeKeyListener(outer_key_listener)
}
@@ -422,8 +402,8 @@
instant_popups: Boolean = false,
enter_adds_to_history: Boolean = true,
syntax: Outer_Syntax = Outer_Syntax.empty) extends
- HistoryTextField(name, instant_popups, enter_adds_to_history)
- {
+ HistoryTextField(name, instant_popups, enter_adds_to_history
+ ) {
text_field =>
// see https://forums.oracle.com/thread/1361677
@@ -435,8 +415,7 @@
/* dismiss */
- private def dismissed(): Boolean =
- {
+ private def dismissed(): Boolean = {
completion_popup match {
case Some(completion) =>
completion.hide_popup()
@@ -450,8 +429,7 @@
/* insert */
- private def insert(item: Completion.Item): Unit =
- {
+ private def insert(item: Completion.Item): Unit = {
GUI_Thread.require {}
val range = item.range
@@ -472,8 +450,7 @@
/* completion action: text field */
- def action(): Unit =
- {
+ def action(): Unit = {
GUI.layered_pane(text_field) match {
case Some(layered) if text_field.isEditable =>
val history = PIDE.plugin.completion_history.value
@@ -491,10 +468,8 @@
val items = result.items.map(new Item(_))
val completion =
- new Completion_Popup(None, layered, loc, text_field.getFont, items)
- {
- override def complete(item: Completion.Item): Unit =
- {
+ new Completion_Popup(None, layered, loc, text_field.getFont, items) {
+ override def complete(item: Completion.Item): Unit = {
PIDE.plugin.completion_history.update(item)
insert(item)
}
@@ -516,8 +491,7 @@
/* process key event */
- private def process(evt: KeyEvent): Unit =
- {
+ private def process(evt: KeyEvent): Unit = {
if (PIDE.options.bool("jedit_completion")) {
dismissed()
if (evt.getKeyChar != '\b') {
@@ -536,8 +510,7 @@
action()
}
- override def processKeyEvent(evt0: KeyEvent): Unit =
- {
+ override def processKeyEvent(evt0: KeyEvent): Unit = {
val evt = KeyEventWorkaround.processKeyEvent(evt0)
if (evt != null) {
evt.getID match {
@@ -564,8 +537,8 @@
layered: JLayeredPane,
location: Point,
font: Font,
- items: List[Completion_Popup.Item]) extends JPanel(new BorderLayout)
-{
+ items: List[Completion_Popup.Item]
+) extends JPanel(new BorderLayout) {
completion =>
GUI_Thread.require {}
@@ -595,16 +568,14 @@
JComponent.WHEN_ANCESTOR_OF_FOCUSED_COMPONENT,
JComponent.WHEN_IN_FOCUSED_WINDOW)) list_view.peer.setInputMap(cond, null)
- private def complete_selected(): Boolean =
- {
+ private def complete_selected(): Boolean = {
list_view.selection.items.toList match {
case item :: _ => complete(item.item); true
case _ => false
}
}
- private def move_items(n: Int): Unit =
- {
+ private def move_items(n: Int): Unit = {
val i = list_view.peer.getSelectedIndex
val j = ((i + n) min (items.length - 1)) max 0
if (i != j) {
@@ -613,8 +584,7 @@
}
}
- private def move_pages(n: Int): Unit =
- {
+ private def move_pages(n: Int): Unit = {
val page_size = list_view.peer.getVisibleRowCount - 1
move_items(page_size * n)
}
@@ -624,46 +594,44 @@
val inner_key_listener: KeyListener =
JEdit_Lib.key_listener(
- key_pressed = (e: KeyEvent) =>
- {
- if (!e.isConsumed) {
- e.getKeyCode match {
- case KeyEvent.VK_ENTER if PIDE.options.bool("jedit_completion_select_enter") =>
- if (complete_selected()) e.consume
- hide_popup()
- case KeyEvent.VK_TAB if PIDE.options.bool("jedit_completion_select_tab") =>
- if (complete_selected()) e.consume
- hide_popup()
- case KeyEvent.VK_ESCAPE =>
+ key_pressed = (e: KeyEvent) => {
+ if (!e.isConsumed) {
+ e.getKeyCode match {
+ case KeyEvent.VK_ENTER if PIDE.options.bool("jedit_completion_select_enter") =>
+ if (complete_selected()) e.consume
+ hide_popup()
+ case KeyEvent.VK_TAB if PIDE.options.bool("jedit_completion_select_tab") =>
+ if (complete_selected()) e.consume
+ hide_popup()
+ case KeyEvent.VK_ESCAPE =>
+ hide_popup()
+ e.consume
+ case KeyEvent.VK_UP | KeyEvent.VK_KP_UP if multi =>
+ move_items(-1)
+ e.consume
+ case KeyEvent.VK_DOWN | KeyEvent.VK_KP_DOWN if multi =>
+ move_items(1)
+ e.consume
+ case KeyEvent.VK_PAGE_UP if multi =>
+ move_pages(-1)
+ e.consume
+ case KeyEvent.VK_PAGE_DOWN if multi =>
+ move_pages(1)
+ e.consume
+ case _ =>
+ if (e.isActionKey || e.isAltDown || e.isMetaDown || e.isControlDown)
hide_popup()
- e.consume
- case KeyEvent.VK_UP | KeyEvent.VK_KP_UP if multi =>
- move_items(-1)
- e.consume
- case KeyEvent.VK_DOWN | KeyEvent.VK_KP_DOWN if multi =>
- move_items(1)
- e.consume
- case KeyEvent.VK_PAGE_UP if multi =>
- move_pages(-1)
- e.consume
- case KeyEvent.VK_PAGE_DOWN if multi =>
- move_pages(1)
- e.consume
- case _ =>
- if (e.isActionKey || e.isAltDown || e.isMetaDown || e.isControlDown)
- hide_popup()
- }
}
- propagate(e)
- },
+ }
+ propagate(e)
+ },
key_typed = propagate
)
list_view.peer.addKeyListener(inner_key_listener)
list_view.peer.addMouseListener(new MouseAdapter {
- override def mouseClicked(e: MouseEvent): Unit =
- {
+ override def mouseClicked(e: MouseEvent): Unit = {
if (complete_selected()) e.consume
hide_popup()
}
@@ -686,11 +654,9 @@
def active_range: Option[Text.Range] =
if (isDisplayable) opt_range else None
- private val popup =
- {
+ private val popup = {
val screen = GUI.screen_location(layered, location)
- val size =
- {
+ val size = {
val geometry = JEdit_Lib.window_geometry(completion, completion)
val bounds = JEdit_Rendering.popup_bounds
val w = geometry.width min (screen.bounds.width * bounds).toInt min layered.getWidth
@@ -700,14 +666,12 @@
new Popup(layered, completion, screen.relative(layered, size), size)
}
- private def show_popup(focus: Boolean): Unit =
- {
+ private def show_popup(focus: Boolean): Unit = {
popup.show
if (focus) list_view.requestFocus()
}
- private def hide_popup(): Unit =
- {
+ private def hide_popup(): Unit = {
shutdown(list_view.peer.isFocusOwner)
popup.hide
}
--- a/src/Tools/jEdit/src/context_menu.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/context_menu.scala Fri Apr 01 17:06:10 2022 +0200
@@ -18,8 +18,7 @@
import org.gjt.sp.jedit.textarea.JEditTextArea
-class Context_Menu extends DynamicContextMenuService
-{
+class Context_Menu extends DynamicContextMenuService {
def createMenu(text_area: JEditTextArea, evt: MouseEvent): Array[JMenuItem] =
if (evt == null) null
else {
--- a/src/Tools/jEdit/src/debugger_dockable.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/debugger_dockable.scala Fri Apr 01 17:06:10 2022 +0200
@@ -26,12 +26,10 @@
import org.gjt.sp.jedit.textarea.JEditTextArea
-object Debugger_Dockable
-{
+object Debugger_Dockable {
/* breakpoints */
- def get_breakpoint(text_area: JEditTextArea, offset: Text.Offset): Option[(Command, Long)] =
- {
+ def get_breakpoint(text_area: JEditTextArea, offset: Text.Offset): Option[(Command, Long)] = {
GUI_Thread.require {}
Document_View.get(text_area) match {
@@ -55,8 +53,7 @@
else Nil
}
-class Debugger_Dockable(view: View, position: String) extends Dockable(view, position)
-{
+class Debugger_Dockable(view: View, position: String) extends Dockable(view, position) {
GUI_Thread.require {}
private val debugger = PIDE.session.debugger
@@ -75,16 +72,14 @@
override def detach_operation: Option[() => Unit] = pretty_text_area.detach_operation
- private def handle_resize(): Unit =
- {
+ private def handle_resize(): Unit = {
GUI_Thread.require {}
pretty_text_area.resize(
Font_Info.main(PIDE.options.real("jedit_font_scale") * zoom.factor / 100))
}
- private def handle_update(): Unit =
- {
+ private def handle_update(): Unit = {
GUI_Thread.require {}
val new_snapshot = PIDE.editor.current_node_snapshot(view).getOrElse(current_snapshot)
@@ -122,8 +117,7 @@
def thread_selection(): Option[String] = tree_selection().map(_.thread_name)
- private def update_tree(threads: Debugger.Threads): Unit =
- {
+ private def update_tree(threads: Debugger.Threads): Unit = {
val thread_contexts =
(for ((a, b) <- threads.iterator)
yield Debugger.Context(a, b)).toList
@@ -163,8 +157,7 @@
tree.revalidate()
}
- def update_vals(): Unit =
- {
+ def update_vals(): Unit = {
tree_selection() match {
case Some(c) if c.stack_state.isDefined =>
debugger.print_vals(c, sml_button.selected, context_field.getText)
@@ -176,16 +169,14 @@
tree.addTreeSelectionListener(
new TreeSelectionListener {
- override def valueChanged(e: TreeSelectionEvent): Unit =
- {
+ override def valueChanged(e: TreeSelectionEvent): Unit = {
update_focus()
update_vals()
}
})
tree.addMouseListener(
new MouseAdapter {
- override def mouseClicked(e: MouseEvent): Unit =
- {
+ override def mouseClicked(e: MouseEvent): Unit = {
val click = tree.getPathForLocation(e.getX, e.getY)
if (click != null && e.getClickCount == 1)
update_focus()
@@ -230,10 +221,8 @@
tooltip = "Isabelle/ML context: type theory, Proof.context, Context.generic"
}
private val context_field =
- new Completion_Popup.History_Text_Field("isabelle-debugger-context")
- {
- override def processKeyEvent(evt: KeyEvent): Unit =
- {
+ new Completion_Popup.History_Text_Field("isabelle-debugger-context") {
+ override def processKeyEvent(evt: KeyEvent): Unit = {
if (evt.getID == KeyEvent.KEY_PRESSED && evt.getKeyCode == KeyEvent.VK_ENTER)
eval_expression()
super.processKeyEvent(evt)
@@ -247,10 +236,8 @@
tooltip = "Isabelle/ML or Standard ML expression"
}
private val expression_field =
- new Completion_Popup.History_Text_Field("isabelle-debugger-expression")
- {
- override def processKeyEvent(evt: KeyEvent): Unit =
- {
+ new Completion_Popup.History_Text_Field("isabelle-debugger-expression") {
+ override def processKeyEvent(evt: KeyEvent): Unit = {
if (evt.getID == KeyEvent.KEY_PRESSED && evt.getKeyCode == KeyEvent.VK_ENTER)
eval_expression()
super.processKeyEvent(evt)
@@ -266,8 +253,7 @@
reactions += { case ButtonClicked(_) => eval_expression() }
}
- private def eval_expression(): Unit =
- {
+ private def eval_expression(): Unit = {
context_field.addCurrentToHistory()
expression_field.addCurrentToHistory()
tree_selection() match {
@@ -303,8 +289,7 @@
override def focusGained(e: FocusEvent): Unit = update_focus()
})
- private def update_focus(): Unit =
- {
+ private def update_focus(): Unit = {
for (c <- tree_selection()) {
debugger.set_focus(c)
for {
@@ -340,8 +325,7 @@
}
}
- override def init(): Unit =
- {
+ override def init(): Unit = {
PIDE.session.global_options += main
PIDE.session.debugger_updates += main
debugger.init()
@@ -349,8 +333,7 @@
jEdit.propertiesChanged()
}
- override def exit(): Unit =
- {
+ override def exit(): Unit = {
PIDE.session.global_options -= main
PIDE.session.debugger_updates -= main
delay_resize.revoke()
--- a/src/Tools/jEdit/src/dockable.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/dockable.scala Fri Apr 01 17:06:10 2022 +0200
@@ -17,8 +17,7 @@
class Dockable(view: View, position: String)
- extends JPanel(new BorderLayout) with DefaultFocusComponent
-{
+extends JPanel(new BorderLayout) with DefaultFocusComponent {
if (position == DockableWindowManager.FLOATING)
setPreferredSize(new Dimension(500, 250))
@@ -32,14 +31,12 @@
protected def init(): Unit = {}
protected def exit(): Unit = {}
- override def addNotify(): Unit =
- {
+ override def addNotify(): Unit = {
super.addNotify()
init()
}
- override def removeNotify(): Unit =
- {
+ override def removeNotify(): Unit = {
exit()
super.removeNotify()
}
--- a/src/Tools/jEdit/src/document_model.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/document_model.scala Fri Apr 01 17:06:10 2022 +0200
@@ -21,15 +21,14 @@
import org.gjt.sp.jedit.buffer.{BufferAdapter, BufferListener, JEditBuffer}
-object Document_Model
-{
+object Document_Model {
/* document models */
sealed case class State(
models: Map[Document.Node.Name, Document_Model] = Map.empty,
buffer_models: Map[JEditBuffer, Buffer_Model] = Map.empty,
- overlays: Document.Overlays = Document.Overlays.empty)
- {
+ overlays: Document.Overlays = Document.Overlays.empty
+ ) {
def file_models_iterator: Iterator[(Document.Node.Name, File_Model)] =
for {
(node_name, model) <- models.iterator
@@ -43,9 +42,11 @@
blob <- model.get_blob
} yield (node_name -> blob)).toMap)
- def open_buffer(session: Session, node_name: Document.Node.Name, buffer: Buffer)
- : (Buffer_Model, State) =
- {
+ def open_buffer(
+ session: Session,
+ node_name: Document.Node.Name,
+ buffer: Buffer
+ ) : (Buffer_Model, State) = {
val old_model =
models.get(node_name) match {
case Some(file_model: File_Model) => Some(file_model)
@@ -58,8 +59,7 @@
buffer_models = buffer_models + (buffer -> buffer_model)))
}
- def close_buffer(buffer: JEditBuffer): State =
- {
+ def close_buffer(buffer: JEditBuffer): State = {
buffer_models.get(buffer) match {
case None => this
case Some(buffer_model) =>
@@ -113,31 +113,28 @@
/* sync external files */
- def sync_files(changed_files: Set[JFile]): Boolean =
- {
- state.change_result(st =>
- {
- val changed_models =
- (for {
- (node_name, model) <- st.file_models_iterator
- file <- model.file if changed_files(file)
- text <- PIDE.resources.read_file_content(node_name)
- if model.content.text != text
- } yield {
- val content = Document_Model.File_Content(text)
- val edits = Text.Edit.replace(0, model.content.text, text)
- (node_name, model.copy(content = content, pending_edits = model.pending_edits ::: edits))
- }).toList
- if (changed_models.isEmpty) (false, st)
- else (true, st.copy(models = changed_models.foldLeft(st.models)(_ + _)))
- })
+ def sync_files(changed_files: Set[JFile]): Boolean = {
+ state.change_result(st => {
+ val changed_models =
+ (for {
+ (node_name, model) <- st.file_models_iterator
+ file <- model.file if changed_files(file)
+ text <- PIDE.resources.read_file_content(node_name)
+ if model.content.text != text
+ } yield {
+ val content = Document_Model.File_Content(text)
+ val edits = Text.Edit.replace(0, model.content.text, text)
+ (node_name, model.copy(content = content, pending_edits = model.pending_edits ::: edits))
+ }).toList
+ if (changed_models.isEmpty) (false, st)
+ else (true, st.copy(models = changed_models.foldLeft(st.models)(_ + _)))
+ })
}
/* syntax */
- def syntax_changed(names: List[Document.Node.Name]): Unit =
- {
+ def syntax_changed(names: List[Document.Node.Name]): Unit = {
GUI_Thread.require {}
val models = state.value.models
@@ -149,8 +146,7 @@
/* init and exit */
- def init(session: Session, node_name: Document.Node.Name, buffer: Buffer): Buffer_Model =
- {
+ def init(session: Session, node_name: Document.Node.Name, buffer: Buffer): Buffer_Model = {
GUI_Thread.require {}
state.change_result(st =>
st.buffer_models.get(buffer) match {
@@ -164,8 +160,7 @@
})
}
- def exit(buffer: Buffer): Unit =
- {
+ def exit(buffer: Buffer): Unit = {
GUI_Thread.require {}
state.change(st =>
if (st.buffer_models.isDefinedAt(buffer)) {
@@ -176,8 +171,7 @@
else st)
}
- def provide_files(session: Session, files: List[(Document.Node.Name, String)]): Unit =
- {
+ def provide_files(session: Session, files: List[(Document.Node.Name, String)]): Unit = {
GUI_Thread.require {}
state.change(st =>
files.foldLeft(st) {
@@ -195,8 +189,10 @@
} yield node_name).toSet
def node_required(
- name: Document.Node.Name, toggle: Boolean = false, set: Boolean = false): Unit =
- {
+ name: Document.Node.Name,
+ toggle: Boolean = false,
+ set: Boolean = false
+ ) : Unit = {
GUI_Thread.require {}
val changed =
@@ -226,65 +222,61 @@
/* flushed edits */
- def flush_edits(hidden: Boolean, purge: Boolean): (Document.Blobs, List[Document.Edit_Text]) =
- {
+ def flush_edits(hidden: Boolean, purge: Boolean): (Document.Blobs, List[Document.Edit_Text]) = {
GUI_Thread.require {}
- state.change_result(st =>
- {
- val doc_blobs = st.document_blobs
+ state.change_result(st => {
+ val doc_blobs = st.document_blobs
- val buffer_edits =
- (for {
- (_, model) <- st.buffer_models.iterator
- edit <- model.flush_edits(doc_blobs, hidden).iterator
- } yield edit).toList
+ val buffer_edits =
+ (for {
+ (_, model) <- st.buffer_models.iterator
+ edit <- model.flush_edits(doc_blobs, hidden).iterator
+ } yield edit).toList
- val file_edits =
- (for {
- (node_name, model) <- st.file_models_iterator
- (edits, model1) <- model.flush_edits(doc_blobs, hidden)
- } yield (edits, node_name -> model1)).toList
+ val file_edits =
+ (for {
+ (node_name, model) <- st.file_models_iterator
+ (edits, model1) <- model.flush_edits(doc_blobs, hidden)
+ } yield (edits, node_name -> model1)).toList
- val model_edits = buffer_edits ::: file_edits.flatMap(_._1)
+ val model_edits = buffer_edits ::: file_edits.flatMap(_._1)
- val purge_edits =
- if (purge) {
- val purged =
- (for ((node_name, model) <- st.file_models_iterator)
- yield (node_name -> model.purge_edits(doc_blobs))).toList
+ val purge_edits =
+ if (purge) {
+ val purged =
+ (for ((node_name, model) <- st.file_models_iterator)
+ yield (node_name -> model.purge_edits(doc_blobs))).toList
- val imports =
- {
- val open_nodes =
- (for ((_, model) <- st.buffer_models.iterator) yield model.node_name).toList
- val touched_nodes = model_edits.map(_._1)
- val pending_nodes = for ((node_name, None) <- purged) yield node_name
- (open_nodes ::: touched_nodes ::: pending_nodes).map((_, Position.none))
- }
- val retain = PIDE.resources.dependencies(imports).theories.toSet
+ val imports = {
+ val open_nodes =
+ (for ((_, model) <- st.buffer_models.iterator) yield model.node_name).toList
+ val touched_nodes = model_edits.map(_._1)
+ val pending_nodes = for ((node_name, None) <- purged) yield node_name
+ (open_nodes ::: touched_nodes ::: pending_nodes).map((_, Position.none))
+ }
+ val retain = PIDE.resources.dependencies(imports).theories.toSet
- for ((node_name, Some(edits)) <- purged if !retain(node_name); edit <- edits)
- yield edit
- }
- else Nil
+ for ((node_name, Some(edits)) <- purged if !retain(node_name); edit <- edits)
+ yield edit
+ }
+ else Nil
- val st1 = st.copy(models = st.models ++ file_edits.map(_._2) -- purge_edits.map(_._1))
- PIDE.plugin.file_watcher.purge(
- (for {
- (_, model) <- st1.file_models_iterator
- file <- model.file
- } yield file.getParentFile).toSet)
+ val st1 = st.copy(models = st.models ++ file_edits.map(_._2) -- purge_edits.map(_._1))
+ PIDE.plugin.file_watcher.purge(
+ (for {
+ (_, model) <- st1.file_models_iterator
+ file <- model.file
+ } yield file.getParentFile).toSet)
- ((doc_blobs, model_edits ::: purge_edits), st1)
- })
+ ((doc_blobs, model_edits ::: purge_edits), st1)
+ })
}
/* file content */
- sealed case class File_Content(text: String)
- {
+ sealed case class File_Content(text: String) {
lazy val bytes: Bytes = Bytes(Symbol.encode(text))
lazy val chunk: Symbol.Text_Chunk = Symbol.Text_Chunk(text)
lazy val bibtex_entries: List[Text.Info[String]] =
@@ -295,8 +287,7 @@
/* HTTP preview */
- def open_preview(view: View, plain_text: Boolean): Unit =
- {
+ def open_preview(view: View, plain_text: Boolean): Unit = {
Document_Model.get(view.getBuffer) match {
case Some(model) =>
val url = Preview_Service.server_url(plain_text, model.node_name)
@@ -305,8 +296,7 @@
}
}
- object Preview_Service extends HTTP.Service("preview")
- {
+ object Preview_Service extends HTTP.Service("preview") {
service =>
private val plain_text_prefix = "plain_text="
@@ -341,15 +331,15 @@
}
}
-sealed abstract class Document_Model extends Document.Model
-{
+sealed abstract class Document_Model extends Document.Model {
/* perspective */
def document_view_ranges(snapshot: Document.Snapshot): List[Text.Range] = Nil
def node_perspective(
- doc_blobs: Document.Blobs, hidden: Boolean): (Boolean, Document.Node.Perspective_Text) =
- {
+ doc_blobs: Document.Blobs,
+ hidden: Boolean
+ ): (Boolean, Document.Node.Perspective_Text) = {
GUI_Thread.require {}
if (Isabelle.continuous_checking && is_theory) {
@@ -373,8 +363,7 @@
/* snapshot */
- @tailrec final def await_stable_snapshot(): Document.Snapshot =
- {
+ @tailrec final def await_stable_snapshot(): Document.Snapshot = {
val snapshot = this.snapshot()
if (snapshot.is_outdated) {
PIDE.options.seconds("editor_output_delay").sleep()
@@ -384,8 +373,7 @@
}
}
-object File_Model
-{
+object File_Model {
def empty(session: Session): File_Model =
File_Model(session, Document.Node.Name.empty, None, Document_Model.File_Content(""),
false, Document.Node.no_perspective_text, Nil)
@@ -395,8 +383,8 @@
text: String,
node_required: Boolean = false,
last_perspective: Document.Node.Perspective_Text = Document.Node.no_perspective_text,
- pending_edits: List[Text.Edit] = Nil): File_Model =
- {
+ pending_edits: List[Text.Edit] = Nil
+ ): File_Model = {
val file = JEdit_Lib.check_file(node_name.node)
file.foreach(PIDE.plugin.file_watcher.register_parent(_))
@@ -413,8 +401,8 @@
content: Document_Model.File_Content,
node_required: Boolean,
last_perspective: Document.Node.Perspective_Text,
- pending_edits: List[Text.Edit]) extends Document_Model
-{
+ pending_edits: List[Text.Edit]
+) extends Document_Model {
/* text */
def get_text(range: Text.Range): Option[String] =
@@ -453,9 +441,10 @@
Some(copy(content = content1, pending_edits = pending_edits1))
}
- def flush_edits(doc_blobs: Document.Blobs, hidden: Boolean)
- : Option[(List[Document.Edit_Text], File_Model)] =
- {
+ def flush_edits(
+ doc_blobs: Document.Blobs,
+ hidden: Boolean
+ ) : Option[(List[Document.Edit_Text], File_Model)] = {
val (reparse, perspective) = node_perspective(doc_blobs, hidden)
if (reparse || pending_edits.nonEmpty || last_perspective != perspective) {
val edits = node_edits(node_header, pending_edits, perspective)
@@ -481,8 +470,7 @@
}
case class Buffer_Model(session: Session, node_name: Document.Node.Name, buffer: Buffer)
- extends Document_Model
-{
+extends Document_Model {
/* text */
def get_text(range: Text.Range): Option[String] =
@@ -491,8 +479,7 @@
/* header */
- def node_header(): Document.Node.Header =
- {
+ def node_header(): Document.Node.Header = {
GUI_Thread.require {}
PIDE.resources.special_header(node_name) getOrElse
@@ -515,8 +502,7 @@
doc_view <- Document_View.get(text_area)
} yield doc_view
- override def document_view_ranges(snapshot: Document.Snapshot): List[Text.Range] =
- {
+ override def document_view_ranges(snapshot: Document.Snapshot): List[Text.Range] = {
GUI_Thread.require {}
(for {
@@ -581,8 +567,7 @@
/* pending edits */
- private object pending_edits
- {
+ private object pending_edits {
private val pending = new mutable.ListBuffer[Text.Edit]
private var last_perspective = Document.Node.no_perspective_text
@@ -629,17 +614,24 @@
/* buffer listener */
- private val buffer_listener: BufferListener = new BufferAdapter
- {
- override def contentInserted(buffer: JEditBuffer,
- start_line: Int, offset: Int, num_lines: Int, length: Int): Unit =
- {
+ private val buffer_listener: BufferListener = new BufferAdapter {
+ override def contentInserted(
+ buffer: JEditBuffer,
+ start_line: Int,
+ offset: Int,
+ num_lines: Int,
+ length: Int
+ ): Unit = {
pending_edits.edit(List(Text.Edit.insert(offset, buffer.getText(offset, length))))
}
- override def preContentRemoved(buffer: JEditBuffer,
- start_line: Int, offset: Int, num_lines: Int, removed_length: Int): Unit =
- {
+ override def preContentRemoved(
+ buffer: JEditBuffer,
+ start_line: Int,
+ offset: Int,
+ num_lines: Int,
+ removed_length: Int
+ ): Unit = {
pending_edits.edit(List(Text.Edit.remove(offset, buffer.getText(offset, removed_length))))
}
}
@@ -647,8 +639,7 @@
/* syntax */
- def syntax_changed(): Unit =
- {
+ def syntax_changed(): Unit = {
JEdit_Lib.buffer_line_manager(buffer).setFirstInvalidLineContext(0)
for (text_area <- JEdit_Lib.jedit_text_areas(buffer))
Untyped.method(Class.forName("org.gjt.sp.jedit.textarea.TextArea"), "foldStructureChanged").
@@ -656,8 +647,7 @@
buffer.invalidateCachedFoldLevels()
}
- def init_token_marker(): Unit =
- {
+ def init_token_marker(): Unit = {
Isabelle.buffer_token_marker(buffer) match {
case Some(marker) if marker != buffer.getTokenMarker =>
buffer.setTokenMarker(marker)
@@ -669,8 +659,7 @@
/* init */
- def init(old_model: Option[File_Model]): Buffer_Model =
- {
+ def init(old_model: Option[File_Model]): Buffer_Model = {
GUI_Thread.require {}
old_model match {
@@ -693,8 +682,7 @@
/* exit */
- def exit(): File_Model =
- {
+ def exit(): File_Model = {
GUI_Thread.require {}
buffer.removeBufferListener(buffer_listener)
--- a/src/Tools/jEdit/src/document_view.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/document_view.scala Fri Apr 01 17:06:10 2022 +0200
@@ -19,14 +19,12 @@
import org.gjt.sp.jedit.textarea.{JEditTextArea, TextArea, TextAreaExtension, TextAreaPainter}
-object Document_View
-{
+object Document_View {
/* document view of text area */
private val key = new Object
- def get(text_area: TextArea): Option[Document_View] =
- {
+ def get(text_area: TextArea): Option[Document_View] = {
GUI_Thread.require {}
text_area.getClientProperty(key) match {
case doc_view: Document_View => Some(doc_view)
@@ -34,8 +32,7 @@
}
}
- def exit(text_area: JEditTextArea): Unit =
- {
+ def exit(text_area: JEditTextArea): Unit = {
GUI_Thread.require {}
get(text_area) match {
case None =>
@@ -45,8 +42,7 @@
}
}
- def init(model: Buffer_Model, text_area: JEditTextArea): Document_View =
- {
+ def init(model: Buffer_Model, text_area: JEditTextArea): Document_View = {
exit(text_area)
val doc_view = new Document_View(model, text_area)
text_area.putClientProperty(key, doc_view)
@@ -56,8 +52,7 @@
}
-class Document_View(val model: Buffer_Model, val text_area: JEditTextArea)
-{
+class Document_View(val model: Buffer_Model, val text_area: JEditTextArea) {
private val session = model.session
def get_rendering(): JEdit_Rendering =
@@ -70,12 +65,10 @@
/* perspective */
- def perspective(snapshot: Document.Snapshot): Text.Perspective =
- {
+ def perspective(snapshot: Document.Snapshot): Text.Perspective = {
GUI_Thread.require {}
- val active_command =
- {
+ val active_command = {
val view = text_area.getView
if (view != null && view.getTextArea == text_area) {
PIDE.editor.current_command(view, snapshot) match {
@@ -105,12 +98,17 @@
Text.Perspective(active_command ::: visible_lines)
}
- private def update_view = new TextAreaExtension
- {
- override def paintScreenLineRange(gfx: Graphics2D,
- first_line: Int, last_line: Int, physical_lines: Array[Int],
- start: Array[Int], end: Array[Int], y: Int, line_height: Int): Unit =
- {
+ private def update_view = new TextAreaExtension {
+ override def paintScreenLineRange(
+ gfx: Graphics2D,
+ first_line: Int,
+ last_line: Int,
+ physical_lines: Array[Int],
+ start: Array[Int],
+ end: Array[Int],
+ y: Int,
+ line_height: Int
+ ): Unit = {
// no robust_body
PIDE.editor.invoke_generated()
}
@@ -119,12 +117,17 @@
/* gutter */
- private val gutter_painter = new TextAreaExtension
- {
- override def paintScreenLineRange(gfx: Graphics2D,
- first_line: Int, last_line: Int, physical_lines: Array[Int],
- start: Array[Int], end: Array[Int], y: Int, line_height: Int): Unit =
- {
+ private val gutter_painter = new TextAreaExtension {
+ override def paintScreenLineRange(
+ gfx: Graphics2D,
+ first_line: Int,
+ last_line: Int,
+ physical_lines: Array[Int],
+ start: Array[Int],
+ end: Array[Int],
+ y: Int,
+ line_height: Int
+ ): Unit = {
rich_text_area.robust_body(()) {
GUI_Thread.assert {}
@@ -145,8 +148,7 @@
case Some((icon, color)) =>
// icons within selection area
if (!gutter.isExpanded &&
- gutter.isSelectionAreaEnabled && sel_width >= 12 && line_height >= 12)
- {
+ gutter.isSelectionAreaEnabled && sel_width >= 12 && line_height >= 12) {
val x0 =
(FOLD_MARKER_SIZE + sel_width - border_width - icon.getIconWidth) max 10
val y0 =
@@ -173,11 +175,10 @@
private val key_listener =
JEdit_Lib.key_listener(
- key_pressed = (evt: KeyEvent) =>
- {
- if (evt.getKeyCode == KeyEvent.VK_ESCAPE && Isabelle.dismissed_popups(text_area.getView))
- evt.consume
- }
+ key_pressed = (evt: KeyEvent) => {
+ if (evt.getKeyCode == KeyEvent.VK_ESCAPE && Isabelle.dismissed_popups(text_area.getView))
+ evt.consume
+ }
)
@@ -228,8 +229,7 @@
/* activation */
- private def activate(): Unit =
- {
+ private def activate(): Unit = {
val painter = text_area.getPainter
painter.addExtension(TextAreaPainter.LOWEST_LAYER, update_view)
@@ -246,8 +246,7 @@
session.commands_changed += main
}
- private def deactivate(): Unit =
- {
+ private def deactivate(): Unit = {
val painter = text_area.getPainter
session.raw_edits -= main
--- a/src/Tools/jEdit/src/documentation_dockable.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/documentation_dockable.scala Fri Apr 01 17:06:10 2022 +0200
@@ -17,12 +17,10 @@
import org.gjt.sp.jedit.{View, OperatingSystem}
-class Documentation_Dockable(view: View, position: String) extends Dockable(view, position)
-{
+class Documentation_Dockable(view: View, position: String) extends Dockable(view, position) {
private val doc_contents = Doc.contents()
- private case class Node(string: String, entry: Doc.Entry)
- {
+ private case class Node(string: String, entry: Doc.Entry) {
override def toString: String = string
}
@@ -47,8 +45,7 @@
override def focusOnDefaultComponent(): Unit = tree.requestFocusInWindow
- private def action(node: DefaultMutableTreeNode): Unit =
- {
+ private def action(node: DefaultMutableTreeNode): Unit = {
node.getUserObject match {
case Node(_, Doc.Doc(_, _, path)) =>
PIDE.editor.goto_doc(view, path)
@@ -59,8 +56,7 @@
}
tree.addKeyListener(new KeyAdapter {
- override def keyPressed(e: KeyEvent): Unit =
- {
+ override def keyPressed(e: KeyEvent): Unit = {
if (e.getKeyCode == KeyEvent.VK_ENTER) {
e.consume
val path = tree.getSelectionPath
@@ -74,8 +70,7 @@
}
})
tree.addMouseListener(new MouseAdapter {
- override def mouseClicked(e: MouseEvent): Unit =
- {
+ override def mouseClicked(e: MouseEvent): Unit = {
val click = tree.getPathForLocation(e.getX, e.getY)
if (click != null && e.getClickCount == 1) {
(click.getLastPathComponent, tree.getLastSelectedPathComponent) match {
--- a/src/Tools/jEdit/src/fold_handling.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/fold_handling.scala Fri Apr 01 17:06:10 2022 +0200
@@ -18,12 +18,10 @@
import org.gjt.sp.jedit.buffer.{JEditBuffer, FoldHandler}
-object Fold_Handling
-{
+object Fold_Handling {
/* input: dynamic line context */
- class Fold_Handler extends FoldHandler("isabelle")
- {
+ class Fold_Handler extends FoldHandler("isabelle") {
override def equals(other: Any): Boolean =
other match {
case that: Fold_Handler => true
@@ -34,8 +32,10 @@
Token_Markup.Line_Context.after(buffer, line).structure.depth max 0
override def getPrecedingFoldLevels(
- buffer: JEditBuffer, line: Int, seg: Segment, level: Int): JList[Integer] =
- {
+ buffer: JEditBuffer,
+ line: Int,
+ seg: Segment, level: Int
+ ): JList[Integer] = {
val structure = Token_Markup.Line_Context.after(buffer, line).structure
val result =
if (line > 0 && structure.command)
@@ -52,16 +52,14 @@
/* output: static document rendering */
class Document_Fold_Handler(private val rendering: JEdit_Rendering)
- extends FoldHandler("isabelle-document")
- {
+ extends FoldHandler("isabelle-document") {
override def equals(other: Any): Boolean =
other match {
case that: Document_Fold_Handler => this.rendering == that.rendering
case _ => false
}
- override def getFoldLevel(buffer: JEditBuffer, line: Int, seg: Segment): Int =
- {
+ override def getFoldLevel(buffer: JEditBuffer, line: Int, seg: Segment): Int = {
def depth(i: Text.Offset): Int =
if (i < 0) 0
else {
--- a/src/Tools/jEdit/src/font_info.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/font_info.scala Fri Apr 01 17:06:10 2022 +0200
@@ -15,8 +15,7 @@
import org.gjt.sp.jedit.{jEdit, View}
-object Font_Info
-{
+object Font_Info {
/* size range */
val min_size = 5
@@ -38,10 +37,8 @@
/* incremental size change */
- object main_change
- {
- private def change_size(change: Float => Float): Unit =
- {
+ object main_change {
+ private def change_size(change: Float => Float): Unit = {
GUI_Thread.require {}
val size0 = main_size()
@@ -56,30 +53,24 @@
// owned by GUI thread
private var steps = 0
- private val delay = Delay.last(PIDE.options.seconds("editor_input_delay"), gui = true)
- {
- change_size(size =>
- {
- var i = size.round
- while (steps != 0 && i > 0) {
- if (steps > 0)
- { i += (i / 10) max 1; steps -= 1 }
- else
- { i -= (i / 10) max 1; steps += 1 }
- }
- steps = 0
- i.toFloat
- })
+ private val delay = Delay.last(PIDE.options.seconds("editor_input_delay"), gui = true) {
+ change_size(size => {
+ var i = size.round
+ while (steps != 0 && i > 0) {
+ if (steps > 0) { i += (i / 10) max 1; steps -= 1 }
+ else { i -= (i / 10) max 1; steps += 1 }
+ }
+ steps = 0
+ i.toFloat
+ })
}
- def step(i: Int): Unit =
- {
+ def step(i: Int): Unit = {
steps += i
delay.invoke()
}
- def reset(size: Float): Unit =
- {
+ def reset(size: Float): Unit = {
delay.revoke()
steps = 0
change_size(_ => size)
@@ -92,8 +83,6 @@
abstract class Zoom_Box extends GUI.Zoom_Box { tooltip = "Zoom factor for output font size" }
}
-sealed case class Font_Info(family: String, size: Float)
-{
+sealed case class Font_Info(family: String, size: Float) {
def font: Font = new Font(family, Font.PLAIN, size.round)
}
-
--- a/src/Tools/jEdit/src/graphview_dockable.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/graphview_dockable.scala Fri Apr 01 17:06:10 2022 +0200
@@ -18,8 +18,7 @@
import scala.swing.TextArea
-object Graphview_Dockable
-{
+object Graphview_Dockable {
/* implicit arguments -- owned by GUI thread */
private var implicit_snapshot = Document.Snapshot.init
@@ -28,8 +27,9 @@
private var implicit_graph = no_graph
private def set_implicit(
- snapshot: Document.Snapshot, graph: Exn.Result[Graph_Display.Graph]): Unit =
- {
+ snapshot: Document.Snapshot,
+ graph: Exn.Result[Graph_Display.Graph]
+ ): Unit = {
GUI_Thread.require {}
implicit_snapshot = snapshot
@@ -39,12 +39,14 @@
private def reset_implicit(): Unit =
set_implicit(Document.Snapshot.init, no_graph)
- class Handler extends Active.Handler
- {
+ class Handler extends Active.Handler {
override def handle(
- view: View, text: String, elem: XML.Elem,
- doc_view: Document_View, snapshot: Document.Snapshot): Boolean =
- {
+ view: View,
+ text: String,
+ elem: XML.Elem,
+ doc_view: Document_View,
+ snapshot: Document.Snapshot
+ ): Boolean = {
elem match {
case XML.Elem(Markup(Markup.GRAPHVIEW, _), body) =>
Isabelle_Thread.fork(name = "graphview") {
@@ -64,8 +66,7 @@
}
}
-class Graphview_Dockable(view: View, position: String) extends Dockable(view, position)
-{
+class Graphview_Dockable(view: View, position: String) extends Dockable(view, position) {
private val snapshot = Graphview_Dockable.implicit_snapshot
private val graph_result = Graphview_Dockable.implicit_graph
@@ -83,8 +84,7 @@
val graphview = new isabelle.graphview.Graphview(graph) {
def options: Options = PIDE.options.value
- override def make_tooltip(parent: JComponent, x: Int, y: Int, body: XML.Body): String =
- {
+ override def make_tooltip(parent: JComponent, x: Int, y: Int, body: XML.Body): String = {
Pretty_Tooltip.invoke(() =>
{
val model = File_Model.empty(PIDE.session)
@@ -128,8 +128,7 @@
set_content(graphview)
- override def focusOnDefaultComponent(): Unit =
- {
+ override def focusOnDefaultComponent(): Unit = {
graphview match {
case main_panel: isabelle.graphview.Main_Panel =>
main_panel.tree_panel.tree.requestFocusInWindow
@@ -152,14 +151,12 @@
}
}
- override def init(): Unit =
- {
+ override def init(): Unit = {
GUI.parent_window(this).foreach(_.addWindowFocusListener(window_focus_listener))
PIDE.session.global_options += main
}
- override def exit(): Unit =
- {
+ override def exit(): Unit = {
GUI.parent_window(this).foreach(_.removeWindowFocusListener(window_focus_listener))
PIDE.session.global_options -= main
}
--- a/src/Tools/jEdit/src/info_dockable.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/info_dockable.scala Fri Apr 01 17:06:10 2022 +0200
@@ -15,8 +15,7 @@
import org.gjt.sp.jedit.View
-object Info_Dockable
-{
+object Info_Dockable {
/* implicit arguments -- owned by GUI thread */
private var implicit_snapshot = Document.Snapshot.init
@@ -24,8 +23,10 @@
private var implicit_info: XML.Body = Nil
private def set_implicit(
- snapshot: Document.Snapshot, results: Command.Results, info: XML.Body): Unit =
- {
+ snapshot: Document.Snapshot,
+ results: Command.Results,
+ info: XML.Body
+ ): Unit = {
GUI_Thread.require {}
implicit_snapshot = snapshot
@@ -37,16 +38,18 @@
set_implicit(Document.Snapshot.init, Command.Results.empty, Nil)
def apply(
- view: View, snapshot: Document.Snapshot, results: Command.Results, info: XML.Body): Unit =
- {
+ view: View,
+ snapshot: Document.Snapshot,
+ results: Command.Results,
+ info: XML.Body
+ ): Unit = {
set_implicit(snapshot, results, info)
view.getDockableWindowManager.floatDockableWindow("isabelle-info")
}
}
-class Info_Dockable(view: View, position: String) extends Dockable(view, position)
-{
+class Info_Dockable(view: View, position: String) extends Dockable(view, position) {
/* component state -- owned by GUI thread */
private val snapshot = Info_Dockable.implicit_snapshot
@@ -71,8 +74,7 @@
private val zoom = new Font_Info.Zoom_Box { def changed = handle_resize() }
- private def handle_resize(): Unit =
- {
+ private def handle_resize(): Unit = {
GUI_Thread.require {}
pretty_text_area.resize(
@@ -103,15 +105,13 @@
case _: Session.Global_Options => GUI_Thread.later { handle_resize() }
}
- override def init(): Unit =
- {
+ override def init(): Unit = {
GUI.parent_window(this).map(_.addWindowFocusListener(window_focus_listener))
PIDE.session.global_options += main
handle_resize()
}
- override def exit(): Unit =
- {
+ override def exit(): Unit = {
GUI.parent_window(this).map(_.removeWindowFocusListener(window_focus_listener))
PIDE.session.global_options -= main
delay_resize.revoke()
--- a/src/Tools/jEdit/src/isabelle.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/isabelle.scala Fri Apr 01 17:06:10 2022 +0200
@@ -24,8 +24,7 @@
import org.jedit.options.CombinedOptions
-object Isabelle
-{
+object Isabelle {
/* editor modes */
val modes =
@@ -79,8 +78,7 @@
def mode_token_marker(mode: String): Option[TokenMarker] = mode_markers.get(mode)
- def buffer_token_marker(buffer: Buffer): Option[TokenMarker] =
- {
+ def buffer_token_marker(buffer: Buffer): Option[TokenMarker] = {
val mode = JEdit_Lib.buffer_mode(buffer)
if (mode == "isabelle") Some(new Token_Markup.Marker(mode, Some(buffer)))
else mode_token_marker(mode)
@@ -207,8 +205,7 @@
def reset_continuous_checking(): Unit = { continuous_checking = false }
def toggle_continuous_checking(): Unit = { continuous_checking = !continuous_checking }
- class Continuous_Checking extends CheckBox("Continuous checking")
- {
+ class Continuous_Checking extends CheckBox("Continuous checking") {
tooltip = "Continuous checking of proof document (visible and required parts)"
reactions += { case ButtonClicked(_) => continuous_checking = selected }
def load(): Unit = { selected = continuous_checking }
@@ -232,8 +229,7 @@
/* full screen */
// see toggleFullScreen() method in jEdit/org/gjt/sp/jedit/View.java
- def toggle_full_screen(view: View): Unit =
- {
+ def toggle_full_screen(view: View): Unit = {
if (PIDE.options.bool("jedit_toggle_full_screen") ||
Untyped.get[Boolean](view, "fullScreenMode")) view.toggleFullScreen()
else {
@@ -271,8 +267,7 @@
buffer.getStringProperty("autoIndent") == "full" &&
PIDE.options.bool(option)
- def indent_input(text_area: TextArea): Unit =
- {
+ def indent_input(text_area: TextArea): Unit = {
val buffer = text_area.getBuffer
val line = text_area.getCaretLine
val caret = text_area.getCaretPosition
@@ -292,8 +287,7 @@
}
}
- def newline(text_area: TextArea): Unit =
- {
+ def newline(text_area: TextArea): Unit = {
if (!text_area.isEditable()) text_area.getToolkit().beep()
else {
val buffer = text_area.getBuffer
@@ -324,8 +318,7 @@
}
}
- def insert_line_padding(text_area: JEditTextArea, text: String): Unit =
- {
+ def insert_line_padding(text_area: JEditTextArea, text: String): Unit = {
val buffer = text_area.getBuffer
JEdit_Lib.buffer_edit(buffer) {
val text1 =
@@ -347,8 +340,8 @@
text_area: TextArea,
padding: Boolean,
id: Document_ID.Generic,
- text: String): Unit =
- {
+ text: String
+ ): Unit = {
val buffer = text_area.getBuffer
if (!snapshot.is_outdated && text != "") {
(snapshot.find_command(id), Document_Model.get(buffer)) match {
@@ -385,8 +378,7 @@
/* formal entities */
- def goto_entity(view: View): Unit =
- {
+ def goto_entity(view: View): Unit = {
val text_area = view.getTextArea
for {
doc_view <- Document_View.get(text_area)
@@ -396,8 +388,7 @@
} link.info.follow(view)
}
- def select_entity(text_area: JEditTextArea): Unit =
- {
+ def select_entity(text_area: JEditTextArea): Unit = {
for (doc_view <- Document_View.get(text_area)) {
val rendering = doc_view.get_rendering()
val caret_range = JEdit_Lib.caret_range(text_area)
@@ -438,8 +429,7 @@
/* block styles */
- private def enclose_input(text_area: JEditTextArea, s1: String, s2: String): Unit =
- {
+ private def enclose_input(text_area: JEditTextArea, s1: String, s2: String): Unit = {
s1.foreach(text_area.userInput)
s2.foreach(text_area.userInput)
s2.foreach(_ => text_area.goToPrevCharacter(false))
@@ -454,8 +444,7 @@
/* antiquoted cartouche */
- def antiquoted_cartouche(text_area: TextArea): Unit =
- {
+ def antiquoted_cartouche(text_area: TextArea): Unit = {
val buffer = text_area.getBuffer
for {
doc_view <- Document_View.get(text_area)
@@ -485,8 +474,7 @@
/* spell-checker dictionary */
- def update_dictionary(text_area: JEditTextArea, include: Boolean, permanent: Boolean): Unit =
- {
+ def update_dictionary(text_area: JEditTextArea, include: Boolean, permanent: Boolean): Unit = {
for {
spell_checker <- PIDE.plugin.spell_checker.get
doc_view <- Document_View.get(text_area)
@@ -499,10 +487,8 @@
}
}
- def reset_dictionary(): Unit =
- {
- for (spell_checker <- PIDE.plugin.spell_checker.get)
- {
+ def reset_dictionary(): Unit = {
+ for (spell_checker <- PIDE.plugin.spell_checker.get) {
spell_checker.reset()
JEdit_Lib.jedit_views().foreach(_.repaint())
}
@@ -511,8 +497,7 @@
/* debugger */
- def toggle_breakpoint(text_area: JEditTextArea): Unit =
- {
+ def toggle_breakpoint(text_area: JEditTextArea): Unit = {
GUI_Thread.require {}
if (PIDE.session.debugger.is_active()) {
@@ -528,8 +513,7 @@
/* plugin options */
- def plugin_options(frame: Frame): Unit =
- {
+ def plugin_options(frame: Frame): Unit = {
GUI_Thread.require {}
jEdit.setProperty("Plugin Options.last", "isabelle-general")
new CombinedOptions(frame, 1)
@@ -538,8 +522,7 @@
/* popups */
- def dismissed_popups(view: View): Boolean =
- {
+ def dismissed_popups(view: View): Boolean = {
var dismissed = false
JEdit_Lib.jedit_text_areas(view).foreach(text_area =>
@@ -553,8 +536,7 @@
/* tooltips */
- def show_tooltip(view: View, control: Boolean): Unit =
- {
+ def show_tooltip(view: View, control: Boolean): Unit = {
GUI_Thread.require {}
val text_area = view.getTextArea
@@ -579,8 +561,9 @@
view: View,
range: Text.Range,
avoid_range: Text.Range = Text.Range.offside,
- which: String = "")(get: List[Text.Markup] => Option[Text.Markup]): Unit =
- {
+ which: String = "")(
+ get: List[Text.Markup] => Option[Text.Markup]
+ ): Unit = {
GUI_Thread.require {}
val text_area = view.getTextArea
@@ -602,15 +585,13 @@
def goto_last_error(view: View): Unit =
goto_error(view, JEdit_Lib.buffer_range(view.getBuffer))(_.lastOption)
- def goto_prev_error(view: View): Unit =
- {
+ def goto_prev_error(view: View): Unit = {
val caret_range = JEdit_Lib.caret_range(view.getTextArea)
val range = Text.Range(0, caret_range.stop)
goto_error(view, range, avoid_range = caret_range, which = "previous ")(_.lastOption)
}
- def goto_next_error(view: View): Unit =
- {
+ def goto_next_error(view: View): Unit = {
val caret_range = JEdit_Lib.caret_range(view.getTextArea)
val range = Text.Range(caret_range.start, view.getBuffer.getLength)
goto_error(view, range, avoid_range = caret_range, which = "next ")(_.headOption)
--- a/src/Tools/jEdit/src/isabelle_encoding.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/isabelle_encoding.scala Fri Apr 01 17:06:10 2022 +0200
@@ -19,8 +19,7 @@
import scala.io.{Codec, BufferedSource}
-object Isabelle_Encoding
-{
+object Isabelle_Encoding {
def is_active(buffer: JEditBuffer): Boolean =
buffer.getStringProperty(JEditBuffer.ENCODING).asInstanceOf[String] == "UTF-8-Isabelle"
@@ -28,12 +27,10 @@
if (is_active(buffer)) Symbol.decode(s) else s
}
-class Isabelle_Encoding extends Encoding
-{
+class Isabelle_Encoding extends Encoding {
private val BUFSIZE = 32768
- private def text_reader(in: InputStream, codec: Codec, strict: Boolean): Reader =
- {
+ private def text_reader(in: InputStream, codec: Codec, strict: Boolean): Reader = {
val source = (new BufferedSource(in)(codec)).mkString
if (strict && Codepoint.iterator(source).exists(Symbol.symbols.code_defined))
@@ -45,19 +42,16 @@
override def getTextReader(in: InputStream): Reader =
text_reader(in, UTF8.codec(), true)
- override def getPermissiveTextReader(in: InputStream): Reader =
- {
+ override def getPermissiveTextReader(in: InputStream): Reader = {
val codec = UTF8.codec()
codec.onMalformedInput(CodingErrorAction.REPLACE)
codec.onUnmappableCharacter(CodingErrorAction.REPLACE)
text_reader(in, codec, false)
}
- override def getTextWriter(out: OutputStream): Writer =
- {
+ override def getTextWriter(out: OutputStream): Writer = {
val buffer = new ByteArrayOutputStream(BUFSIZE) {
- override def flush(): Unit =
- {
+ override def flush(): Unit = {
val text = Symbol.encode(toString(UTF8.charset_name))
out.write(UTF8.bytes(text))
out.flush()
--- a/src/Tools/jEdit/src/isabelle_export.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/isabelle_export.scala Fri Apr 01 17:06:10 2022 +0200
@@ -17,17 +17,24 @@
import org.gjt.sp.jedit.browser.VFSBrowser
-object Isabelle_Export
-{
+object Isabelle_Export {
/* virtual file-system */
val vfs_prefix = "isabelle-export:"
- class VFS extends Isabelle_VFS(vfs_prefix,
- read = true, browse = true, low_latency = true, non_awt_session = true)
- {
- override def _listFiles(vfs_session: AnyRef, url: String, component: Component): Array[VFSFile] =
- {
+ class VFS
+ extends Isabelle_VFS(
+ vfs_prefix,
+ read = true,
+ browse = true,
+ low_latency = true,
+ non_awt_session = true
+ ) {
+ override def _listFiles(
+ vfs_session: AnyRef,
+ url: String,
+ component: Component
+ ): Array[VFSFile] = {
explode_url(url, component = component) match {
case None => null
case Some(elems) =>
@@ -63,8 +70,11 @@
}
override def _createInputStream(
- vfs_session: AnyRef, url: String, ignore_errors: Boolean, component: Component): InputStream =
- {
+ vfs_session: AnyRef,
+ url: String,
+ ignore_errors: Boolean,
+ component: Component
+ ): InputStream = {
explode_url(url, component = if (ignore_errors) null else component) match {
case None | Some(Nil) => Bytes.empty.stream()
case Some(theory :: name_elems) =>
@@ -86,8 +96,7 @@
/* open browser */
- def open_browser(view: View): Unit =
- {
+ def open_browser(view: View): Unit = {
val path =
PIDE.maybe_snapshot(view) match {
case None => ""
--- a/src/Tools/jEdit/src/isabelle_options.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/isabelle_options.scala Fri Apr 01 17:06:10 2022 +0200
@@ -12,12 +12,10 @@
import org.gjt.sp.jedit.{jEdit, AbstractOptionPane}
-abstract class Isabelle_Options(name: String) extends AbstractOptionPane(name)
-{
+abstract class Isabelle_Options(name: String) extends AbstractOptionPane(name) {
protected val components: List[(String, List[Option_Component])]
- override def _init(): Unit =
- {
+ override def _init(): Unit = {
val dummy_property = "options.isabelle.dummy"
for ((s, cs) <- components) {
@@ -30,15 +28,13 @@
}
}
- override def _save(): Unit =
- {
+ override def _save(): Unit = {
for ((_, cs) <- components) cs.foreach(_.save())
}
}
-class Isabelle_Options1 extends Isabelle_Options("isabelle-general")
-{
+class Isabelle_Options1 extends Isabelle_Options("isabelle-general") {
val options: JEdit_Options = PIDE.options
private val predefined =
@@ -51,8 +47,7 @@
}
-class Isabelle_Options2 extends Isabelle_Options("isabelle-rendering")
-{
+class Isabelle_Options2 extends Isabelle_Options("isabelle-rendering") {
private val predefined =
(for {
(name, opt) <- PIDE.options.value.options.toList
--- a/src/Tools/jEdit/src/isabelle_session.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/isabelle_session.scala Fri Apr 01 17:06:10 2022 +0200
@@ -17,8 +17,7 @@
import org.gjt.sp.jedit.browser.VFSBrowser
-object Isabelle_Session
-{
+object Isabelle_Session {
/* sessions structure */
def sessions_structure(): Sessions.Structure =
@@ -30,8 +29,7 @@
val vfs_prefix = "isabelle-session:"
class Session_Entry(name: String, path: String, marker: String)
- extends VFSFile(name, path, vfs_prefix + name, VFSFile.FILE, 0L, false)
- {
+ extends VFSFile(name, path, vfs_prefix + name, VFSFile.FILE, 0L, false) {
override def getPathMarker: String = marker
override def getExtendedAttribute(att: String): String =
@@ -39,11 +37,19 @@
else super.getExtendedAttribute(att)
}
- class VFS extends Isabelle_VFS(vfs_prefix,
- read = true, browse = true, low_latency = true, non_awt_session = true)
- {
- override def _listFiles(vfs_session: AnyRef, url: String, component: Component): Array[VFSFile] =
- {
+ class VFS
+ extends Isabelle_VFS(
+ vfs_prefix,
+ read = true,
+ browse = true,
+ low_latency = true,
+ non_awt_session = true
+ ) {
+ override def _listFiles(
+ vfs_session: AnyRef,
+ url: String,
+ component: Component
+ ): Array[VFSFile] = {
explode_url(url, component = component) match {
case None => null
case Some(elems) =>
@@ -55,8 +61,7 @@
sessions.chapters.get(chapter) match {
case None => null
case Some(infos) =>
- infos.map(info =>
- {
+ infos.map(info => {
val name = chapter + "/" + info.name
val path =
Position.File.unapply(info.pos) match {
@@ -80,8 +85,7 @@
/* open browser */
- def open_browser(view: View): Unit =
- {
+ def open_browser(view: View): Unit = {
val path =
PIDE.maybe_snapshot(view) match {
case None => ""
--- a/src/Tools/jEdit/src/isabelle_vfs.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/isabelle_vfs.scala Fri Apr 01 17:06:10 2022 +0200
@@ -33,15 +33,14 @@
(if (mkdir) VFS.MKDIR_CAP else 0) |
(if (low_latency) VFS.LOW_LATENCY_CAP else 0) |
(if (case_insensitive) VFS.CASE_INSENSITIVE_CAP else 0) |
- (if (non_awt_session) VFS.NON_AWT_SESSION_CAP else 0))
-{
+ (if (non_awt_session) VFS.NON_AWT_SESSION_CAP else 0)
+) {
/* URL structure */
def explode_name(s: String): List[String] = space_explode(getFileSeparator, s)
def implode_name(elems: Iterable[String]): String = elems.mkString(getFileSeparator.toString)
- def explode_url(url: String, component: Component = null): Option[List[String]] =
- {
+ def explode_url(url: String, component: Component = null): Option[List[String]] = {
Library.try_unprefix(prefix, url) match {
case Some(path) => Some(explode_name(path).filter(_.nonEmpty))
case None =>
@@ -51,8 +50,7 @@
}
def implode_url(elems: Iterable[String]): String = prefix + implode_name(elems)
- override def constructPath(parent: String, path: String): String =
- {
+ override def constructPath(parent: String, path: String): String = {
if (parent == "") path
else if (parent(parent.length - 1) == getFileSeparator || parent == prefix) parent + path
else parent + getFileSeparator + path
@@ -63,15 +61,13 @@
override def isMarkersFileSupported: Boolean = false
- def make_entry(path: String, is_dir: Boolean = false, size: Long = 0L): VFSFile =
- {
+ def make_entry(path: String, is_dir: Boolean = false, size: Long = 0L): VFSFile = {
val entry = explode_name(path).lastOption getOrElse ""
val url = prefix + path
new VFSFile(entry, url, url, if (is_dir) VFSFile.DIRECTORY else VFSFile.FILE, size, false)
}
- override def _getFile(vfs_session: AnyRef, url: String, component: Component): VFSFile =
- {
+ override def _getFile(vfs_session: AnyRef, url: String, component: Component): VFSFile = {
val parent = getParentOfPath(url)
if (parent == prefix) new VFSFile(prefix, prefix, prefix, VFSFile.DIRECTORY, 0L, false)
else {
--- a/src/Tools/jEdit/src/jedit_bibtex.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/jedit_bibtex.scala Fri Apr 01 17:06:10 2022 +0200
@@ -23,12 +23,10 @@
import org.gjt.sp.jedit.syntax.{Token => JEditToken, TokenMarker, TokenHandler}
-object JEdit_Bibtex
-{
+object JEdit_Bibtex {
/** context menu **/
- def context_menu(text_area0: JEditTextArea): List[JMenuItem] =
- {
+ def context_menu(text_area0: JEditTextArea): List[JMenuItem] = {
text_area0 match {
case text_area: TextArea =>
text_area.getBuffer match {
@@ -83,8 +81,7 @@
private val mode_rule_set = Token_Markup.mode_rule_set("bibtex")
private class Line_Context(val context: Option[Bibtex.Line_Context])
- extends TokenMarker.LineContext(mode_rule_set, null)
- {
+ extends TokenMarker.LineContext(mode_rule_set, null) {
override def hashCode: Int = context.hashCode
override def equals(that: Any): Boolean =
that match {
@@ -96,13 +93,14 @@
/* token marker */
- class Token_Marker extends TokenMarker
- {
+ class Token_Marker extends TokenMarker {
addRuleSet(mode_rule_set)
- override def markTokens(context: TokenMarker.LineContext,
- handler: TokenHandler, raw_line: Segment): TokenMarker.LineContext =
- {
+ override def markTokens(
+ context: TokenMarker.LineContext,
+ handler: TokenHandler,
+ raw_line: Segment
+ ): TokenMarker.LineContext = {
val line_ctxt =
context match {
case c: Line_Context => c.context
@@ -110,14 +108,12 @@
}
val line = if (raw_line == null) new Segment else raw_line
- def no_markup =
- {
+ def no_markup = {
val styled_token = (JEditToken.NULL, line.subSequence(0, line.count).toString)
(List(styled_token), new Line_Context(None))
}
- val context1 =
- {
+ val context1 = {
val (styled_tokens, context1) =
line_ctxt match {
case Some(ctxt) =>
--- a/src/Tools/jEdit/src/jedit_editor.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/jedit_editor.scala Fri Apr 01 17:06:10 2022 +0200
@@ -17,8 +17,7 @@
import org.gjt.sp.jedit.io.{VFSManager, VFSFile}
-class JEdit_Editor extends Editor[View]
-{
+class JEdit_Editor extends Editor[View] {
/* session */
override def session: Session = PIDE.session
@@ -62,8 +61,7 @@
override def current_node_snapshot(view: View): Option[Document.Snapshot] =
GUI_Thread.require { Document_Model.get(view.getBuffer).map(_.snapshot()) }
- override def node_snapshot(name: Document.Node.Name): Document.Snapshot =
- {
+ override def node_snapshot(name: Document.Node.Name): Document.Snapshot = {
GUI_Thread.require {}
Document_Model.get(name) match {
case Some(model) => model.snapshot()
@@ -71,8 +69,7 @@
}
}
- override def current_command(view: View, snapshot: Document.Snapshot): Option[Command] =
- {
+ override def current_command(view: View, snapshot: Document.Snapshot): Option[Command] = {
GUI_Thread.require {}
val text_area = view.getTextArea
@@ -105,8 +102,7 @@
/* navigation */
- def push_position(view: View): Unit =
- {
+ def push_position(view: View): Unit = {
val navigator = jEdit.getPlugin("ise.plugin.nav.NavigatorPlugin")
if (navigator != null) {
try { Untyped.method(navigator.getClass, "pushPosition", view.getClass).invoke(null, view) }
@@ -114,8 +110,7 @@
}
}
- def goto_buffer(focus: Boolean, view: View, buffer: Buffer, offset: Text.Offset): Unit =
- {
+ def goto_buffer(focus: Boolean, view: View, buffer: Buffer, offset: Text.Offset): Unit = {
GUI_Thread.require {}
push_position(view)
@@ -131,8 +126,7 @@
def goto_file(focus: Boolean, view: View, name: String): Unit =
goto_file(focus, view, Line.Node_Position.offside(name))
- def goto_file(focus: Boolean, view: View, pos: Line.Node_Position): Unit =
- {
+ def goto_file(focus: Boolean, view: View, pos: Line.Node_Position): Unit = {
GUI_Thread.require {}
push_position(view)
@@ -176,8 +170,7 @@
}
}
- def goto_doc(view: View, path: Path): Unit =
- {
+ def goto_doc(view: View, path: Path): Unit = {
if (path.is_pdf) Doc.view(path)
else goto_file(true, view, File.platform_path(path))
}
@@ -234,9 +227,12 @@
}
}
- def hyperlink_source_file(focus: Boolean, source_name: String, line1: Int, offset: Symbol.Offset)
- : Option[Hyperlink] =
- {
+ def hyperlink_source_file(
+ focus: Boolean,
+ source_name: String,
+ line1: Int,
+ offset: Symbol.Offset
+ ) : Option[Hyperlink] = {
for (name <- PIDE.resources.source_file(source_name)) yield {
def hyperlink(pos: Line.Position) =
hyperlink_file(focus, Line.Node_Position(name, pos))
@@ -256,16 +252,20 @@
}
override def hyperlink_command(
- focus: Boolean, snapshot: Document.Snapshot, id: Document_ID.Generic, offset: Symbol.Offset = 0)
- : Option[Hyperlink] =
- {
+ focus: Boolean,
+ snapshot: Document.Snapshot,
+ id: Document_ID.Generic,
+ offset: Symbol.Offset = 0
+ ) : Option[Hyperlink] = {
if (snapshot.is_outdated) None
else snapshot.find_command_position(id, offset).map(hyperlink_file(focus, _))
}
- def is_hyperlink_position(snapshot: Document.Snapshot,
- text_offset: Text.Offset, pos: Position.T): Boolean =
- {
+ def is_hyperlink_position(
+ snapshot: Document.Snapshot,
+ text_offset: Text.Offset,
+ pos: Position.T
+ ): Boolean = {
pos match {
case Position.Item_Id(id, range) if range.start > 0 =>
snapshot.find_command(id) match {
--- a/src/Tools/jEdit/src/jedit_lib.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/jedit_lib.scala Fri Apr 01 17:06:10 2022 +0200
@@ -24,8 +24,7 @@
import org.gjt.sp.jedit.textarea.{JEditTextArea, TextArea, TextAreaPainter}
-object JEdit_Lib
-{
+object JEdit_Lib {
/* jEdit directories */
def directories: List[JFile] =
@@ -36,14 +35,12 @@
private lazy val dummy_window = new JWindow
- final case class Window_Geometry(width: Int, height: Int, inner_width: Int, inner_height: Int)
- {
+ final case class Window_Geometry(width: Int, height: Int, inner_width: Int, inner_height: Int) {
def deco_width: Int = width - inner_width
def deco_height: Int = height - inner_height
}
- def window_geometry(outer: Container, inner: Component): Window_Geometry =
- {
+ def window_geometry(outer: Container, inner: Component): Window_Geometry = {
GUI_Thread.require {}
val old_content = dummy_window.getContentPane
@@ -79,8 +76,7 @@
def buffer_reader(buffer: JEditBuffer): CharSequenceReader =
Scan.char_reader(buffer.getSegment(0, buffer.getLength))
- def buffer_mode(buffer: JEditBuffer): String =
- {
+ def buffer_mode(buffer: JEditBuffer): String = {
val mode = buffer.getMode
if (mode == null) ""
else {
@@ -96,8 +92,7 @@
def buffer_file(buffer: Buffer): Option[JFile] = check_file(buffer_name(buffer))
- def buffer_undo_in_progress[A](buffer: JEditBuffer, body: => A): A =
- {
+ def buffer_undo_in_progress[A](buffer: JEditBuffer, body: => A): A = {
val undo_in_progress = buffer.isUndoInProgress
def set(b: Boolean): Unit = Untyped.set[Boolean](buffer, "undoInProgress", b)
try { set(true); body } finally { set(undo_in_progress) }
@@ -135,14 +130,12 @@
def jedit_text_areas(buffer: JEditBuffer): Iterator[JEditTextArea] =
jedit_text_areas().filter(_.getBuffer == buffer)
- def buffer_lock[A](buffer: JEditBuffer)(body: => A): A =
- {
+ def buffer_lock[A](buffer: JEditBuffer)(body: => A): A = {
try { buffer.readLock(); body }
finally { buffer.readUnlock() }
}
- def buffer_edit[A](buffer: JEditBuffer)(body: => A): A =
- {
+ def buffer_edit[A](buffer: JEditBuffer)(body: => A): A = {
try { buffer.beginCompoundEdit(); body }
finally { buffer.endCompoundEdit() }
}
@@ -186,8 +179,7 @@
def caret_range(text_area: TextArea): Text.Range =
point_range(text_area.getBuffer, text_area.getCaretPosition)
- def visible_range(text_area: TextArea): Option[Text.Range] =
- {
+ def visible_range(text_area: TextArea): Option[Text.Range] = {
val buffer = text_area.getBuffer
val n = text_area.getVisibleLines
if (n > 0) {
@@ -199,8 +191,7 @@
else None
}
- def invalidate_range(text_area: TextArea, range: Text.Range): Unit =
- {
+ def invalidate_range(text_area: TextArea, range: Text.Range): Unit = {
val buffer = text_area.getBuffer
buffer_range(buffer).try_restrict(range) match {
case Some(range1) if !range1.is_singularity =>
@@ -214,8 +205,7 @@
}
}
- def invalidate(text_area: TextArea): Unit =
- {
+ def invalidate(text_area: TextArea): Unit = {
val visible_lines = text_area.getVisibleLines
if (visible_lines > 0) text_area.invalidateScreenLineRange(0, visible_lines)
}
@@ -227,8 +217,7 @@
// NB: jEdit always normalizes \r\n and \r to \n
// NB: last line lacks \n
- def gfx_range(text_area: TextArea, range: Text.Range): Option[Gfx_Range] =
- {
+ def gfx_range(text_area: TextArea, range: Text.Range): Option[Gfx_Range] = {
val metric = pretty_metric(text_area.getPainter)
val char_width = (metric.unit * metric.average).round.toInt
@@ -258,8 +247,7 @@
/* pixel range */
- def pixel_range(text_area: TextArea, x: Int, y: Int): Option[Text.Range] =
- {
+ def pixel_range(text_area: TextArea, x: Int, y: Int): Option[Text.Range] = {
// coordinates wrt. inner painter component
val painter = text_area.getPainter
if (0 <= x && x < painter.getWidth && 0 <= y && y < painter.getHeight) {
@@ -279,8 +267,7 @@
/* pretty text metric */
- abstract class Pretty_Metric extends Pretty.Metric
- {
+ abstract class Pretty_Metric extends Pretty.Metric {
def average: Double
}
@@ -297,8 +284,7 @@
/* icons */
- def load_icon(name: String): Icon =
- {
+ def load_icon(name: String): Icon = {
val name1 =
if (name.startsWith("idea-icons/")) {
val file = Path.explode("$ISABELLE_IDEA_ICONS").file.toURI.toASCIIString
@@ -319,8 +305,7 @@
/* key event handling */
- def request_focus_view(alt_view: View = null): Unit =
- {
+ def request_focus_view(alt_view: View = null): Unit = {
val view = if (alt_view != null) alt_view else jEdit.getActiveView()
if (view != null) {
val text_area = view.getTextArea
@@ -328,8 +313,7 @@
}
}
- def propagate_key(view: View, evt: KeyEvent): Unit =
- {
+ def propagate_key(view: View, evt: KeyEvent): Unit = {
if (view != null && !evt.isConsumed)
view.getInputHandler().processKeyEvent(evt, View.ACTION_BAR, false)
}
@@ -337,24 +321,21 @@
def key_listener(
key_typed: KeyEvent => Unit = _ => (),
key_pressed: KeyEvent => Unit = _ => (),
- key_released: KeyEvent => Unit = _ => ()): KeyListener =
- {
- def process_key_event(evt0: KeyEvent, handle: KeyEvent => Unit): Unit =
- {
+ key_released: KeyEvent => Unit = _ => ()
+ ): KeyListener = {
+ def process_key_event(evt0: KeyEvent, handle: KeyEvent => Unit): Unit = {
val evt = KeyEventWorkaround.processKeyEvent(evt0)
if (evt != null) handle(evt)
}
- new KeyListener
- {
+ new KeyListener {
def keyTyped(evt: KeyEvent): Unit = process_key_event(evt, key_typed)
def keyPressed(evt: KeyEvent): Unit = process_key_event(evt, key_pressed)
def keyReleased(evt: KeyEvent): Unit = process_key_event(evt, key_released)
}
}
- def special_key(evt: KeyEvent): Boolean =
- {
+ def special_key(evt: KeyEvent): Boolean = {
// cf. 5.2.0/jEdit/org/gjt/sp/jedit/gui/KeyEventWorkaround.java
val mod = evt.getModifiersEx
(mod & InputEvent.CTRL_DOWN_MASK) != 0 && (mod & InputEvent.ALT_DOWN_MASK) == 0 ||
--- a/src/Tools/jEdit/src/jedit_main.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/jedit_main.scala Fri Apr 01 17:06:10 2022 +0200
@@ -12,18 +12,15 @@
import org.gjt.sp.jedit.{MiscUtilities, jEdit}
-object JEdit_Main
-{
+object JEdit_Main {
/* main entry point */
- def main(args: Array[String]): Unit =
- {
+ def main(args: Array[String]): Unit = {
if (args.nonEmpty && args(0) == "-init") {
Isabelle_System.init()
}
else {
- val start =
- {
+ val start = {
try {
Isabelle_System.init()
Isabelle_Fonts.init()
@@ -99,8 +96,7 @@
val jedit_options =
Isabelle_System.getenv_strict("JEDIT_OPTIONS").split(" +")
- val more_args =
- {
+ val more_args = {
args.toList.dropWhile(arg => arg.startsWith("-") && arg != "--") match {
case Nil | List("--") =>
args ++ Array(File.platform_path(Path.explode("$USER_HOME/Scratch.thy")))
--- a/src/Tools/jEdit/src/jedit_options.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/jedit_options.scala Fri Apr 01 17:06:10 2022 +0200
@@ -19,19 +19,16 @@
import org.gjt.sp.jedit.gui.ColorWellButton
-trait Option_Component extends Component
-{
+trait Option_Component extends Component {
val title: String
def load(): Unit
def save(): Unit
}
-object JEdit_Options
-{
+object JEdit_Options {
val RENDERING_SECTION = "Rendering of Document Content"
- class Check_Box(name: String, label: String, description: String) extends CheckBox(label)
- {
+ class Check_Box(name: String, label: String, description: String) extends CheckBox(label) {
tooltip = description
reactions += { case ButtonClicked(_) => update(selected) }
@@ -49,12 +46,10 @@
}
}
-class JEdit_Options(init_options: Options) extends Options_Variable(init_options)
-{
+class JEdit_Options(init_options: Options) extends Options_Variable(init_options) {
def color_value(s: String): Color = Color_Value(string(s))
- def make_color_component(opt: Options.Opt): Option_Component =
- {
+ def make_color_component(opt: Options.Opt): Option_Component = {
GUI_Thread.require {}
val opt_name = opt.name
@@ -72,8 +67,7 @@
component
}
- def make_component(opt: Options.Opt): Option_Component =
- {
+ def make_component(opt: Options.Opt): Option_Component = {
GUI_Thread.require {}
val opt_name = opt.name
@@ -120,9 +114,10 @@
component
}
- def make_components(predefined: List[Option_Component], filter: String => Boolean)
- : List[(String, List[Option_Component])] =
- {
+ def make_components(
+ predefined: List[Option_Component],
+ filter: String => Boolean
+ ) : List[(String, List[Option_Component])] = {
def mk_component(opt: Options.Opt): List[Option_Component] =
predefined.find(opt.name == _.name) match {
case Some(c) => List(c)
--- a/src/Tools/jEdit/src/jedit_rendering.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/jedit_rendering.scala Fri Apr 01 17:06:10 2022 +0200
@@ -19,16 +19,17 @@
import scala.collection.immutable.SortedMap
-object JEdit_Rendering
-{
+object JEdit_Rendering {
/* make rendering */
def apply(snapshot: Document.Snapshot, model: Document_Model, options: Options): JEdit_Rendering =
new JEdit_Rendering(snapshot, model, options)
- def text(snapshot: Document.Snapshot, formatted_body: XML.Body,
- results: Command.Results = Command.Results.empty): (String, JEdit_Rendering) =
- {
+ def text(
+ snapshot: Document.Snapshot,
+ formatted_body: XML.Body,
+ results: Command.Results = Command.Results.empty
+ ): (String, JEdit_Rendering) = {
val command = Command.rich_text(Document_ID.make(), results, formatted_body)
val snippet = snapshot.snippet(command)
val model = File_Model.empty(PIDE.session)
@@ -44,8 +45,7 @@
/* Isabelle/Isar token markup */
- private val command_style: Map[String, Byte] =
- {
+ private val command_style: Map[String, Byte] = {
import JEditToken._
Map[String, Byte](
Keyword.THY_END -> KEYWORD2,
@@ -54,8 +54,7 @@
).withDefaultValue(KEYWORD1)
}
- private val token_style: Map[Token.Kind.Value, Byte] =
- {
+ private val token_style: Map[Token.Kind.Value, Byte] = {
import JEditToken._
Map[Token.Kind.Value, Byte](
Token.Kind.KEYWORD -> KEYWORD2,
@@ -86,8 +85,7 @@
/* Isabelle/ML token markup */
- private val ml_token_style: Map[ML_Lex.Kind.Value, Byte] =
- {
+ private val ml_token_style: Map[ML_Lex.Kind.Value, Byte] = {
import JEditToken._
Map[ML_Lex.Kind.Value, Byte](
ML_Lex.Kind.KEYWORD -> NULL,
@@ -162,8 +160,7 @@
class JEdit_Rendering(snapshot: Document.Snapshot, model: Document_Model, options: Options)
- extends Rendering(snapshot, options, PIDE.session)
-{
+extends Rendering(snapshot, options, PIDE.session) {
override def get_text(range: Text.Range): Option[String] = model.get_text(range)
@@ -244,8 +241,7 @@
/* hyperlinks */
- def hyperlink(range: Text.Range): Option[Text.Info[PIDE.editor.Hyperlink]] =
- {
+ def hyperlink(range: Text.Range): Option[Text.Info[PIDE.editor.Hyperlink]] = {
snapshot.cumulate[Vector[Text.Info[PIDE.editor.Hyperlink]]](
range, Vector.empty, JEdit_Rendering.hyperlink_elements, _ =>
{
@@ -285,8 +281,7 @@
}) match { case Text.Info(_, _ :+ info) :: _ => Some(info) case _ => None }
}
- def hyperlink_entity(range: Text.Range): Option[Text.Info[PIDE.editor.Hyperlink]] =
- {
+ def hyperlink_entity(range: Text.Range): Option[Text.Info[PIDE.editor.Hyperlink]] = {
snapshot.cumulate[Vector[Text.Info[PIDE.editor.Hyperlink]]](
range, Vector.empty, Rendering.entity_elements, _ =>
{
@@ -321,8 +316,7 @@
def tooltip_margin: Int = options.int("jedit_tooltip_margin")
override def timing_threshold: Double = options.real("jedit_timing_threshold")
- def tooltip(range: Text.Range, control: Boolean): Option[Text.Info[XML.Body]] =
- {
+ def tooltip(range: Text.Range, control: Boolean): Option[Text.Info[XML.Body]] = {
val elements = if (control) Rendering.tooltip_elements else Rendering.tooltip_message_elements
tooltips(elements, range).map(info => info.map(Pretty.fbreaks))
}
@@ -354,8 +348,7 @@
(JEdit_Lib.load_icon(options.string("gutter_error_icon")),
color(Rendering.Color.error_message)))
- def gutter_content(range: Text.Range): Option[(Icon, Color)] =
- {
+ def gutter_content(range: Text.Range): Option[(Icon, Color)] = {
val pris =
snapshot.cumulate[Int](range, 0, JEdit_Rendering.gutter_elements, _ =>
{
@@ -372,8 +365,7 @@
def squiggly_underline(range: Text.Range): List[Text.Info[Rendering.Color.Value]] =
message_underline_color(JEdit_Rendering.squiggly_elements, range)
- def line_background(range: Text.Range): Option[(Rendering.Color.Value, Boolean)] =
- {
+ def line_background(range: Text.Range): Option[(Rendering.Color.Value, Boolean)] = {
val results =
snapshot.cumulate[Int](range, 0, JEdit_Rendering.line_background_elements, _ =>
{
@@ -381,22 +373,20 @@
})
val pri = results.foldLeft(0) { case (p1, Text.Info(_, p2)) => p1 max p2 }
- Rendering.message_background_color.get(pri).map(message_color =>
- {
- val is_separator =
- snapshot.cumulate[Boolean](range, false, JEdit_Rendering.separator_elements, _ =>
- {
- case _ => Some(true)
- }).exists(_.info)
- (message_color, is_separator)
- })
+ Rendering.message_background_color.get(pri).map(message_color => {
+ val is_separator =
+ snapshot.cumulate[Boolean](range, false, JEdit_Rendering.separator_elements, _ =>
+ {
+ case _ => Some(true)
+ }).exists(_.info)
+ (message_color, is_separator)
+ })
}
/* text color */
- def text_color(range: Text.Range, current_color: Color): List[Text.Info[Color]] =
- {
+ def text_color(range: Text.Range, current_color: Color): List[Text.Info[Color]] = {
if (current_color == Syntax_Style.hidden_color) List(Text.Info(range, current_color))
else
snapshot.cumulate(range, current_color, Rendering.text_color_elements, _ =>
--- a/src/Tools/jEdit/src/jedit_resources.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/jedit_resources.scala Fri Apr 01 17:06:10 2022 +0200
@@ -21,15 +21,13 @@
import org.gjt.sp.jedit.bufferio.BufferIORequest
-object JEdit_Resources
-{
+object JEdit_Resources {
def apply(options: Options): JEdit_Resources =
new JEdit_Resources(JEdit_Sessions.session_base_info(options))
}
class JEdit_Resources private(val session_base_info: Sessions.Base_Info)
- extends Resources(session_base_info.sessions_structure, session_base_info.base)
-{
+extends Resources(session_base_info.sessions_structure, session_base_info.base) {
def session_name: String = session_base_info.session
def session_errors: List[String] = session_base_info.errors
@@ -51,14 +49,12 @@
def node_name(buffer: Buffer): Document.Node.Name =
node_name(JEdit_Lib.buffer_name(buffer))
- def theory_node_name(buffer: Buffer): Option[Document.Node.Name] =
- {
+ def theory_node_name(buffer: Buffer): Option[Document.Node.Name] = {
val name = node_name(buffer)
if (name.is_theory) Some(name) else None
}
- override def append(dir: String, source_path: Path): String =
- {
+ override def append(dir: String, source_path: Path): String = {
val path = source_path.expand
if (dir == "" || path.is_absolute)
MiscUtilities.resolveSymlinks(File.platform_path(path))
@@ -75,8 +71,7 @@
/* file content */
- def read_file_content(node_name: Document.Node.Name): Option[String] =
- {
+ def read_file_content(node_name: Document.Node.Name): Option[String] = {
make_theory_content(node_name) orElse {
val name = node_name.node
try {
@@ -96,8 +91,7 @@
case None => read_file_content(node_name)
}
- override def with_thy_reader[A](name: Document.Node.Name, f: Reader[Char] => A): A =
- {
+ override def with_thy_reader[A](name: Document.Node.Name, f: Reader[Char] => A): A = {
GUI_Thread.now {
Document_Model.get(name) match {
case Some(model: Buffer_Model) =>
@@ -116,18 +110,15 @@
}
- private class File_Content_Output(buffer: Buffer) extends
- ByteArrayOutputStream(buffer.getLength + 1)
- {
+ private class File_Content_Output(buffer: Buffer)
+ extends ByteArrayOutputStream(buffer.getLength + 1) {
def content(): Bytes = Bytes(this.buf, 0, this.count)
}
- private class File_Content(buffer: Buffer) extends
- BufferIORequest(null, buffer, null, VFSManager.getVFSForPath(buffer.getPath), buffer.getPath)
- {
+ private class File_Content(buffer: Buffer)
+ extends BufferIORequest(null, buffer, null, VFSManager.getVFSForPath(buffer.getPath), buffer.getPath) {
def _run(): Unit = {}
- def content(): Bytes =
- {
+ def content(): Bytes = {
val out = new File_Content_Output(buffer)
write(buffer, out)
out.content()
@@ -139,8 +130,7 @@
/* theory text edits */
- override def commit(change: Session.Change): Unit =
- {
+ override def commit(change: Session.Change): Unit = {
if (change.syntax_changed.nonEmpty)
GUI_Thread.later { Document_Model.syntax_changed(change.syntax_changed) }
if (change.deps_changed ||
--- a/src/Tools/jEdit/src/jedit_sessions.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/jedit_sessions.scala Fri Apr 01 17:06:10 2022 +0200
@@ -14,8 +14,7 @@
import scala.swing.event.SelectionChanged
-object JEdit_Sessions
-{
+object JEdit_Sessions {
/* session options */
def session_dirs: List[Path] =
@@ -24,8 +23,7 @@
def session_no_build: Boolean =
Isabelle_System.getenv("JEDIT_NO_BUILD") == "true"
- def session_options(options: Options): Options =
- {
+ def session_options(options: Options): Options = {
val options1 =
Isabelle_System.getenv("JEDIT_BUILD_MODE") match {
case "default" => options
@@ -70,17 +68,14 @@
/* logic selector */
- private class Logic_Entry(val name: String, val description: String)
- {
+ private class Logic_Entry(val name: String, val description: String) {
override def toString: String = description
}
- def logic_selector(options: Options_Variable, autosave: Boolean): Option_Component =
- {
+ def logic_selector(options: Options_Variable, autosave: Boolean): Option_Component = {
GUI_Thread.require {}
- val session_list =
- {
+ val session_list = {
val sessions = sessions_structure(options.value)
val (main_sessions, other_sessions) =
sessions.imports_topological_order.partition(name => sessions(name).groups.contains("main"))
@@ -94,8 +89,7 @@
val component = new ComboBox(entries) with Option_Component {
name = jedit_logic_option
val title = "Logic"
- def load(): Unit =
- {
+ def load(): Unit = {
val logic = options.string(jedit_logic_option)
entries.find(_.name == logic) match {
case Some(entry) => selection.item = entry
@@ -126,16 +120,17 @@
session_requirements = logic_requirements)
def session_build(
- options: Options, progress: Progress = new Progress, no_build: Boolean = false): Int =
- {
+ options: Options,
+ progress: Progress = new Progress,
+ no_build: Boolean = false
+ ): Int = {
Build.build(session_options(options),
selection = Sessions.Selection.session(PIDE.resources.session_name),
progress = progress, build_heap = true, no_build = no_build, dirs = session_dirs,
infos = PIDE.resources.session_base_info.infos).rc
}
- def session_start(options0: Options): Unit =
- {
+ def session_start(options0: Options): Unit = {
val session = PIDE.session
val options = session_options(options0)
val sessions_structure = PIDE.resources.session_base_info.sessions_structure
--- a/src/Tools/jEdit/src/jedit_spell_checker.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/jedit_spell_checker.scala Fri Apr 01 17:06:10 2022 +0200
@@ -17,13 +17,14 @@
import org.gjt.sp.jedit.textarea.{JEditTextArea, TextArea}
-object JEdit_Spell_Checker
-{
+object JEdit_Spell_Checker {
/* completion */
def completion(
- rendering: JEdit_Rendering, explicit: Boolean, caret: Text.Offset): Option[Completion.Result] =
- {
+ rendering: JEdit_Rendering,
+ explicit: Boolean,
+ caret: Text.Offset
+ ): Option[Completion.Result] = {
PIDE.plugin.spell_checker.get match {
case Some(spell_checker) if explicit =>
spell_checker.completion(rendering, caret)
@@ -34,8 +35,7 @@
/* context menu */
- def context_menu(text_area: JEditTextArea, offset: Text.Offset): List[JMenuItem] =
- {
+ def context_menu(text_area: JEditTextArea, offset: Text.Offset): List[JMenuItem] = {
val result =
for {
spell_checker <- PIDE.plugin.spell_checker.get
@@ -80,8 +80,7 @@
/* dictionaries */
- def dictionaries_selector(): Option_Component =
- {
+ def dictionaries_selector(): Option_Component = {
GUI_Thread.require {}
val option_name = "spell_checker_dictionary"
@@ -91,8 +90,7 @@
val component = new ComboBox(entries) with Option_Component {
name = option_name
val title = opt.title()
- def load(): Unit =
- {
+ def load(): Unit = {
val lang = PIDE.options.string(option_name)
entries.find(_.lang == lang) match {
case Some(entry) => selection.item = entry
--- a/src/Tools/jEdit/src/keymap_merge.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/keymap_merge.scala Fri Apr 01 17:06:10 2022 +0200
@@ -21,16 +21,14 @@
import org.jedit.keymap.{KeymapManager, Keymap}
-object Keymap_Merge
-{
+object Keymap_Merge {
/** shortcuts **/
private def is_shortcut(property: String): Boolean =
(property.endsWith(".shortcut") || property.endsWith(".shortcut2")) &&
!property.startsWith("options.shortcuts.")
- class Shortcut(val property: String, val binding: String)
- {
+ class Shortcut(val property: String, val binding: String) {
override def toString: String = Properties.Eq(property, binding)
def primary: Boolean = property.endsWith(".shortcut")
@@ -79,8 +77,10 @@
result.sortBy(_.property)
}
- def get_shortcut_conflicts(keymap_name: String, keymap: Keymap): List[(Shortcut, List[Shortcut])] =
- {
+ def get_shortcut_conflicts(
+ keymap_name: String,
+ keymap: Keymap
+ ): List[(Shortcut, List[Shortcut])] = {
val keymap_shortcuts =
if (keymap == null) Nil
else convert_properties(Untyped.get[JProperties](keymap, "props"))
@@ -101,8 +101,7 @@
private def conflict_color: Color =
PIDE.options.color_value("error_color")
- private sealed case class Table_Entry(shortcut: Shortcut, head: Option[Int], tail: List[Int])
- {
+ private sealed case class Table_Entry(shortcut: Shortcut, head: Option[Int], tail: List[Int]) {
override def toString: String =
if (head.isEmpty) "<html>" + HTML.output(shortcut.toString) + "</html>"
else
@@ -111,8 +110,7 @@
"</font></html>"
}
- private class Table_Model(entries: List[Table_Entry]) extends AbstractTableModel
- {
+ private class Table_Model(entries: List[Table_Entry]) extends AbstractTableModel {
private val entries_count = entries.length
private def has_entry(row: Int): Boolean = 0 <= row && row <= entries_count
private def get_entry(row: Int): Option[Table_Entry] =
@@ -128,8 +126,7 @@
private def select(head: Int, tail: List[Int], b: Boolean): Unit =
selected.change(set => if (b) set + head -- tail else set - head ++ tail)
- def apply(keymap_name: String, keymap: Keymap): Unit =
- {
+ def apply(keymap_name: String, keymap: Keymap): Unit = {
GUI_Thread.require {}
for ((entry, row) <- entries.iterator.zipWithIndex if entry.head.isEmpty) {
@@ -153,8 +150,7 @@
override def getRowCount: Int = entries_count
- override def getValueAt(row: Int, column: Int): AnyRef =
- {
+ override def getValueAt(row: Int, column: Int): AnyRef = {
get_entry(row) match {
case Some(entry) if column == 0 => java.lang.Boolean.valueOf(is_selected(row))
case Some(entry) if column == 1 => entry
@@ -165,8 +161,7 @@
override def isCellEditable(row: Int, column: Int): Boolean =
has_entry(row) && column == 0
- override def setValueAt(value: AnyRef, row: Int, column: Int): Unit =
- {
+ override def setValueAt(value: AnyRef, row: Int, column: Int): Unit = {
value match {
case obj: java.lang.Boolean if has_entry(row) && column == 0 =>
val b = obj.booleanValue
@@ -183,8 +178,7 @@
}
}
- private class Table(table_model: Table_Model) extends JPanel(new BorderLayout)
- {
+ private class Table(table_model: Table_Model) extends JPanel(new BorderLayout) {
private val cell_size = GenericGUIUtilities.defaultTableCellSize()
private val table_size = new Dimension(cell_size.width * 2, cell_size.height * 15)
@@ -213,8 +207,7 @@
/** check with optional dialog **/
- def check_dialog(view: View): Unit =
- {
+ def check_dialog(view: View): Unit = {
GUI_Thread.require {}
val keymap_manager = jEdit.getKeymapManager
@@ -248,8 +241,9 @@
"jEdit keymap " + quote(keymap_name) + ":",
new Table(table_model),
"Selected shortcuts will be applied, unselected changes will be ignored.",
- "Results are stored in $JEDIT_SETTINGS/properties and $JEDIT_SETTINGS/keymaps/.") == 0)
- { table_model.apply(keymap_name, keymap) }
+ "Results are stored in $JEDIT_SETTINGS/properties and $JEDIT_SETTINGS/keymaps/.") == 0) {
+ table_model.apply(keymap_name, keymap)
+ }
no_shortcut_conflicts.foreach(_.set(keymap))
--- a/src/Tools/jEdit/src/main_plugin.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/main_plugin.scala Fri Apr 01 17:06:10 2022 +0200
@@ -22,18 +22,15 @@
import org.gjt.sp.util.Log
-object PIDE
-{
+object PIDE {
/* semantic document content */
- def maybe_snapshot(view: View = null): Option[Document.Snapshot] = GUI_Thread.now
- {
+ def maybe_snapshot(view: View = null): Option[Document.Snapshot] = GUI_Thread.now {
val buffer = JEdit_Lib.jedit_view(view).getBuffer
Document_Model.get(buffer).map(_.snapshot())
}
- def maybe_rendering(view: View = null): Option[JEdit_Rendering] = GUI_Thread.now
- {
+ def maybe_rendering(view: View = null): Option[JEdit_Rendering] = GUI_Thread.now {
val text_area = JEdit_Lib.jedit_view(view).getTextArea
Document_View.get(text_area).map(_.get_rendering())
}
@@ -60,8 +57,7 @@
object editor extends JEdit_Editor
}
-class Main_Plugin extends EBPlugin
-{
+class Main_Plugin extends EBPlugin {
/* options */
private var _options: JEdit_Options = null
@@ -92,14 +88,12 @@
/* global changes */
- def options_changed(): Unit =
- {
+ def options_changed(): Unit = {
session.global_options.post(Session.Global_Options(options.value))
delay_load.invoke()
}
- def deps_changed(): Unit =
- {
+ def deps_changed(): Unit = {
delay_load.invoke()
}
@@ -107,23 +101,19 @@
/* theory files */
lazy val delay_init: Delay =
- Delay.last(options.seconds("editor_load_delay"), gui = true)
- {
+ Delay.last(options.seconds("editor_load_delay"), gui = true) {
init_models()
}
private val delay_load_active = Synchronized(false)
private def delay_load_activated(): Boolean =
delay_load_active.guarded_access(a => Some((!a, true)))
- private def delay_load_action(): Unit =
- {
+ private def delay_load_action(): Unit = {
if (Isabelle.continuous_checking && delay_load_activated() &&
- PerspectiveManager.isPerspectiveEnabled)
- {
+ PerspectiveManager.isPerspectiveEnabled) {
if (JEdit_Lib.jedit_buffers().exists(_.isLoading)) delay_load.invoke()
else {
- val required_files =
- {
+ val required_files = {
val models = Document_Model.get_models()
val thys =
@@ -189,8 +179,7 @@
/* session phase */
- val session_phase_changed: Session.Consumer[Session.Phase] = Session.Consumer("Isabelle/jEdit")
- {
+ val session_phase_changed: Session.Consumer[Session.Phase] = Session.Consumer("Isabelle/jEdit") {
case Session.Terminated(result) if !result.ok =>
GUI_Thread.later {
GUI.error_dialog(jEdit.getActiveView, "Prover process terminated with error",
@@ -229,8 +218,7 @@
/* document model and view */
- def exit_models(buffers: List[Buffer]): Unit =
- {
+ def exit_models(buffers: List[Buffer]): Unit = {
GUI_Thread.now {
buffers.foreach(buffer =>
JEdit_Lib.buffer_lock(buffer) {
@@ -240,8 +228,7 @@
}
}
- def init_models(): Unit =
- {
+ def init_models(): Unit = {
GUI_Thread.now {
PIDE.editor.flush()
@@ -289,14 +276,12 @@
@volatile private var startup_failure: Option[Throwable] = None
@volatile private var startup_notified = false
- private def init_editor(view: View): Unit =
- {
+ private def init_editor(view: View): Unit = {
Keymap_Merge.check_dialog(view)
Session_Build.check_dialog(view)
}
- private def init_title(view: View): Unit =
- {
+ private def init_title(view: View): Unit = {
val title =
proper_string(Isabelle_System.getenv("ISABELLE_IDENTIFIER")).getOrElse("Isabelle") +
"/" + PIDE.resources.session_name
@@ -308,8 +293,7 @@
}
}
- override def handleMessage(message: EBMessage): Unit =
- {
+ override def handleMessage(message: EBMessage): Unit = {
GUI_Thread.assert {}
if (startup_failure.isDefined && !startup_notified) {
@@ -409,8 +393,7 @@
private var orig_mode_provider: ModeProvider = null
private var pide_mode_provider: ModeProvider = null
- def init_mode_provider(): Unit =
- {
+ def init_mode_provider(): Unit = {
orig_mode_provider = ModeProvider.instance
if (orig_mode_provider.isInstanceOf[ModeProvider]) {
pide_mode_provider = new Token_Markup.Mode_Provider(orig_mode_provider)
@@ -418,8 +401,7 @@
}
}
- def exit_mode_provider(): Unit =
- {
+ def exit_mode_provider(): Unit = {
if (ModeProvider.instance == pide_mode_provider)
ModeProvider.instance = orig_mode_provider
}
@@ -437,8 +419,7 @@
private val shutting_down = Synchronized(false)
- override def start(): Unit =
- {
+ override def start(): Unit = {
/* strict initialization */
init_options()
@@ -476,8 +457,7 @@
if (view != null) init_editor(view)
}
- override def stop(): Unit =
- {
+ override def stop(): Unit = {
http_server.stop()
Syntax_Style.set_extender(Syntax_Style.Base_Extender)
--- a/src/Tools/jEdit/src/monitor_dockable.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/monitor_dockable.scala Fri Apr 01 17:06:10 2022 +0200
@@ -21,15 +21,13 @@
import org.gjt.sp.jedit.View
-class Monitor_Dockable(view: View, position: String) extends Dockable(view, position)
-{
+class Monitor_Dockable(view: View, position: String) extends Dockable(view, position) {
/* chart data -- owned by GUI thread */
private var statistics = Queue.empty[Properties.T]
private var statistics_length = 0
- private def add_statistics(stats: Properties.T): Unit =
- {
+ private def add_statistics(stats: Properties.T): Unit = {
statistics = statistics.enqueue(stats)
statistics_length += 1
limit_data.text match {
@@ -41,8 +39,7 @@
case _ =>
}
}
- private def clear_statistics(): Unit =
- {
+ private def clear_statistics(): Unit = {
statistics = Queue.empty
statistics_length = 0
}
@@ -51,8 +48,7 @@
private val chart = ML_Statistics.empty.chart(null, Nil)
private val data = chart.getXYPlot.getDataset.asInstanceOf[XYSeriesCollection]
- private def update_chart(): Unit =
- {
+ private def update_chart(): Unit = {
ML_Statistics.all_fields.find(_._1 == data_name) match {
case None =>
case Some((_, fields)) => ML_Statistics(statistics.toList).update_data(data, fields)
@@ -68,8 +64,7 @@
/* controls */
- private val select_data = new ComboBox[String](ML_Statistics.all_fields.map(_._1))
- {
+ private val select_data = new ComboBox[String](ML_Statistics.all_fields.map(_._1)) {
tooltip = "Select visualized data collection"
listenTo(selection)
reactions += {
@@ -132,13 +127,11 @@
update_delay.invoke()
}
- override def init(): Unit =
- {
+ override def init(): Unit = {
PIDE.session.runtime_statistics += main
}
- override def exit(): Unit =
- {
+ override def exit(): Unit = {
PIDE.session.runtime_statistics -= main
}
}
--- a/src/Tools/jEdit/src/output_dockable.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/output_dockable.scala Fri Apr 01 17:06:10 2022 +0200
@@ -18,8 +18,7 @@
import org.gjt.sp.jedit.View
-class Output_Dockable(view: View, position: String) extends Dockable(view, position)
-{
+class Output_Dockable(view: View, position: String) extends Dockable(view, position) {
/* component state -- owned by GUI thread */
private var do_update = true
@@ -34,16 +33,14 @@
override def detach_operation: Option[() => Unit] = pretty_text_area.detach_operation
- private def handle_resize(): Unit =
- {
+ private def handle_resize(): Unit = {
GUI_Thread.require {}
pretty_text_area.resize(
Font_Info.main(PIDE.options.real("jedit_font_scale") * zoom.factor / 100))
}
- private def handle_update(follow: Boolean, restriction: Option[Set[Command]]): Unit =
- {
+ private def handle_update(follow: Boolean, restriction: Option[Set[Command]]): Unit = {
GUI_Thread.require {}
for {
@@ -72,8 +69,7 @@
/* controls */
private def output_state: Boolean = PIDE.options.bool("editor_output_state")
- private def output_state_=(b: Boolean): Unit =
- {
+ private def output_state_=(b: Boolean): Unit = {
if (output_state != b) {
PIDE.options.bool("editor_output_state") = b
PIDE.session.update_options(PIDE.options.value)
@@ -82,8 +78,7 @@
}
}
- private val output_state_button = new CheckBox("Proof state")
- {
+ private val output_state_button = new CheckBox("Proof state") {
tooltip = "Output of proof state (normally shown on State panel)"
reactions += { case ButtonClicked(_) => output_state = selected }
selected = output_state
@@ -130,16 +125,14 @@
GUI_Thread.later { handle_update(do_update, None) }
}
- override def init(): Unit =
- {
+ override def init(): Unit = {
PIDE.session.global_options += main
PIDE.session.commands_changed += main
PIDE.session.caret_focus += main
handle_update(true, None)
}
- override def exit(): Unit =
- {
+ override def exit(): Unit = {
PIDE.session.global_options -= main
PIDE.session.commands_changed -= main
PIDE.session.caret_focus -= main
--- a/src/Tools/jEdit/src/pide_docking_framework.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/pide_docking_framework.scala Fri Apr 01 17:06:10 2022 +0200
@@ -19,19 +19,17 @@
FloatingWindowContainer, PanelWindowContainer}
-class PIDE_Docking_Framework extends DockableWindowManagerProvider
-{
+class PIDE_Docking_Framework extends DockableWindowManagerProvider {
override def create(
view: View,
instance: DockableWindowFactory,
config: View.ViewConfig): DockableWindowManager =
- new DockableWindowManagerImpl(view, instance, config)
- {
+ new DockableWindowManagerImpl(view, instance, config) {
override def createPopupMenu(
container: DockableWindowContainer,
dockable_name: String,
- clone: Boolean): JPopupMenu =
- {
+ clone: Boolean
+ ): JPopupMenu = {
val menu = super.createPopupMenu(container, dockable_name, clone)
val detach_operation: Option[() => Unit] =
--- a/src/Tools/jEdit/src/pretty_text_area.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/pretty_text_area.scala Fri Apr 01 17:06:10 2022 +0200
@@ -30,8 +30,8 @@
class Pretty_Text_Area(
view: View,
close_action: () => Unit = () => (),
- propagate_keys: Boolean = false) extends JEditEmbeddedTextArea
-{
+ propagate_keys: Boolean = false
+) extends JEditEmbeddedTextArea {
text_area =>
GUI_Thread.require {}
@@ -53,8 +53,7 @@
def get_background(): Option[Color] = None
- def refresh(): Unit =
- {
+ def refresh(): Unit = {
GUI_Thread.require {}
val font = current_font_info.font
@@ -116,8 +115,7 @@
}
}
- def resize(font_info: Font_Info): Unit =
- {
+ def resize(font_info: Font_Info): Unit = {
GUI_Thread.require {}
current_font_info = font_info
@@ -125,8 +123,10 @@
}
def update(
- base_snapshot: Document.Snapshot, base_results: Command.Results, body: XML.Body): Unit =
- {
+ base_snapshot: Document.Snapshot,
+ base_results: Command.Results,
+ body: XML.Body
+ ): Unit = {
GUI_Thread.require {}
require(!base_snapshot.is_outdated, "document snapshot outdated")
@@ -136,8 +136,7 @@
refresh()
}
- def detach: Unit =
- {
+ def detach: Unit = {
GUI_Thread.require {}
Info_Dockable(view, current_base_snapshot, current_base_results, current_body)
}
@@ -153,26 +152,24 @@
}
val search_field: Component =
- Component.wrap(new Completion_Popup.History_Text_Field("isabelle-search")
- {
- private val input_delay =
- Delay.last(PIDE.options.seconds("editor_input_delay"), gui = true) {
- search_action(this)
- }
- getDocument.addDocumentListener(new DocumentListener {
- def changedUpdate(e: DocumentEvent): Unit = input_delay.invoke()
- def insertUpdate(e: DocumentEvent): Unit = input_delay.invoke()
- def removeUpdate(e: DocumentEvent): Unit = input_delay.invoke()
- })
- setColumns(20)
- setToolTipText(search_label.tooltip)
- setFont(GUI.imitate_font(getFont, scale = 1.2))
+ Component.wrap(new Completion_Popup.History_Text_Field("isabelle-search") {
+ private val input_delay =
+ Delay.last(PIDE.options.seconds("editor_input_delay"), gui = true) {
+ search_action(this)
+ }
+ getDocument.addDocumentListener(new DocumentListener {
+ def changedUpdate(e: DocumentEvent): Unit = input_delay.invoke()
+ def insertUpdate(e: DocumentEvent): Unit = input_delay.invoke()
+ def removeUpdate(e: DocumentEvent): Unit = input_delay.invoke()
})
+ setColumns(20)
+ setToolTipText(search_label.tooltip)
+ setFont(GUI.imitate_font(getFont, scale = 1.2))
+ })
private val search_field_foreground = search_field.foreground
- private def search_action(text_field: JTextField): Unit =
- {
+ private def search_action(text_field: JTextField): Unit = {
val (pattern, ok) =
text_field.getText match {
case null | "" => (None, true)
@@ -203,34 +200,31 @@
})
addKeyListener(JEdit_Lib.key_listener(
- key_pressed = (evt: KeyEvent) =>
- {
- val strict_control =
- JEdit_Lib.command_modifier(evt) && !JEdit_Lib.shift_modifier(evt)
+ key_pressed = (evt: KeyEvent) => {
+ val strict_control =
+ JEdit_Lib.command_modifier(evt) && !JEdit_Lib.shift_modifier(evt)
- evt.getKeyCode match {
- case KeyEvent.VK_C | KeyEvent.VK_INSERT
- if strict_control && text_area.getSelectionCount != 0 =>
- Registers.copy(text_area, '$')
- evt.consume
+ evt.getKeyCode match {
+ case KeyEvent.VK_C | KeyEvent.VK_INSERT
+ if strict_control && text_area.getSelectionCount != 0 =>
+ Registers.copy(text_area, '$')
+ evt.consume
- case KeyEvent.VK_A
- if strict_control =>
- text_area.selectAll
- evt.consume
+ case KeyEvent.VK_A
+ if strict_control =>
+ text_area.selectAll
+ evt.consume
- case KeyEvent.VK_ESCAPE =>
- if (Isabelle.dismissed_popups(view)) evt.consume
+ case KeyEvent.VK_ESCAPE =>
+ if (Isabelle.dismissed_popups(view)) evt.consume
- case _ =>
- }
- if (propagate_keys) JEdit_Lib.propagate_key(view, evt)
- },
- key_typed = (evt: KeyEvent) =>
- {
- if (propagate_keys) JEdit_Lib.propagate_key(view, evt)
+ case _ =>
}
- )
+ if (propagate_keys) JEdit_Lib.propagate_key(view, evt)
+ },
+ key_typed = (evt: KeyEvent) => {
+ if (propagate_keys) JEdit_Lib.propagate_key(view, evt)
+ })
)
--- a/src/Tools/jEdit/src/pretty_tooltip.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/pretty_tooltip.scala Fri Apr 01 17:06:10 2022 +0200
@@ -21,15 +21,15 @@
import org.gjt.sp.jedit.textarea.TextArea
-object Pretty_Tooltip
-{
+object Pretty_Tooltip {
/* tooltip hierarchy */
// owned by GUI thread
private var stack: List[Pretty_Tooltip] = Nil
- private def hierarchy(tip: Pretty_Tooltip): Option[(List[Pretty_Tooltip], List[Pretty_Tooltip])] =
- {
+ private def hierarchy(
+ tip: Pretty_Tooltip
+ ): Option[(List[Pretty_Tooltip], List[Pretty_Tooltip])] = {
GUI_Thread.require {}
if (stack.contains(tip)) Some(stack.span(_ != tip))
@@ -45,8 +45,8 @@
location: Point,
rendering: JEdit_Rendering,
results: Command.Results,
- info: Text.Info[XML.Body]): Unit =
- {
+ info: Text.Info[XML.Body]
+ ): Unit = {
GUI_Thread.require {}
stack match {
@@ -113,13 +113,12 @@
/* dismiss */
- private lazy val focus_delay = Delay.last(PIDE.options.seconds("editor_input_delay"), gui = true)
- {
- dismiss_unfocused()
- }
+ private lazy val focus_delay =
+ Delay.last(PIDE.options.seconds("editor_input_delay"), gui = true) {
+ dismiss_unfocused()
+ }
- def dismiss_unfocused(): Unit =
- {
+ def dismiss_unfocused(): Unit = {
stack.span(tip => !tip.pretty_text_area.isFocusOwner) match {
case (Nil, _) =>
case (unfocused, rest) =>
@@ -129,8 +128,7 @@
}
}
- def dismiss(tip: Pretty_Tooltip): Unit =
- {
+ def dismiss(tip: Pretty_Tooltip): Unit = {
deactivate()
hierarchy(tip) match {
case Some((old, _ :: rest)) =>
@@ -148,8 +146,7 @@
def dismiss_descendant(parent: JComponent): Unit =
descendant(parent).foreach(dismiss)
- def dismissed_all(): Boolean =
- {
+ def dismissed_all(): Boolean = {
deactivate()
if (stack.isEmpty) false
else {
@@ -169,8 +166,8 @@
location: Point,
rendering: JEdit_Rendering,
private val results: Command.Results,
- private val info: Text.Info[XML.Body]) extends JPanel(new BorderLayout)
-{
+ private val info: Text.Info[XML.Body]
+) extends JPanel(new BorderLayout) {
tip =>
GUI_Thread.require {}
@@ -209,13 +206,11 @@
}
pretty_text_area.addFocusListener(new FocusAdapter {
- override def focusGained(e: FocusEvent): Unit =
- {
+ override def focusGained(e: FocusEvent): Unit = {
tip_border(true)
Pretty_Tooltip.focus_delay.invoke()
}
- override def focusLost(e: FocusEvent): Unit =
- {
+ override def focusLost(e: FocusEvent): Unit = {
tip_border(false)
Pretty_Tooltip.focus_delay.invoke()
}
@@ -226,8 +221,7 @@
/* main content */
- def tip_border(has_focus: Boolean): Unit =
- {
+ def tip_border(has_focus: Boolean): Unit = {
tip.setBorder(new LineBorder(if (has_focus) Color.BLACK else Color.GRAY))
tip.repaint()
}
@@ -241,11 +235,9 @@
/* popup */
- private val popup =
- {
+ private val popup = {
val screen = GUI.screen_location(layered, location)
- val size =
- {
+ val size = {
val bounds = JEdit_Rendering.popup_bounds
val w_max = layered.getWidth min (screen.bounds.width * bounds).toInt
@@ -277,8 +269,7 @@
new Popup(layered, tip, screen.relative(layered, size), size)
}
- private def show_popup: Unit =
- {
+ private def show_popup: Unit = {
popup.show
pretty_text_area.requestFocus()
pretty_text_area.update(rendering.snapshot, results, info.info)
--- a/src/Tools/jEdit/src/process_indicator.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/process_indicator.scala Fri Apr 01 17:06:10 2022 +0200
@@ -15,8 +15,7 @@
import scala.swing.{Label, Component}
-class Process_Indicator
-{
+class Process_Indicator {
private val label = new Label
private val passive_icon =
@@ -25,13 +24,11 @@
space_explode(':', PIDE.options.string("process_active_icons")).map(name =>
JEdit_Lib.load_image_icon(name).getImage)
- private class Animation extends ImageIcon(passive_icon)
- {
+ private class Animation extends ImageIcon(passive_icon) {
private var current_frame = 0
private val timer =
new Timer(0, new ActionListener {
- override def actionPerformed(e: ActionEvent): Unit =
- {
+ override def actionPerformed(e: ActionEvent): Unit = {
current_frame = (current_frame + 1) % active_icons.length
setImage(active_icons(current_frame))
label.repaint()
@@ -39,8 +36,7 @@
})
timer.setRepeats(true)
- def update(rate: Int): Unit =
- {
+ def update(rate: Int): Unit = {
if (rate == 0) {
setImage(passive_icon)
timer.stop
@@ -60,8 +56,7 @@
def component: Component = label
- def update(tip: String, rate: Int): Unit =
- {
+ def update(tip: String, rate: Int): Unit = {
label.tooltip = tip
animation.update(rate)
}
--- a/src/Tools/jEdit/src/protocol_dockable.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/protocol_dockable.scala Fri Apr 01 17:06:10 2022 +0200
@@ -16,8 +16,7 @@
import org.gjt.sp.jedit.View
-class Protocol_Dockable(view: View, position: String) extends Dockable(view, position)
-{
+class Protocol_Dockable(view: View, position: String) extends Dockable(view, position) {
/* text area */
private val text_area = new TextArea
@@ -39,13 +38,11 @@
GUI_Thread.later { text_area.append(output.message.toString + "\n\n") }
}
- override def init(): Unit =
- {
+ override def init(): Unit = {
PIDE.session.all_messages += main
}
- override def exit(): Unit =
- {
+ override def exit(): Unit = {
PIDE.session.all_messages -= main
}
}
--- a/src/Tools/jEdit/src/query_dockable.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/query_dockable.scala Fri Apr 01 17:06:10 2022 +0200
@@ -19,10 +19,8 @@
import org.gjt.sp.jedit.View
-object Query_Dockable
-{
- private abstract class Operation(view: View)
- {
+object Query_Dockable {
+ private abstract class Operation(view: View) {
val pretty_text_area = new Pretty_Text_Area(view)
def query_operation: Query_Operation[View]
def query: JComponent
@@ -31,18 +29,18 @@
}
}
-class Query_Dockable(view: View, position: String) extends Dockable(view, position)
-{
+class Query_Dockable(view: View, position: String) extends Dockable(view, position) {
/* common GUI components */
private val zoom = new Font_Info.Zoom_Box { def changed = handle_resize() }
- private def make_query(property: String, tooltip: String, apply_query: () => Unit)
- : Completion_Popup.History_Text_Field =
- new Completion_Popup.History_Text_Field(property)
- {
- override def processKeyEvent(evt: KeyEvent): Unit =
- {
+ private def make_query(
+ property: String,
+ tooltip: String,
+ apply_query: () => Unit
+ ): Completion_Popup.History_Text_Field = {
+ new Completion_Popup.History_Text_Field(property) {
+ override def processKeyEvent(evt: KeyEvent): Unit = {
if (evt.getID == KeyEvent.KEY_PRESSED && evt.getKeyCode == KeyEvent.VK_ENTER) apply_query()
super.processKeyEvent(evt)
}
@@ -51,13 +49,15 @@
setToolTipText(tooltip)
setFont(GUI.imitate_font(getFont, scale = 1.2))
}
+ }
/* consume status */
def consume_status(
- process_indicator: Process_Indicator, status: Query_Operation.Status.Value): Unit =
- {
+ process_indicator: Process_Indicator,
+ status: Query_Operation.Status.Value
+ ): Unit = {
status match {
case Query_Operation.Status.WAITING =>
process_indicator.update("Waiting for evaluation of context ...", 5)
@@ -71,8 +71,7 @@
/* find theorems */
- private val find_theorems = new Query_Dockable.Operation(view)
- {
+ private val find_theorems = new Query_Dockable.Operation(view) {
/* query */
private val process_indicator = new Process_Indicator
@@ -83,8 +82,7 @@
(snapshot, results, body) =>
pretty_text_area.update(snapshot, results, Pretty.separate(body)))
- private def apply_query(): Unit =
- {
+ private def apply_query(): Unit = {
query.addCurrentToHistory()
query_operation.apply_query(List(limit.text, allow_dups.selected.toString, query.getText))
}
@@ -137,8 +135,7 @@
/* find consts */
- private val find_consts = new Query_Dockable.Operation(view)
- {
+ private val find_consts = new Query_Dockable.Operation(view) {
/* query */
private val process_indicator = new Process_Indicator
@@ -149,8 +146,7 @@
(snapshot, results, body) =>
pretty_text_area.update(snapshot, results, Pretty.separate(body)))
- private def apply_query(): Unit =
- {
+ private def apply_query(): Unit = {
query.addCurrentToHistory()
query_operation.apply_query(List(query.getText))
}
@@ -187,12 +183,10 @@
/* print operation */
- private val print_operation = new Query_Dockable.Operation(view)
- {
+ private val print_operation = new Query_Dockable.Operation(view) {
/* items */
- private class Item(val name: String, description: String, sel: Boolean)
- {
+ private class Item(val name: String, description: String, sel: Boolean) {
val checkbox = new CheckBox(name) {
tooltip = "Print " + description
selected = sel
@@ -205,8 +199,7 @@
private def selected_items(): List[String] =
for (item <- _items if item.checkbox.selected) yield item.name
- private def update_items(): List[Item] =
- {
+ private def update_items(): List[Item] = {
val old_items = _items
def was_selected(name: String): Boolean =
old_items.find(item => item.name == name) match {
@@ -255,8 +248,7 @@
private val control_panel = Wrap_Panel()
- def select: Unit =
- {
+ def select: Unit = {
control_panel.contents.clear()
control_panel.contents += query_label
update_items().foreach(item => control_panel.contents += item.checkbox)
@@ -277,8 +269,7 @@
private val operations = List(find_theorems, find_consts, print_operation)
- private val operations_pane = new TabbedPane
- {
+ private val operations_pane = new TabbedPane {
pages ++= operations.map(_.page)
listenTo(selection)
reactions += { case SelectionChanged(_) => select_operation() }
@@ -288,14 +279,12 @@
try { Some(operations(operations_pane.selection.index)) }
catch { case _: IndexOutOfBoundsException => None }
- private def select_operation(): Unit =
- {
+ private def select_operation(): Unit = {
for (op <- get_operation()) { op.select; op.query.requestFocus() }
operations_pane.revalidate()
}
- override def focusOnDefaultComponent(): Unit =
- {
+ override def focusOnDefaultComponent(): Unit = {
for (op <- get_operation()) op.query.requestFocus()
}
@@ -335,15 +324,13 @@
case _: Session.Global_Options => GUI_Thread.later { handle_resize() }
}
- override def init(): Unit =
- {
+ override def init(): Unit = {
PIDE.session.global_options += main
handle_resize()
operations.foreach(op => op.query_operation.activate())
}
- override def exit(): Unit =
- {
+ override def exit(): Unit = {
operations.foreach(op => op.query_operation.deactivate())
PIDE.session.global_options -= main
delay_resize.revoke()
--- a/src/Tools/jEdit/src/raw_output_dockable.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/raw_output_dockable.scala Fri Apr 01 17:06:10 2022 +0200
@@ -14,8 +14,7 @@
import org.gjt.sp.jedit.View
-class Raw_Output_Dockable(view: View, position: String) extends Dockable(view, position)
-{
+class Raw_Output_Dockable(view: View, position: String) extends Dockable(view, position) {
private val text_area = new TextArea
set_content(new ScrollPane(text_area))
--- a/src/Tools/jEdit/src/rich_text_area.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/rich_text_area.scala Fri Apr 01 17:06:10 2022 +0200
@@ -33,8 +33,8 @@
get_search_pattern: () => Option[Regex],
caret_update: () => Unit,
caret_visible: Boolean,
- enable_hovering: Boolean)
-{
+ enable_hovering: Boolean
+) {
private val buffer = text_area.getBuffer
@@ -43,8 +43,7 @@
def check_robust_body: Boolean =
GUI_Thread.require { buffer == text_area.getBuffer }
- def robust_body[A](default: A)(body: => A): A =
- {
+ def robust_body[A](default: A)(body: => A): A = {
try {
if (check_robust_body) body
else {
@@ -58,8 +57,7 @@
/* original painters */
- private def pick_extension(name: String): TextAreaExtension =
- {
+ private def pick_extension(name: String): TextAreaExtension = {
text_area.getPainter.getExtensions.iterator.filter(x => x.getClass.getName == name).toList
match {
case List(x) => x
@@ -81,15 +79,13 @@
private val key_listener =
JEdit_Lib.key_listener(
- key_pressed = (evt: KeyEvent) =>
- {
+ key_pressed = (evt: KeyEvent) => {
val mod = PIDE.options.string("jedit_focus_modifier")
val old = caret_focus_modifier
caret_focus_modifier = (mod.nonEmpty && mod == JEdit_Lib.modifier_string(evt))
if (caret_focus_modifier != old) caret_update()
},
- key_released = _ =>
- {
+ key_released = _ => {
if (caret_focus_modifier) {
caret_focus_modifier = false
caret_update()
@@ -103,12 +99,17 @@
@volatile private var painter_clip: Shape = null
@volatile private var caret_focus = Rendering.Focus.empty
- private val set_state = new TextAreaExtension
- {
- override def paintScreenLineRange(gfx: Graphics2D,
- first_line: Int, last_line: Int, physical_lines: Array[Int],
- start: Array[Int], end: Array[Int], y: Int, line_height: Int): Unit =
- {
+ private val set_state = new TextAreaExtension {
+ override def paintScreenLineRange(
+ gfx: Graphics2D,
+ first_line: Int,
+ last_line: Int,
+ physical_lines: Array[Int],
+ start: Array[Int],
+ end: Array[Int],
+ y: Int,
+ line_height: Int
+ ): Unit = {
painter_rendering = get_rendering()
painter_clip = gfx.getClip
caret_focus =
@@ -119,20 +120,24 @@
}
}
- private val reset_state = new TextAreaExtension
- {
- override def paintScreenLineRange(gfx: Graphics2D,
- first_line: Int, last_line: Int, physical_lines: Array[Int],
- start: Array[Int], end: Array[Int], y: Int, line_height: Int): Unit =
- {
+ private val reset_state = new TextAreaExtension {
+ override def paintScreenLineRange(
+ gfx: Graphics2D,
+ first_line: Int,
+ last_line: Int,
+ physical_lines: Array[Int],
+ start: Array[Int],
+ end: Array[Int],
+ y: Int,
+ line_height: Int
+ ): Unit = {
painter_rendering = null
painter_clip = null
caret_focus = Rendering.Focus.empty
}
}
- def robust_rendering(body: JEdit_Rendering => Unit): Unit =
- {
+ def robust_rendering(body: JEdit_Rendering => Unit): Unit = {
robust_body(()) { body(painter_rendering) }
}
@@ -141,16 +146,15 @@
private class Active_Area[A](
rendering: JEdit_Rendering => Text.Range => Option[Text.Info[A]],
- cursor: Option[Int])
- {
+ cursor: Option[Int]
+ ) {
private var the_text_info: Option[(String, Text.Info[A])] = None
def is_active: Boolean = the_text_info.isDefined
def text_info: Option[(String, Text.Info[A])] = the_text_info
def info: Option[Text.Info[A]] = the_text_info.map(_._2)
- def update(new_info: Option[Text.Info[A]]): Unit =
- {
+ def update(new_info: Option[Text.Info[A]]): Unit = {
val old_text_info = the_text_info
val new_text_info =
new_info.map(info => (text_area.getText(info.range.start, info.range.length), info))
@@ -210,8 +214,7 @@
}
private val mouse_listener = new MouseAdapter {
- override def mouseClicked(e: MouseEvent): Unit =
- {
+ override def mouseClicked(e: MouseEvent): Unit = {
robust_body(()) {
hyperlink_area.info match {
case Some(Text.Info(range, link)) =>
@@ -247,8 +250,7 @@
}
private val mouse_motion_listener = new MouseMotionAdapter {
- override def mouseDragged(evt: MouseEvent): Unit =
- {
+ override def mouseDragged(evt: MouseEvent): Unit = {
robust_body(()) {
active_reset()
Completion_Popup.Text_Area.dismissed(text_area)
@@ -256,8 +258,7 @@
}
}
- override def mouseMoved(evt: MouseEvent): Unit =
- {
+ override def mouseMoved(evt: MouseEvent): Unit = {
robust_body(()) {
val x = evt.getX
val y = evt.getY
@@ -269,8 +270,7 @@
case None => active_reset()
case Some(range) =>
val rendering = get_rendering()
- for ((area, require_control) <- active_areas)
- {
+ for ((area, require_control) <- active_areas) {
if (control == require_control && !rendering.snapshot.is_outdated)
area.update_rendering(rendering, range)
else area.reset()
@@ -310,12 +310,17 @@
/* text background */
- private val background_painter = new TextAreaExtension
- {
- override def paintScreenLineRange(gfx: Graphics2D,
- first_line: Int, last_line: Int, physical_lines: Array[Int],
- start: Array[Int], end: Array[Int], y: Int, line_height: Int): Unit =
- {
+ private val background_painter = new TextAreaExtension {
+ override def paintScreenLineRange(
+ gfx: Graphics2D,
+ first_line: Int,
+ last_line: Int,
+ physical_lines: Array[Int],
+ start: Array[Int],
+ end: Array[Int],
+ y: Int,
+ line_height: Int
+ ): Unit = {
robust_rendering { rendering =>
val fm = text_area.getPainter.getFontMetrics
@@ -324,8 +329,7 @@
val line_range = Text.Range(start(i), end(i) min buffer.getLength)
// line background color
- for { (c, separator) <- rendering.line_background(line_range) }
- {
+ for { (c, separator) <- rendering.line_background(line_range) } {
gfx.setColor(rendering.color(c))
val sep = if (separator) 2 min (line_height / 2) else 0
gfx.fillRect(0, y + i * line_height, text_area.getWidth, line_height - sep)
@@ -389,8 +393,7 @@
private def caret_enabled: Boolean =
caret_visible && (!text_area.hasFocus || text_area.isCaretVisible)
- private def caret_color(rendering: JEdit_Rendering, offset: Text.Offset): Color =
- {
+ private def caret_color(rendering: JEdit_Rendering, offset: Text.Offset): Color = {
if (text_area.isCaretVisible) text_area.getPainter.getCaretColor
else {
val debug_positions =
@@ -404,8 +407,7 @@
}
}
- private class Font_Subst
- {
+ private class Font_Subst {
private var cache = Map.empty[Int, Option[Font]]
def get(codepoint: Int): Option[Font] =
@@ -420,9 +422,15 @@
})
}
- private def paint_chunk_list(rendering: JEdit_Rendering, font_subst: Font_Subst,
- gfx: Graphics2D, line_start: Text.Offset, head: Chunk, x: Float, y: Float): Float =
- {
+ private def paint_chunk_list(
+ rendering: JEdit_Rendering,
+ font_subst: Font_Subst,
+ gfx: Graphics2D,
+ line_start: Text.Offset,
+ head: Chunk,
+ x: Float,
+ y: Float
+ ): Float = {
val clip_rect = gfx.getClipBounds
val caret_range =
@@ -434,8 +442,7 @@
while (chunk != null) {
val chunk_offset = line_start + chunk.offset
if (x + w + chunk.width > clip_rect.x &&
- x + w < clip_rect.x + clip_rect.width && chunk.length > 0)
- {
+ x + w < clip_rect.x + clip_rect.width && chunk.length > 0) {
val chunk_range = Text.Range(chunk_offset, chunk_offset + chunk.length)
val chunk_str =
if (chunk.chars == null) Symbol.spaces(chunk.length)
@@ -486,12 +493,17 @@
w
}
- private val text_painter = new TextAreaExtension
- {
- override def paintScreenLineRange(gfx: Graphics2D,
- first_line: Int, last_line: Int, physical_lines: Array[Int],
- start: Array[Int], end: Array[Int], y: Int, line_height: Int): Unit =
- {
+ private val text_painter = new TextAreaExtension {
+ override def paintScreenLineRange(
+ gfx: Graphics2D,
+ first_line: Int,
+ last_line: Int,
+ physical_lines: Array[Int],
+ start: Array[Int],
+ end: Array[Int],
+ y: Int,
+ line_height: Int
+ ): Unit = {
robust_rendering { rendering =>
val painter = text_area.getPainter
val fm = painter.getFontMetrics
@@ -501,8 +513,7 @@
val x0 = text_area.getHorizontalOffset
var y0 = y + painter.getLineHeight - (fm.getLeading + 1) - fm.getDescent
- val (bullet_x, bullet_y, bullet_w, bullet_h) =
- {
+ val (bullet_x, bullet_y, bullet_w, bullet_h) = {
val w = fm.charWidth(' ')
val b = (w / 2) max 1
val c = (lm.getAscent + lm.getStrikethroughOffset).round
@@ -551,12 +562,17 @@
/* foreground */
- private val foreground_painter = new TextAreaExtension
- {
- override def paintScreenLineRange(gfx: Graphics2D,
- first_line: Int, last_line: Int, physical_lines: Array[Int],
- start: Array[Int], end: Array[Int], y: Int, line_height: Int): Unit =
- {
+ private val foreground_painter = new TextAreaExtension {
+ override def paintScreenLineRange(
+ gfx: Graphics2D,
+ first_line: Int,
+ last_line: Int,
+ physical_lines: Array[Int],
+ start: Array[Int],
+ end: Array[Int],
+ y: Int,
+ line_height: Int
+ ): Unit = {
robust_rendering { rendering =>
val search_pattern = get_search_pattern()
for (i <- physical_lines.indices) {
@@ -635,11 +651,15 @@
/* caret -- outside of text range */
- private class Caret_Painter(before: Boolean) extends TextAreaExtension
- {
- override def paintValidLine(gfx: Graphics2D,
- screen_line: Int, physical_line: Int, start: Int, end: Int, y: Int): Unit =
- {
+ private class Caret_Painter(before: Boolean) extends TextAreaExtension {
+ override def paintValidLine(
+ gfx: Graphics2D,
+ screen_line: Int,
+ physical_line: Int,
+ start: Int,
+ end: Int,
+ y: Int
+ ): Unit = {
robust_rendering { _ =>
if (before) gfx.clipRect(0, 0, 0, 0)
else gfx.setClip(painter_clip)
@@ -652,11 +672,15 @@
private val before_caret_painter2 = new Caret_Painter(true)
private val after_caret_painter2 = new Caret_Painter(false)
- private val caret_painter = new TextAreaExtension
- {
- override def paintValidLine(gfx: Graphics2D,
- screen_line: Int, physical_line: Int, start: Int, end: Int, y: Int): Unit =
- {
+ private val caret_painter = new TextAreaExtension {
+ override def paintValidLine(
+ gfx: Graphics2D,
+ screen_line: Int,
+ physical_line: Int,
+ start: Int,
+ end: Int,
+ y: Int
+ ): Unit = {
robust_rendering { rendering =>
if (caret_visible) {
val caret = text_area.getCaretPosition
@@ -688,8 +712,7 @@
/* activation */
- def activate(): Unit =
- {
+ def activate(): Unit = {
val painter = text_area.getPainter
painter.addExtension(TextAreaPainter.LOWEST_LAYER, set_state)
painter.addExtension(TextAreaPainter.LINE_BACKGROUND_LAYER + 1, background_painter)
@@ -709,8 +732,7 @@
view.addWindowListener(window_listener)
}
- def deactivate(): Unit =
- {
+ def deactivate(): Unit = {
active_reset()
val painter = text_area.getPainter
view.removeWindowListener(window_listener)
--- a/src/Tools/jEdit/src/session_build.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/session_build.scala Fri Apr 01 17:06:10 2022 +0200
@@ -19,10 +19,8 @@
import org.gjt.sp.jedit.View
-object Session_Build
-{
- def check_dialog(view: View): Unit =
- {
+object Session_Build {
+ def check_dialog(view: View): Unit = {
val options = PIDE.options.value
Isabelle_Thread.fork() {
try {
@@ -38,8 +36,7 @@
}
}
- private class Dialog(view: View) extends JDialog(view)
- {
+ private class Dialog(view: View) extends JDialog(view) {
val options: Options = PIDE.options.value
@@ -78,8 +75,7 @@
setContentPane(layout_panel.peer)
- private def set_actions(cs: Component*): Unit =
- {
+ private def set_actions(cs: Component*): Unit = {
action_panel.contents.clear()
action_panel.contents ++= cs
layout_panel.revalidate()
@@ -98,8 +94,7 @@
}
private val delay_exit =
- Delay.first(Time.seconds(1.0), gui = true)
- {
+ Delay.first(Time.seconds(1.0), gui = true) {
if (can_auto_close) conclude()
else {
val button = new Button("Close") { reactions += { case ButtonClicked(_) => conclude() } }
@@ -108,8 +103,7 @@
}
}
- private def conclude(): Unit =
- {
+ private def conclude(): Unit = {
setVisible(false)
dispose()
}
@@ -120,15 +114,13 @@
setDefaultCloseOperation(WindowConstants.DO_NOTHING_ON_CLOSE)
addWindowListener(new WindowAdapter {
- override def windowClosing(e: WindowEvent): Unit =
- {
+ override def windowClosing(e: WindowEvent): Unit = {
if (_return_code.isDefined) conclude()
else stopping()
}
})
- private def stopping(): Unit =
- {
+ private def stopping(): Unit = {
progress.stop()
set_actions(new Label("Stopping ..."))
}
--- a/src/Tools/jEdit/src/simplifier_trace_dockable.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/simplifier_trace_dockable.scala Fri Apr 01 17:06:10 2022 +0200
@@ -18,8 +18,7 @@
import org.gjt.sp.jedit.View
-class Simplifier_Trace_Dockable(view: View, position: String) extends Dockable(view, position)
-{
+class Simplifier_Trace_Dockable(view: View, position: String) extends Dockable(view, position) {
GUI_Thread.require {}
@@ -35,8 +34,7 @@
private val text_area = new Pretty_Text_Area(view)
set_content(text_area)
- private def update_contents(): Unit =
- {
+ private def update_contents(): Unit = {
val snapshot = current_snapshot
val context = Simplifier_Trace.handle_results(PIDE.session, current_id, current_results)
@@ -61,14 +59,12 @@
do_paint()
}
- private def show_trace(): Unit =
- {
+ private def show_trace(): Unit = {
val trace = Simplifier_Trace.generate_trace(PIDE.session, current_results)
new Simplifier_Trace_Window(view, current_snapshot, trace)
}
- private def do_paint(): Unit =
- {
+ private def do_paint(): Unit = {
GUI_Thread.later {
text_area.resize(Font_Info.main(PIDE.options.real("jedit_font_scale")))
}
@@ -76,8 +72,7 @@
private def handle_resize(): Unit = do_paint()
- private def handle_update(follow: Boolean): Unit =
- {
+ private def handle_update(follow: Boolean): Unit = {
val (new_snapshot, new_command, new_results, new_id) =
PIDE.editor.current_node_snapshot(view) match {
case Some(snapshot) =>
@@ -118,8 +113,7 @@
GUI_Thread.later { handle_update(do_update) }
}
- override def init(): Unit =
- {
+ override def init(): Unit = {
PIDE.session.global_options += main
PIDE.session.commands_changed += main
PIDE.session.caret_focus += main
@@ -127,8 +121,7 @@
handle_update(true)
}
- override def exit(): Unit =
- {
+ override def exit(): Unit = {
PIDE.session.global_options -= main
PIDE.session.commands_changed -= main
PIDE.session.caret_focus -= main
--- a/src/Tools/jEdit/src/simplifier_trace_window.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/simplifier_trace_window.scala Fri Apr 01 17:06:10 2022 +0200
@@ -23,10 +23,8 @@
import org.gjt.sp.jedit.View
-private object Simplifier_Trace_Window
-{
- sealed trait Trace_Tree
- {
+private object Simplifier_Trace_Window {
+ sealed trait Trace_Tree {
// FIXME replace with immutable tree builder
var children: SortedMap[Long, Either[Simplifier_Trace.Item.Data, Elem_Tree]] = SortedMap.empty
val serial: Long
@@ -39,8 +37,7 @@
}
}
- final class Root_Tree(val serial: Long) extends Trace_Tree
- {
+ final class Root_Tree(val serial: Long) extends Trace_Tree {
val parent = None
val interesting = true
val markup = ""
@@ -50,8 +47,7 @@
}
final class Elem_Tree(data: Simplifier_Trace.Item.Data, val parent: Option[Trace_Tree])
- extends Trace_Tree
- {
+ extends Trace_Tree {
val serial = data.serial
val markup = data.markup
@@ -63,8 +59,7 @@
private def body_contains(regex: Regex, body: XML.Body): Boolean =
body.exists(tree => regex.findFirstIn(XML.content(tree)).isDefined)
- def format: Option[XML.Tree] =
- {
+ def format: Option[XML.Tree] = {
def format_hint(data: Simplifier_Trace.Item.Data): XML.Tree =
Pretty.block(Pretty.separate(XML.Text(data.text) :: data.content))
@@ -94,8 +89,7 @@
case head :: tail =>
lookup.get(head.parent) match {
case Some(parent) =>
- if (head.markup == Markup.SIMP_TRACE_HINT)
- {
+ if (head.markup == Markup.SIMP_TRACE_HINT) {
head.props match {
case Simplifier_Trace.Success(x)
if x ||
@@ -107,8 +101,7 @@
}
walk_trace(tail, lookup)
}
- else if (head.markup == Markup.SIMP_TRACE_IGNORE)
- {
+ else if (head.markup == Markup.SIMP_TRACE_IGNORE) {
parent.parent match {
case None =>
Output.error_message(
@@ -118,8 +111,7 @@
walk_trace(tail, lookup)
}
}
- else
- {
+ else {
val entry = new Elem_Tree(head, Some(parent))
parent.children += ((head.serial, Right(entry)))
walk_trace(tail, lookup + (head.serial -> entry))
@@ -134,8 +126,10 @@
class Simplifier_Trace_Window(
- view: View, snapshot: Document.Snapshot, trace: Simplifier_Trace.Trace) extends Frame
-{
+ view: View,
+ snapshot: Document.Snapshot,
+ trace: Simplifier_Trace.Trace
+) extends Frame {
GUI_Thread.require {}
private val pretty_text_area = new Pretty_Text_Area(view)
@@ -159,14 +153,12 @@
open()
do_paint()
- def do_update(): Unit =
- {
+ def do_update(): Unit = {
val xml = tree.format
pretty_text_area.update(snapshot, Command.Results.empty, xml)
}
- def do_paint(): Unit =
- {
+ def do_paint(): Unit = {
GUI_Thread.later {
pretty_text_area.resize(
Font_Info.main(PIDE.options.real("jedit_font_scale") * zoom.factor / 100))
--- a/src/Tools/jEdit/src/sledgehammer_dockable.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/sledgehammer_dockable.scala Fri Apr 01 17:06:10 2022 +0200
@@ -19,8 +19,7 @@
import org.gjt.sp.jedit.gui.HistoryTextField
-class Sledgehammer_Dockable(view: View, position: String) extends Dockable(view, position)
-{
+class Sledgehammer_Dockable(view: View, position: String) extends Dockable(view, position) {
GUI_Thread.require {}
@@ -36,8 +35,7 @@
private val process_indicator = new Process_Indicator
- private def consume_status(status: Query_Operation.Status.Value): Unit =
- {
+ private def consume_status(status: Query_Operation.Status.Value): Unit = {
status match {
case Query_Operation.Status.WAITING =>
process_indicator.update("Waiting for evaluation of context ...", 5)
@@ -64,8 +62,7 @@
override def componentShown(e: ComponentEvent): Unit = delay_resize.invoke()
})
- private def handle_resize(): Unit =
- {
+ private def handle_resize(): Unit = {
GUI_Thread.require {}
pretty_text_area.resize(
@@ -75,8 +72,7 @@
/* controls */
- private def clicked: Unit =
- {
+ private def clicked: Unit = {
provers.addCurrentToHistory()
PIDE.options.string("sledgehammer_provers") = provers.getText
sledgehammer.apply_query(
@@ -91,8 +87,7 @@
}
private val provers = new HistoryTextField("isabelle-sledgehammer-provers") {
- override def processKeyEvent(evt: KeyEvent): Unit =
- {
+ override def processKeyEvent(evt: KeyEvent): Unit = {
if (evt.getID == KeyEvent.KEY_PRESSED && evt.getKeyCode == KeyEvent.VK_ENTER) clicked
super.processKeyEvent(evt)
}
@@ -100,8 +95,7 @@
setColumns(30)
}
- private def update_provers(): Unit =
- {
+ private def update_provers(): Unit = {
val new_provers = PIDE.options.string("sledgehammer_provers")
if (new_provers != provers.getText) {
provers.setText(new_provers)
@@ -154,16 +148,14 @@
case _: Session.Global_Options => GUI_Thread.later { update_provers(); handle_resize() }
}
- override def init(): Unit =
- {
+ override def init(): Unit = {
PIDE.session.global_options += main
update_provers()
handle_resize()
sledgehammer.activate()
}
- override def exit(): Unit =
- {
+ override def exit(): Unit = {
sledgehammer.deactivate()
PIDE.session.global_options -= main
delay_resize.revoke()
--- a/src/Tools/jEdit/src/state_dockable.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/state_dockable.scala Fri Apr 01 17:06:10 2022 +0200
@@ -18,8 +18,7 @@
import org.gjt.sp.jedit.View
-class State_Dockable(view: View, position: String) extends Dockable(view, position)
-{
+class State_Dockable(view: View, position: String) extends Dockable(view, position) {
GUI_Thread.require {}
@@ -46,8 +45,7 @@
override def componentShown(e: ComponentEvent): Unit = delay_resize.invoke()
})
- private def handle_resize(): Unit =
- {
+ private def handle_resize(): Unit = {
GUI_Thread.require {}
pretty_text_area.resize(
@@ -60,8 +58,7 @@
def update_request(): Unit =
GUI_Thread.require { print_state.apply_query(Nil) }
- def update(): Unit =
- {
+ def update(): Unit = {
GUI_Thread.require {}
PIDE.editor.current_node_snapshot(view) match {
@@ -125,8 +122,7 @@
GUI_Thread.later { auto_update() }
}
- override def init(): Unit =
- {
+ override def init(): Unit = {
PIDE.session.global_options += main
PIDE.session.commands_changed += main
PIDE.session.caret_focus += main
@@ -135,8 +131,7 @@
auto_update()
}
- override def exit(): Unit =
- {
+ override def exit(): Unit = {
print_state.deactivate()
PIDE.session.caret_focus -= main
PIDE.session.global_options -= main
--- a/src/Tools/jEdit/src/status_widget.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/status_widget.scala Fri Apr 01 17:06:10 2022 +0200
@@ -17,14 +17,12 @@
import org.gjt.sp.jedit.gui.statusbar.{StatusWidgetFactory, Widget}
-object Status_Widget
-{
+object Status_Widget {
/** generic GUI **/
private val template = "ABC: 99999/99999MB"
- abstract class GUI(view: View) extends JComponent
- {
+ abstract class GUI(view: View) extends JComponent {
/* init */
setFont(new JLabel().getFont)
@@ -50,8 +48,7 @@
def get_status: (String, Double)
- override def paintComponent(gfx: Graphics): Unit =
- {
+ override def paintComponent(gfx: Graphics): Unit = {
super.paintComponent(gfx)
val insets = new Insets(0, 0, 0, 0)
@@ -89,8 +86,7 @@
def action: String
addMouseListener(new MouseAdapter {
- override def mouseClicked(evt: MouseEvent): Unit =
- {
+ override def mouseClicked(evt: MouseEvent): Unit = {
if (evt.getClickCount == 2) {
view.getInputHandler.invokeAction(action)
}
@@ -102,20 +98,17 @@
/** Java **/
- class Java_GUI(view: View) extends GUI(view)
- {
+ class Java_GUI(view: View) extends GUI(view) {
/* component state -- owned by GUI thread */
private var status = Java_Statistics.memory_status()
- def get_status: (String, Double) =
- {
+ def get_status: (String, Double) = {
("JVM: " + (status.heap_used / 1024 / 1024) + "/" + (status.heap_size / 1024 / 1024) + "MB",
status.heap_used_fraction)
}
- private def update_status(new_status: Java_Statistics.Memory_Status): Unit =
- {
+ private def update_status(new_status: Java_Statistics.Memory_Status): Unit = {
if (status != new_status) {
status = new_status
repaint()
@@ -131,14 +124,12 @@
update_status(Java_Statistics.memory_status())
})
- override def addNotify(): Unit =
- {
+ override def addNotify(): Unit = {
super.addNotify()
timer.start()
}
- override def removeNotify(): Unit =
- {
+ override def removeNotify(): Unit = {
super.removeNotify()
timer.stop()
}
@@ -151,8 +142,7 @@
override def action: String = "isabelle.java-monitor"
}
- class Java_Factory extends StatusWidgetFactory
- {
+ class Java_Factory extends StatusWidgetFactory {
override def getWidget(view: View): Widget =
new Widget { override def getComponent: JComponent = new Java_GUI(view) }
}
@@ -161,8 +151,7 @@
/** ML **/
- class ML_GUI(view: View) extends GUI(view)
- {
+ class ML_GUI(view: View) extends GUI(view) {
/* component state -- owned by GUI thread */
private var status = ML_Statistics.memory_status(Nil)
@@ -175,8 +164,7 @@
status.heap_used_fraction)
}
- private def update_status(new_status: ML_Statistics.Memory_Status): Unit =
- {
+ private def update_status(new_status: ML_Statistics.Memory_Status): Unit = {
if (status != new_status) {
status = new_status
repaint()
@@ -193,14 +181,12 @@
GUI_Thread.later { update_status(status) }
}
- override def addNotify(): Unit =
- {
+ override def addNotify(): Unit = {
super.addNotify()
PIDE.session.runtime_statistics += main
}
- override def removeNotify(): Unit =
- {
+ override def removeNotify(): Unit = {
super.removeNotify()
PIDE.session.runtime_statistics -= main
}
@@ -213,8 +199,7 @@
override def action: String = "isabelle-monitor-float"
}
- class ML_Factory extends StatusWidgetFactory
- {
+ class ML_Factory extends StatusWidgetFactory {
override def getWidget(view: View): Widget =
new Widget { override def getComponent: JComponent = new ML_GUI(view) }
}
--- a/src/Tools/jEdit/src/symbols_dockable.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/symbols_dockable.scala Fri Apr 01 17:06:10 2022 +0200
@@ -16,8 +16,7 @@
import org.gjt.sp.jedit.{EditBus, EBComponent, EBMessage, View}
-class Symbols_Dockable(view: View, position: String) extends Dockable(view, position)
-{
+class Symbols_Dockable(view: View, position: String) extends Dockable(view, position) {
private def font_size: Int =
Font_Info.main_size(PIDE.options.real("jedit_font_scale")).round
@@ -29,8 +28,7 @@
private val abbrevs_refresh_delay =
Delay.last(PIDE.options.seconds("editor_update_delay"), gui = true) { abbrevs_panel.refresh() }
- private class Abbrev_Component(txt: String, abbrs: List[String]) extends Button
- {
+ private class Abbrev_Component(txt: String, abbrs: List[String]) extends Button {
def update_font(): Unit = { font = GUI.font(size = font_size) }
update_font()
@@ -47,8 +45,7 @@
tooltip = GUI.tooltip_lines(cat_lines(txt :: abbrs.map(a => "abbrev: " + a)))
}
- private class Abbrevs_Panel extends Wrap_Panel(Nil, Wrap_Panel.Alignment.Center)
- {
+ private class Abbrevs_Panel extends Wrap_Panel(Nil, Wrap_Panel.Alignment.Center) {
private var abbrevs: Thy_Header.Abbrevs = Nil
def refresh(): Unit = GUI_Thread.require {
@@ -81,10 +78,8 @@
/* symbols */
- private class Symbol_Component(val symbol: String, is_control: Boolean) extends Button
- {
- def update_font(): Unit =
- {
+ private class Symbol_Component(val symbol: String, is_control: Boolean) extends Button {
+ def update_font(): Unit = {
font =
Symbol.symbols.fonts.get(symbol) match {
case None => GUI.font(size = font_size)
@@ -107,8 +102,7 @@
cat_lines(symbol :: Symbol.symbols.get_abbrevs(symbol).map(a => "abbrev: " + a)))
}
- private class Reset_Component extends Button
- {
+ private class Reset_Component extends Button {
action = Action("Reset") {
val text_area = view.getTextArea
Syntax_Style.edit_control_style(text_area, "")
@@ -207,15 +201,13 @@
case _: Session.Commands_Changed => abbrevs_refresh_delay.invoke()
}
- override def init(): Unit =
- {
+ override def init(): Unit = {
EditBus.addToBus(edit_bus_handler)
PIDE.session.global_options += main
PIDE.session.commands_changed += main
}
- override def exit(): Unit =
- {
+ override def exit(): Unit = {
EditBus.removeFromBus(edit_bus_handler)
PIDE.session.global_options -= main
PIDE.session.commands_changed -= main
--- a/src/Tools/jEdit/src/syntax_style.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/syntax_style.scala Fri Apr 01 17:06:10 2022 +0200
@@ -20,8 +20,7 @@
import org.gjt.sp.jedit.textarea.TextArea
-object Syntax_Style
-{
+object Syntax_Style {
/* extended syntax styles */
private val plain_range: Int = JEditToken.ID_COUNT
@@ -39,8 +38,7 @@
private def font_style(style: SyntaxStyle, f: Font => Font): SyntaxStyle =
new SyntaxStyle(style.getForegroundColor, style.getBackgroundColor, f(style.getFont))
- private def script_style(style: SyntaxStyle, i: Int): SyntaxStyle =
- {
+ private def script_style(style: SyntaxStyle, i: Int): SyntaxStyle = {
font_style(style, font0 =>
{
val font1 = font0.deriveFont(JMap.of(TextAttribute.SUPERSCRIPT, java.lang.Integer.valueOf(i)))
@@ -63,16 +61,13 @@
val hidden_color: Color = new Color(255, 255, 255, 0)
- def set_extender(extender: SyntaxUtilities.StyleExtender): Unit =
- {
+ def set_extender(extender: SyntaxUtilities.StyleExtender): Unit = {
SyntaxUtilities.setStyleExtender(extender)
GUI_Thread.later { jEdit.propertiesChanged }
}
- object Base_Extender extends SyntaxUtilities.StyleExtender
- {
- override def extendStyles(styles: Array[SyntaxStyle]): Array[SyntaxStyle] =
- {
+ object Base_Extender extends SyntaxUtilities.StyleExtender {
+ override def extendStyles(styles: Array[SyntaxStyle]): Array[SyntaxStyle] = {
val new_styles = Array.fill[SyntaxStyle](java.lang.Byte.MAX_VALUE)(styles(0))
for (i <- 0 until full_range) {
new_styles(i) = styles(i % plain_range)
@@ -81,15 +76,13 @@
}
}
- object Main_Extender extends SyntaxUtilities.StyleExtender
- {
+ object Main_Extender extends SyntaxUtilities.StyleExtender {
val max_user_fonts = 2
if (Symbol.symbols.font_names.length > max_user_fonts)
error("Too many user symbol fonts (" + max_user_fonts + " permitted): " +
Symbol.symbols.font_names.mkString(", "))
- override def extendStyles(styles: Array[SyntaxStyle]): Array[SyntaxStyle] =
- {
+ override def extendStyles(styles: Array[SyntaxStyle]): Array[SyntaxStyle] = {
val style0 = styles(0)
val font0 = style0.getFont
@@ -125,11 +118,9 @@
else if (sym == Symbol.bold_decoded) Some(bold)
else None
- def extended(text: CharSequence): Map[Text.Offset, Byte => Byte] =
- {
+ def extended(text: CharSequence): Map[Text.Offset, Byte => Byte] = {
var result = Map[Text.Offset, Byte => Byte]()
- def mark(start: Text.Offset, stop: Text.Offset, style: Byte => Byte): Unit =
- {
+ def mark(start: Text.Offset, stop: Text.Offset, style: Byte => Byte): Unit = {
for (i <- start until stop) result += (i -> style)
}
@@ -169,16 +160,14 @@
/* editing support for control symbols */
- def edit_control_style(text_area: TextArea, control_sym: String): Unit =
- {
+ def edit_control_style(text_area: TextArea, control_sym: String): Unit = {
GUI_Thread.assert {}
val buffer = text_area.getBuffer
val control_decoded = Isabelle_Encoding.perhaps_decode(buffer, control_sym)
- def update_style(text: String): String =
- {
+ def update_style(text: String): String = {
val result = new StringBuilder
for (sym <- Symbol.iterator(text) if !HTML.is_control(sym)) {
if (Symbol.is_controllable(sym)) result ++= control_decoded
--- a/src/Tools/jEdit/src/syslog_dockable.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/syslog_dockable.scala Fri Apr 01 17:06:10 2022 +0200
@@ -14,14 +14,12 @@
import org.gjt.sp.jedit.View
-class Syslog_Dockable(view: View, position: String) extends Dockable(view, position)
-{
+class Syslog_Dockable(view: View, position: String) extends Dockable(view, position) {
/* GUI components */
private val syslog = new TextArea()
- private def syslog_delay = Delay.first(PIDE.options.seconds("editor_update_delay"), gui = true)
- {
+ private def syslog_delay = Delay.first(PIDE.options.seconds("editor_update_delay"), gui = true) {
val text = PIDE.session.syslog_content()
if (text != syslog.text) syslog.text = text
}
@@ -34,14 +32,12 @@
private val main =
Session.Consumer[Prover.Output](getClass.getName) { case _ => syslog_delay.invoke() }
- override def init(): Unit =
- {
+ override def init(): Unit = {
PIDE.session.syslog_messages += main
syslog_delay.invoke()
}
- override def exit(): Unit =
- {
+ override def exit(): Unit = {
PIDE.session.syslog_messages -= main
}
}
--- a/src/Tools/jEdit/src/text_overview.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/text_overview.scala Fri Apr 01 17:06:10 2022 +0200
@@ -16,8 +16,7 @@
import javax.swing.{JPanel, ToolTipManager}
-class Text_Overview(doc_view: Document_View) extends JPanel(new BorderLayout)
-{
+class Text_Overview(doc_view: Document_View) extends JPanel(new BorderLayout) {
/* GUI components */
private val text_area = doc_view.text_area
@@ -33,22 +32,19 @@
setRequestFocusEnabled(false)
addMouseListener(new MouseAdapter {
- override def mousePressed(event: MouseEvent): Unit =
- {
+ override def mousePressed(event: MouseEvent): Unit = {
val line = (event.getY * lines()) / getHeight
if (line >= 0 && line < text_area.getLineCount)
text_area.setCaretPosition(text_area.getLineStartOffset(line))
}
})
- override def addNotify(): Unit =
- {
+ override def addNotify(): Unit = {
super.addNotify()
ToolTipManager.sharedInstance.registerComponent(this)
}
- override def removeNotify(): Unit =
- {
+ override def removeNotify(): Unit = {
ToolTipManager.sharedInstance.unregisterComponent(this)
super.removeNotify
}
@@ -80,8 +76,7 @@
private val delay_repaint =
Delay.first(PIDE.options.seconds("editor_update_delay"), gui = true) { repaint() }
- override def paintComponent(gfx: Graphics): Unit =
- {
+ override def paintComponent(gfx: Graphics): Unit = {
super.paintComponent(gfx)
GUI_Thread.assert {}
@@ -128,8 +123,7 @@
if (rendering.snapshot.is_outdated || is_running()) invoke()
else {
- val line_offsets =
- {
+ val line_offsets = {
val line_manager = JEdit_Lib.buffer_line_manager(buffer)
val a = new Array[Int](line_manager.getLineCount)
for (i <- 1 until a.length) a(i) = line_manager.getLineEndOffset(i - 1)
@@ -143,9 +137,13 @@
val L = overview.L
val H = overview.H
- @tailrec def loop(l: Int, h: Int, p: Int, q: Int, colors: List[Color_Info])
- : List[Color_Info] =
- {
+ @tailrec def loop(
+ l: Int,
+ h: Int,
+ p: Int,
+ q: Int,
+ colors: List[Color_Info]
+ ): List[Color_Info] = {
Exn.Interrupt.expose()
if (l < line_count && h < H) {
--- a/src/Tools/jEdit/src/text_structure.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/text_structure.scala Fri Apr 01 17:06:10 2022 +0200
@@ -17,22 +17,18 @@
import org.gjt.sp.jedit.Buffer
-object Text_Structure
-{
+object Text_Structure {
/* token navigator */
- class Navigator(syntax: Outer_Syntax, buffer: JEditBuffer, comments: Boolean)
- {
+ class Navigator(syntax: Outer_Syntax, buffer: JEditBuffer, comments: Boolean) {
val limit: Int = PIDE.options.value.int("jedit_structure_limit") max 0
- def iterator(line: Int, lim: Int = limit): Iterator[Text.Info[Token]] =
- {
+ def iterator(line: Int, lim: Int = limit): Iterator[Text.Info[Token]] = {
val it = Token_Markup.line_token_iterator(syntax, buffer, line, line + lim)
if (comments) it.filterNot(_.info.is_space) else it.filter(_.info.is_proper)
}
- def reverse_iterator(line: Int, lim: Int = limit): Iterator[Text.Info[Token]] =
- {
+ def reverse_iterator(line: Int, lim: Int = limit): Iterator[Text.Info[Token]] = {
val it = Token_Markup.line_token_reverse_iterator(syntax, buffer, line, line - lim)
if (comments) it.filterNot(_.info.is_space) else it.filter(_.info.is_proper)
}
@@ -41,14 +37,17 @@
/* indentation */
- object Indent_Rule extends IndentRule
- {
+ object Indent_Rule extends IndentRule {
private val keyword_open = Keyword.theory_goal ++ Keyword.proof_open
private val keyword_close = Keyword.proof_close
- def apply(buffer: JEditBuffer, current_line: Int, prev_line0: Int, prev_prev_line0: Int,
- actions: JList[IndentAction]): Unit =
- {
+ def apply(
+ buffer: JEditBuffer,
+ current_line: Int,
+ prev_line0: Int,
+ prev_prev_line0: Int,
+ actions: JList[IndentAction]
+ ): Unit = {
Isabelle.buffer_syntax(buffer) match {
case Some(syntax) =>
val keywords = syntax.keywords
@@ -87,8 +86,7 @@
nav.reverse_iterator(prev_line).map(_.info).takeWhile(tok => !tok.is_begin_or_command)
- val script_indent: Text.Info[Token] => Int =
- {
+ val script_indent: Text.Info[Token] => Int = {
val opt_rendering: Option[JEdit_Rendering] =
if (PIDE.options.value.bool("jedit_indent_script"))
GUI_Thread.now {
@@ -193,18 +191,24 @@
}
}
- def line_content(buffer: JEditBuffer, keywords: Keyword.Keywords,
- range: Text.Range, ctxt: Scan.Line_Context): (List[Token], Scan.Line_Context) =
- {
+ def line_content(
+ buffer: JEditBuffer,
+ keywords: Keyword.Keywords,
+ range: Text.Range,
+ ctxt: Scan.Line_Context
+ ): (List[Token], Scan.Line_Context) = {
val text = JEdit_Lib.get_text(buffer, range).getOrElse("")
val (toks, ctxt1) = Token.explode_line(keywords, text, ctxt)
val toks1 = toks.filterNot(_.is_space)
(toks1, ctxt1)
}
- def split_line_content(buffer: JEditBuffer, keywords: Keyword.Keywords, line: Int, caret: Int)
- : (List[Token], List[Token]) =
- {
+ def split_line_content(
+ buffer: JEditBuffer,
+ keywords: Keyword.Keywords,
+ line: Int,
+ caret: Int
+ ): (List[Token], List[Token]) = {
val line_range = JEdit_Lib.line_range(buffer, line)
val ctxt0 = Token_Markup.Line_Context.before(buffer, line).get_context
val (toks1, ctxt1) = line_content(buffer, keywords, Text.Range(line_range.start, caret), ctxt0)
@@ -215,15 +219,14 @@
/* structure matching */
- object Matcher extends StructureMatcher
- {
+ object Matcher extends StructureMatcher {
private def find_block(
open: Token => Boolean,
close: Token => Boolean,
reset: Token => Boolean,
restrict: Token => Boolean,
- it: Iterator[Text.Info[Token]]): Option[(Text.Range, Text.Range)] =
- {
+ it: Iterator[Text.Info[Token]]
+ ): Option[(Text.Range, Text.Range)] = {
val range1 = it.next().range
it.takeWhile(info => !info.info.is_command || restrict(info.info)).
scanLeft((range1, 1))(
@@ -235,8 +238,7 @@
).collectFirst({ case (range2, 0) => (range1, range2) })
}
- private def find_pair(text_area: TextArea): Option[(Text.Range, Text.Range)] =
- {
+ private def find_pair(text_area: TextArea): Option[(Text.Range, Text.Range)] = {
val buffer = text_area.getBuffer
val caret_line = text_area.getCaretLine
val caret = text_area.getCaretPosition
@@ -330,8 +332,7 @@
case None => null
}
- def selectMatch(text_area: TextArea): Unit =
- {
+ def selectMatch(text_area: TextArea): Unit = {
def get_span(offset: Text.Offset): Option[Text.Range] =
for {
syntax <- Isabelle.buffer_syntax(text_area.getBuffer)
--- a/src/Tools/jEdit/src/theories_dockable.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/theories_dockable.scala Fri Apr 01 17:06:10 2022 +0200
@@ -20,13 +20,11 @@
import org.gjt.sp.jedit.{View, jEdit}
-class Theories_Dockable(view: View, position: String) extends Dockable(view, position)
-{
+class Theories_Dockable(view: View, position: String) extends Dockable(view, position) {
/* status */
private val status = new ListView(List.empty[Document.Node.Name]) {
- background =
- {
+ background = {
// enforce default value
val c = UIManager.getDefaults.getColor("Panel.background")
new Color(c.getRed, c.getGreen, c.getBlue, c.getAlpha)
@@ -72,8 +70,7 @@
session_phase.border = new SoftBevelBorder(BevelBorder.LOWERED)
session_phase.tooltip = "Status of prover session"
- private def handle_phase(phase: Session.Phase): Unit =
- {
+ private def handle_phase(phase: Session.Phase): Unit = {
GUI_Thread.require {}
session_phase.text = " " + phase_text(phase) + " "
}
@@ -114,8 +111,7 @@
Node_Renderer_Component != null && in(Node_Renderer_Component.label_geometry, loc0, p)
- private object Node_Renderer_Component extends BorderPanel
- {
+ private object Node_Renderer_Component extends BorderPanel {
opaque = true
border = BorderFactory.createEmptyBorder(2, 2, 2, 2)
@@ -124,8 +120,7 @@
var checkbox_geometry: Option[(Point, Dimension)] = None
val checkbox = new CheckBox {
opaque = false
- override def paintComponent(gfx: Graphics2D): Unit =
- {
+ override def paintComponent(gfx: Graphics2D): Unit = {
super.paintComponent(gfx)
if (location != null && size != null)
checkbox_geometry = Some((location, size))
@@ -139,10 +134,8 @@
opaque = false
xAlignment = Alignment.Leading
- override def paintComponent(gfx: Graphics2D): Unit =
- {
- def paint_segment(x: Int, w: Int, color: Color): Unit =
- {
+ override def paintComponent(gfx: Graphics2D): Unit = {
+ def paint_segment(x: Int, w: Int, color: Color): Unit = {
gfx.setColor(color)
gfx.fillRect(x, 0, w, size.height)
}
@@ -175,8 +168,7 @@
}
}
- def label_border(name: Document.Node.Name): Unit =
- {
+ def label_border(name: Document.Node.Name): Unit = {
val st = nodes_status.overall_node_status(name)
val color =
st match {
@@ -199,11 +191,14 @@
layout(label) = BorderPanel.Position.Center
}
- private class Node_Renderer extends ListView.Renderer[Document.Node.Name]
- {
- def componentFor(list: ListView[_ <: isabelle.Document.Node.Name], isSelected: Boolean,
- focused: Boolean, name: Document.Node.Name, index: Int): Component =
- {
+ private class Node_Renderer extends ListView.Renderer[Document.Node.Name] {
+ def componentFor(
+ list: ListView[_ <: isabelle.Document.Node.Name],
+ isSelected: Boolean,
+ focused: Boolean,
+ name: Document.Node.Name,
+ index: Int
+ ): Component = {
val component = Node_Renderer_Component
component.node_name = name
component.checkbox.selected = nodes_required.contains(name)
@@ -215,8 +210,9 @@
status.renderer = new Node_Renderer
private def handle_update(
- domain: Option[Set[Document.Node.Name]] = None, trim: Boolean = false): Unit =
- {
+ domain: Option[Set[Document.Node.Name]] = None,
+ trim: Boolean = false
+ ): Unit = {
GUI_Thread.require {}
val snapshot = PIDE.session.snapshot()
@@ -255,8 +251,7 @@
GUI_Thread.later { handle_update(domain = Some(changed.nodes), trim = changed.assignment) }
}
- override def init(): Unit =
- {
+ override def init(): Unit = {
PIDE.session.phase_changed += main
PIDE.session.global_options += main
PIDE.session.commands_changed += main
@@ -265,8 +260,7 @@
handle_update()
}
- override def exit(): Unit =
- {
+ override def exit(): Unit = {
PIDE.session.phase_changed -= main
PIDE.session.global_options -= main
PIDE.session.commands_changed -= main
--- a/src/Tools/jEdit/src/timing_dockable.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/timing_dockable.scala Fri Apr 01 17:06:10 2022 +0200
@@ -19,29 +19,23 @@
import org.gjt.sp.jedit.{View, jEdit}
-class Timing_Dockable(view: View, position: String) extends Dockable(view, position)
-{
+class Timing_Dockable(view: View, position: String) extends Dockable(view, position) {
/* entry */
- private object Entry
- {
- object Ordering extends scala.math.Ordering[Entry]
- {
+ private object Entry {
+ object Ordering extends scala.math.Ordering[Entry] {
def compare(entry1: Entry, entry2: Entry): Int =
entry2.timing compare entry1.timing
}
- object Renderer_Component extends Label
- {
+ object Renderer_Component extends Label {
opaque = false
xAlignment = Alignment.Leading
border = BorderFactory.createEmptyBorder(2, 2, 2, 2)
var entry: Entry = null
- override def paintComponent(gfx: Graphics2D): Unit =
- {
- def paint_rectangle(color: Color): Unit =
- {
+ override def paintComponent(gfx: Graphics2D): Unit = {
+ def paint_rectangle(color: Color): Unit = {
val size = peer.getSize()
val insets = border.getBorderInsets(peer)
val x = insets.left
@@ -63,11 +57,13 @@
}
}
- class Renderer extends ListView.Renderer[Entry]
- {
- def componentFor(list: ListView[_ <: Timing_Dockable.this.Entry],
- isSelected: Boolean, focused: Boolean, entry: Entry, index: Int): Component =
- {
+ class Renderer extends ListView.Renderer[Entry] {
+ def componentFor(
+ list: ListView[_ <: Timing_Dockable.this.Entry],
+ isSelected: Boolean,
+ focused: Boolean,
+ entry: Entry, index: Int
+ ): Component = {
val component = Renderer_Component
component.entry = entry
component.text = entry.print
@@ -76,24 +72,22 @@
}
}
- private abstract class Entry
- {
+ private abstract class Entry {
def timing: Double
def print: String
def follow(snapshot: Document.Snapshot): Unit
}
private case class Theory_Entry(name: Document.Node.Name, timing: Double, current: Boolean)
- extends Entry
- {
+ extends Entry {
def print: String =
Time.print_seconds(timing) + "s theory " + quote(name.theory)
def follow(snapshot: Document.Snapshot): Unit =
PIDE.editor.goto_file(true, view, name.node)
}
- private case class Command_Entry(command: Command, timing: Double) extends Entry
- {
+ private case class Command_Entry(command: Command, timing: Double)
+ extends Entry {
def print: String =
" " + Time.print_seconds(timing) + "s command " + quote(command.span.name)
def follow(snapshot: Document.Snapshot): Unit =
@@ -152,8 +146,7 @@
private var nodes_timing = Map.empty[Document.Node.Name, Document_Status.Overall_Timing]
- private def make_entries(): List[Entry] =
- {
+ private def make_entries(): List[Entry] = {
GUI_Thread.require {}
val name =
@@ -175,8 +168,7 @@
else List(entry))
}
- private def handle_update(restriction: Option[Set[Document.Node.Name]] = None): Unit =
- {
+ private def handle_update(restriction: Option[Set[Document.Node.Name]] = None): Unit = {
GUI_Thread.require {}
val snapshot = PIDE.session.snapshot()
@@ -210,14 +202,12 @@
GUI_Thread.later { handle_update(Some(changed.nodes)) }
}
- override def init(): Unit =
- {
+ override def init(): Unit = {
PIDE.session.commands_changed += main
handle_update()
}
- override def exit(): Unit =
- {
+ override def exit(): Unit = {
PIDE.session.commands_changed -= main
}
}
--- a/src/Tools/jEdit/src/token_markup.scala Fri Apr 01 11:51:42 2022 +0200
+++ b/src/Tools/jEdit/src/token_markup.scala Fri Apr 01 17:06:10 2022 +0200
@@ -20,15 +20,13 @@
import org.gjt.sp.jedit.buffer.JEditBuffer
-object Token_Markup
-{
+object Token_Markup {
/* line context */
def mode_rule_set(mode: String): ParserRuleSet =
new ParserRuleSet(mode, "MAIN")
- object Line_Context
- {
+ object Line_Context {
def init(mode: String): Line_Context =
new Line_Context(mode, Some(Scan.Finished), Line_Structure.init)
@@ -39,8 +37,7 @@
if (line == 0) init(JEdit_Lib.buffer_mode(buffer))
else after(buffer, line - 1)
- def after(buffer: JEditBuffer, line: Int): Line_Context =
- {
+ def after(buffer: JEditBuffer, line: Int): Line_Context = {
val line_mgr = JEdit_Lib.buffer_line_manager(buffer)
def context =
line_mgr.getLineContext(line) match {
@@ -56,11 +53,10 @@
}
class Line_Context(
- val mode: String,
- val context: Option[Scan.Line_Context],
- val structure: Line_Structure)
- extends TokenMarker.LineContext(mode_rule_set(mode), null)
- {
+ val mode: String,
+ val context: Option[Scan.Line_Context],
+ val structure: Line_Structure)
+ extends TokenMarker.LineContext(mode_rule_set(mode), null) {
def get_context: Scan.Line_Context = context.getOrElse(Scan.Finished)
override def hashCode: Int = (mode, context, structure).hashCode
@@ -75,9 +71,11 @@
/* tokens from line (inclusive) */
- private def try_line_tokens(syntax: Outer_Syntax, buffer: JEditBuffer, line: Int)
- : Option[List[Token]] =
- {
+ private def try_line_tokens(
+ syntax: Outer_Syntax,
+ buffer: JEditBuffer,
+ line: Int
+ ): Option[List[Token]] = {
val line_context = Line_Context.before(buffer, line)
for {
ctxt <- line_context.context
@@ -133,9 +131,11 @@
/* command spans */
- def command_span(syntax: Outer_Syntax, buffer: JEditBuffer, offset: Text.Offset)
- : Option[Text.Info[Command_Span.Span]] =
- {
+ def command_span(
+ syntax: Outer_Syntax,
+ buffer: JEditBuffer,
+ offset: Text.Offset
+ ): Option[Text.Info[Command_Span.Span]] = {
val keywords = syntax.keywords
def maybe_command_start(i: Text.Offset): Option[Text.Info[Token]] =
@@ -147,8 +147,7 @@
find(info => keywords.is_before_command(info.info) || info.info.is_command)
if (JEdit_Lib.buffer_range(buffer).contains(offset)) {
- val start_info =
- {
+ val start_info = {
val info1 = maybe_command_start(offset)
info1 match {
case Some(Text.Info(range1, tok1)) if tok1.is_command =>
@@ -167,8 +166,7 @@
case None => (false, 0, 0)
}
- val stop_info =
- {
+ val stop_info = {
val info1 = maybe_command_stop(start_next)
info1 match {
case Some(Text.Info(range1, tok1)) if tok1.is_command && start_before_command =>
@@ -193,21 +191,21 @@
}
private def _command_span_iterator(
- syntax: Outer_Syntax,
- buffer: JEditBuffer,
- offset: Text.Offset,
- next_offset: Text.Range => Text.Offset): Iterator[Text.Info[Command_Span.Span]] =
- new Iterator[Text.Info[Command_Span.Span]]
- {
+ syntax: Outer_Syntax,
+ buffer: JEditBuffer,
+ offset: Text.Offset,
+ next_offset: Text.Range => Text.Offset
+ ): Iterator[Text.Info[Command_Span.Span]] = {
+ new Iterator[Text.Info[Command_Span.Span]] {
private var next_span = command_span(syntax, buffer, offset)
def hasNext: Boolean = next_span.isDefined
- def next(): Text.Info[Command_Span.Span] =
- {
+ def next(): Text.Info[Command_Span.Span] = {
val span = next_span.getOrElse(Iterator.empty.next())
next_span = command_span(syntax, buffer, next_offset(span.range))
span
}
}
+ }
def command_span_iterator(syntax: Outer_Syntax, buffer: JEditBuffer, offset: Text.Offset)
: Iterator[Text.Info[Command_Span.Span]] =
@@ -223,8 +221,8 @@
class Marker(
protected val mode: String,
- protected val opt_buffer: Option[Buffer]) extends TokenMarker
- {
+ protected val opt_buffer: Option[Buffer]
+ ) extends TokenMarker {
addRuleSet(mode_rule_set(mode))
override def hashCode: Int = (mode, opt_buffer).hashCode
@@ -240,16 +238,17 @@
case Some(buffer) => "Marker(" + mode + "," + JEdit_Lib.buffer_name(buffer) + ")"
}
- override def markTokens(context: TokenMarker.LineContext,
- handler: TokenHandler, raw_line: Segment): TokenMarker.LineContext =
- {
+ override def markTokens(
+ context: TokenMarker.LineContext,
+ handler: TokenHandler,
+ raw_line: Segment
+ ): TokenMarker.LineContext = {
val line = if (raw_line == null) new Segment else raw_line
val line_context =
context match { case c: Line_Context => c case _ => Line_Context.init(mode) }
val structure = line_context.structure
- val context1 =
- {
+ val context1 = {
val opt_syntax =
opt_buffer match {
case Some(buffer) => Isabelle.buffer_syntax(buffer)
@@ -307,12 +306,10 @@
/* mode provider */
- class Mode_Provider(orig_provider: ModeProvider) extends ModeProvider
- {
+ class Mode_Provider(orig_provider: ModeProvider) extends ModeProvider {
for (mode <- orig_provider.getModes) addMode(mode)
- override def loadMode(mode: Mode, xmh: XModeHandler): Unit =
- {
+ override def loadMode(mode: Mode, xmh: XModeHandler): Unit = {
super.loadMode(mode, xmh)
Isabelle.mode_token_marker(mode.getName).foreach(mode.setTokenMarker)
Isabelle.indent_rule(mode.getName).foreach(indent_rule =>