diff --git a/src/Pure/Admin/afp.scala b/src/Pure/Admin/afp.scala --- a/src/Pure/Admin/afp.scala +++ b/src/Pure/Admin/afp.scala @@ -1,222 +1,226 @@ /* Title: Pure/Admin/afp.scala Author: Makarius Administrative support for the Archive of Formal Proofs. */ package isabelle import java.time.LocalDate import scala.collection.immutable.SortedMap object AFP { val repos_source = "https://isabelle.sketis.net/repos/afp-devel" 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)", "very_slow" -> "elapsed time of many hours (on high-end hardware)") val groups_bulky: List[String] = List("large", "slow") val chapter: String = "AFP" val force_partition1: List[String] = List("Category3", "HOL-ODE") def init(options: Options, base_dir: Path = Path.explode("$AFP_BASE")): AFP = new AFP(options, base_dir) /* entries */ 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]) { 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] = space_explode(',', get_string(prop)).map(_.trim).filter(_.nonEmpty) def title: String = get_string("title") def authors: List[String] = get_strings("author") def date: Date = parse_date(get("date").getOrElse(error("Missing date for entry " + quote(name)))) def topics: List[String] = get_strings("topic") def `abstract`: String = get_string("abstract").trim def maintainers: List[String] = get_strings("notify") def contributors: List[String] = get_strings("contributors") def license: String = get("license").getOrElse("BSD") def rdf_meta_data: Properties.T = RDF.meta_data( proper_string(title).map(Markup.META_TITLE -> _).toList ::: authors.map(Markup.META_CREATOR -> _) ::: contributors.map(Markup.META_CONTRIBUTOR -> _) ::: List(Markup.META_DATE -> RDF.date_format(date)) ::: List(Markup.META_LICENSE -> license) ::: proper_string(`abstract`).map(Markup.META_DESCRIPTION -> _).toList) } } 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") /* metadata */ private val entry_metadata: Map[String, Properties.T] = { val metadata_file = base_dir + Path.explode("metadata/metadata") var result = Map.empty[String, Properties.T] var section = "" var props = List.empty[Properties.Entry] val Section = """^\[(\S+)\]\s*$""".r val Property = """^(\S+)\s*=(.*)$""".r val Extra_Line = """^\s+(.*)$""".r val Blank_Line = """^\s*$""".r 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) { def err(msg: String): Nothing = error(msg + Position.here(Position.Line_File(i + 1, metadata_file.expand.implode))) line match { case Section(name) => flush(); section = name case Property(a, b) => if (section == "") err("Property without a section") props = (a -> b.trim) :: props case Extra_Line(line) => props match { case Nil => err("Extra line without a property") case (a, b) :: rest => props = (a, b + "\n" + line.trim) :: rest } case Blank_Line() => case _ => err("Bad input") } } flush() result } /* entries */ val entries_map: SortedMap[String, AFP.Entry] = { val entries = for (name <- Sessions.parse_roots(main_dir + Sessions.ROOTS)) yield { val metadata = entry_metadata.getOrElse(name, error("Entry without metadata: " + quote(name))) val sessions = Sessions.parse_root_entries(main_dir + Path.explode(name) + Sessions.ROOT).map(_.name) AFP.Entry(name, metadata, sessions) } val entries_map = - (SortedMap.empty[String, AFP.Entry] /: entries)({ case (m, e) => m + (e.name -> e) }) + entries.foldLeft(SortedMap.empty[String, AFP.Entry]) { case (m, e) => m + (e.name -> e) } val extra_metadata = (for ((name, _) <- entry_metadata.iterator if !entries_map.isDefinedAt(name)) yield name). toList.sorted if (extra_metadata.nonEmpty) error("Meta data without entry: " + commas_quote(extra_metadata)) entries_map } val entries: List[AFP.Entry] = entries_map.toList.map(_._2) /* sessions */ val sessions_map: SortedMap[String, AFP.Entry] = - (SortedMap.empty[String, AFP.Entry] /: entries)( - { case (m1, e) => (m1 /: e.sessions)({ case (m2, s) => m2 + (s -> e) }) }) + entries.foldLeft(SortedMap.empty[String, AFP.Entry]) { + case (m1, e) => e.sessions.foldLeft(m1) { case (m2, s) => m2 + (s -> e) } + } val sessions: List[String] = entries.flatMap(_.sessions) val sessions_structure: Sessions.Structure = Sessions.load_structure(options, dirs = List(main_dir)). selection(Sessions.Selection(sessions = sessions.toList)) /* dependency graph */ 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] = { val session_entries = - (Map.empty[String, String] /: entries) { - case (m1, e) => (m1 /: e.sessions) { case (m2, s) => m2 + (s -> e.name) } + entries.foldLeft(Map.empty[String, String]) { + case (m1, e) => e.sessions.foldLeft(m1) { case (m2, s) => m2 + (s -> e.name) } } - (Graph.empty[String, Unit] /: entries) { case (g, entry) => - val e1 = entry.name - (g.default_node(e1, ()) /: sessions_deps(entry)) { case (g1, s) => - (g1 /: session_entries.get(s).filterNot(_ == e1)) { case (g2, e2) => - try { g2.default_node(e2, ()).add_edge_acyclic(e2, e1) } - catch { - case exn: Graph.Cycles[_] => - error(cat_lines(exn.cycles.map(cycle => - "Cyclic dependency of " + cycle.map(c => quote(c.toString)).mkString(" via ") + - " due to session " + quote(s)))) - } + entries.foldLeft(Graph.empty[String, Unit]) { + case (g, entry) => + val e1 = entry.name + sessions_deps(entry).foldLeft(g.default_node(e1, ())) { + case (g1, s) => + session_entries.get(s).filterNot(_ == e1).foldLeft(g1) { + case (g2, e2) => + try { g2.default_node(e2, ()).add_edge_acyclic(e2, e1) } + catch { + case exn: Graph.Cycles[_] => + error(cat_lines(exn.cycles.map(cycle => + "Cyclic dependency of " + cycle.map(c => quote(c.toString)).mkString(" via ") + + " due to session " + quote(s)))) + } + } } - } } } def entries_graph_display: Graph_Display.Graph = Graph_Display.make_graph(entries_graph) def entries_json_text: String = (for (entry <- entries.iterator) yield { val distrib_deps = sessions_deps(entry).filterNot(sessions.contains(_)) val afp_deps = entries_graph.imm_preds(entry.name).toList """ {""" + JSON.Format(entry.name) + """: {"distrib_deps": """ + JSON.Format(distrib_deps) + """, "afp_deps": """ + JSON.Format(afp_deps) + """ } }""" }).mkString("[", ", ", "\n]\n") /* partition sessions */ def partition(n: Int): List[String] = n match { case 0 => Nil case 1 | 2 => val graph = sessions_structure.build_graph.restrict(sessions.toSet) val force_part1 = graph.all_preds(graph.all_succs(AFP.force_partition1.filter(graph.defined))).toSet val (part1, part2) = graph.keys.partition(a => force_part1(a) || graph.is_isolated(a)) if (n == 1) part1 else part2 case _ => error("Bad AFP partition: " + n + " (should be 0, 1, 2)") } } diff --git a/src/Pure/Admin/build_csdp.scala b/src/Pure/Admin/build_csdp.scala --- a/src/Pure/Admin/build_csdp.scala +++ b/src/Pure/Admin/build_csdp.scala @@ -1,207 +1,207 @@ /* Title: Pure/Admin/build_csdp.scala Author: Makarius Build Isabelle CSDP component from official download. */ package isabelle 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 = "") { val changed: List[(String, String)] = List("CFLAGS" -> CFLAGS, "LIBS" -> LIBS).filter(p => p._2.nonEmpty) def print: Option[String] = if (changed.isEmpty) None else Some(" * " + platform + ":\n" + changed.map(p => " " + p._1 + "=" + p._2) .mkString("\n")) def change(path: Path): Unit = { def change_line(line: String, entry: (String, String)): String = line.replaceAll(entry._1 + "=.*", entry._1 + "=" + entry._2) File.change(path, s => - split_lines(s).map(line => (line /: changed)(change_line)).mkString("\n")) + split_lines(s).map(line => changed.foldLeft(line)(change_line)).mkString("\n")) } } val build_flags: List[Flags] = List( Flags("arm64-linux", CFLAGS = "-O3 -ansi -Wall -DNOSHORTS -DBIT64 -DUSESIGTERM -DUSEGETTIME -I../include", LIBS = "-static -L../lib -lsdp -llapack -lblas -lgfortran -lm"), Flags("x86_64-linux", CFLAGS = "-O3 -ansi -Wall -DNOSHORTS -DBIT64 -DUSESIGTERM -DUSEGETTIME -I../include", LIBS = "-static -L../lib -lsdp -llapack -lblas -lgfortran -lquadmath -lm"), Flags("x86_64-darwin", CFLAGS = "-O3 -Wall -DNOSHORTS -DBIT64 -DUSESIGTERM -DUSEGETTIME -I../include", LIBS = "-L../lib -lsdp -llapack -lblas -lm"), Flags("x86_64-windows")) /* build CSDP */ def build_csdp( download_url: String = default_download_url, verbose: Boolean = false, progress: Progress = new Progress, target_dir: Path = Path.current, mingw: MinGW = MinGW.none): Unit = { mingw.check Isabelle_System.with_tmp_dir("build")(tmp_dir => { /* component */ val Archive_Name = """^.*?([^/]+)$""".r val Version = """^[^0-9]*([0-9].*)\.tar.gz$""".r val archive_name = download_url match { case Archive_Name(name) => name case _ => error("Failed to determine source archive name from " + quote(download_url)) } val version = archive_name match { case Version(version) => version case _ => error("Failed to determine component version from " + quote(archive_name)) } val component_name = "csdp-" + version val component_dir = Isabelle_System.new_directory(target_dir + Path.basic(component_name)) progress.echo("Component " + component_dir) /* platform */ val platform_name = proper_string(Isabelle_System.getenv("ISABELLE_WINDOWS_PLATFORM64")) orElse proper_string(Isabelle_System.getenv("ISABELLE_PLATFORM64")) getOrElse error("No 64bit platform") val platform_dir = Isabelle_System.make_directory(component_dir + Path.basic(platform_name)) /* download source */ val archive_path = tmp_dir + Path.basic(archive_name) Isabelle_System.download(download_url, archive_path, progress = progress) Isabelle_System.bash("tar xzf " + File.bash_path(archive_path), cwd = tmp_dir.file).check val source_name = File.get_dir(tmp_dir) Isabelle_System.bash( "tar xzf " + archive_path + " && mv " + Bash.string(source_name) + " src", cwd = component_dir.file).check /* build */ progress.echo("Building CSDP for " + platform_name + " ...") val build_dir = tmp_dir + Path.basic(source_name) build_flags.find(flags => flags.platform == platform_name) match { case None => error("No build flags for platform " + quote(platform_name)) case Some(flags) => File.find_files(build_dir.file, pred = file => file.getName == "Makefile"). foreach(file => flags.change(File.path(file))) } progress.bash(mingw.bash_script("make"), cwd = build_dir.file, echo = verbose).check /* install */ Isabelle_System.copy_file(build_dir + Path.explode("LICENSE"), component_dir) Isabelle_System.copy_file(build_dir + Path.explode("solver/csdp").platform_exe, platform_dir) if (Platform.is_windows) { Executable.libraries_closure(platform_dir + Path.explode("csdp.exe"), mingw = mingw, filter = Set("libblas", "liblapack", "libgfortran", "libgcc_s_seh", "libquadmath", "libwinpthread")) } /* settings */ val etc_dir = Isabelle_System.make_directory(component_dir + Path.basic("etc")) File.write(etc_dir + Path.basic("settings"), """# -*- shell-script -*- :mode=shellscript: ISABELLE_CSDP="$COMPONENT/${ISABELLE_WINDOWS_PLATFORM64:-$ISABELLE_PLATFORM64}/csdp" """) /* README */ File.write(component_dir + Path.basic("README"), """This is CSDP """ + version + """ from """ + download_url + """ Makefile flags have been changed for various platforms as follows: """ + build_flags.flatMap(_.print).mkString("\n\n") + """ The distribution has been built like this: cd src && make Only the bare "solver/csdp" program is used for Isabelle. Makarius """ + Date.Format.date(Date.now()) + "\n") }) } /* Isabelle tool wrapper */ val isabelle_tool = Isabelle_Tool("build_csdp", "build prover component from official download", Scala_Project.here, args => { var target_dir = Path.current var mingw = MinGW.none var download_url = default_download_url var verbose = false val getopts = Getopts(""" Usage: isabelle build_csdp [OPTIONS] Options are: -D DIR target directory (default ".") -M DIR msys/mingw root specification for Windows -U URL download URL (default: """" + default_download_url + """") -v verbose Build prover component from official download. """, "D:" -> (arg => target_dir = Path.explode(arg)), "M:" -> (arg => mingw = MinGW(Path.explode(arg))), "U:" -> (arg => download_url = arg), "v" -> (_ => verbose = true)) val more_args = getopts(args) if (more_args.nonEmpty) getopts.usage() val progress = new Console_Progress() build_csdp(download_url = download_url, verbose = verbose, progress = progress, target_dir = target_dir, mingw = mingw) }) } diff --git a/src/Pure/Admin/build_history.scala b/src/Pure/Admin/build_history.scala --- a/src/Pure/Admin/build_history.scala +++ b/src/Pure/Admin/build_history.scala @@ -1,618 +1,618 @@ /* Title: Pure/Admin/build_history.scala Author: Makarius Build other history versions. */ package isabelle object Build_History { /* log files */ val engine = "build_history" val log_prefix = engine + "_" /* augment settings */ def augment_settings( other_isabelle: Other_Isabelle, threads: Int, arch_64: Boolean, heap: Int, max_heap: Option[Int], more_settings: List[String]): String = { val (ml_platform, ml_settings) = { val windows_32 = "x86-windows" val windows_64 = "x86_64-windows" val windows_64_32 = "x86_64_32-windows" val platform_32 = other_isabelle.getenv("ISABELLE_PLATFORM32") val platform_64 = other_isabelle.getenv("ISABELLE_PLATFORM64") val platform_64_32 = platform_64.replace("x86_64-", "x86_64_32-") val polyml_home = try { Path.explode(other_isabelle.getenv("ML_HOME")).dir } catch { case ERROR(msg) => error("Bad ML_HOME: " + msg) } def ml_home(platform: String): Path = polyml_home + Path.explode(platform) def err(platform: String): Nothing = error("Platform " + platform + " unavailable on this machine") def check_dir(platform: String): Boolean = platform != "" && ml_home(platform).is_dir val ml_platform = if (Platform.is_windows && arch_64) { if (check_dir(windows_64)) windows_64 else err(windows_64) } else if (Platform.is_windows && !arch_64) { if (check_dir(windows_64_32)) windows_64_32 else if (check_dir(windows_32)) windows_32 else platform_32 // x86-cygwin } else if (arch_64) { if (check_dir(platform_64)) platform_64 else err(platform_64) } else if (check_dir(platform_64_32)) platform_64_32 else platform_32 val ml_options = "--minheap " + heap + (if (max_heap.isDefined) " --maxheap " + max_heap.get else "") + " --gcthreads " + threads + (if (ml_platform.endsWith("-windows")) " --codepage utf8" else "") (ml_platform, List( "ML_HOME=" + File.bash_path(ml_home(ml_platform)), "ML_PLATFORM=" + quote(ml_platform), "ML_OPTIONS=" + quote(ml_options))) } val thread_settings = List( "ISABELLE_JAVA_SYSTEM_OPTIONS=\"$ISABELLE_JAVA_SYSTEM_OPTIONS -Disabelle.threads=" + threads + "\"", "ISABELLE_BUILD_OPTIONS=\"threads=" + threads + "\"") val settings = List(ml_settings, thread_settings) ::: (if (more_settings.isEmpty) Nil else List(more_settings)) File.append(other_isabelle.etc_settings, "\n" + cat_lines(settings.map(terminate_lines(_)))) ml_platform } /** build_history **/ private val default_user_home = Path.explode("$USER_HOME") private val default_rev = "tip" private val default_multicore = (1, 1) private val default_heap = 1500 private val default_isabelle_identifier = "build_history" def build_history( options: Options, root: Path, user_home: Path = default_user_home, progress: Progress = new Progress, rev: String = default_rev, afp_rev: Option[String] = None, afp_partition: Int = 0, isabelle_identifier: String = default_isabelle_identifier, ml_statistics_step: Int = 1, component_repository: String = Components.default_component_repository, components_base: Path = Components.default_components_base, fresh: Boolean = false, hostname: String = "", multicore_base: Boolean = false, multicore_list: List[(Int, Int)] = List(default_multicore), arch_64: Boolean = false, heap: Int = default_heap, max_heap: Option[Int] = None, init_settings: List[String] = Nil, more_settings: List[String] = Nil, more_preferences: List[String] = Nil, verbose: Boolean = false, build_tags: List[String] = Nil, build_args: List[String] = Nil): List[(Process_Result, Path)] = { /* sanity checks */ if (File.eq(Path.explode("~~"), root)) error("Repository coincides with ISABELLE_HOME=" + Path.explode("~~").expand) for ((threads, _) <- multicore_list if threads < 1) error("Bad threads value < 1: " + threads) for ((_, processes) <- multicore_list if processes < 1) error("Bad processes value < 1: " + processes) if (heap < 100) error("Bad heap value < 100: " + heap) if (max_heap.isDefined && max_heap.get < heap) error("Bad max_heap value < heap: " + max_heap.get) System.getenv("ISABELLE_SETTINGS_PRESENT") match { case null | "" => case _ => error("Cannot run build_history within existing Isabelle settings environment") } /* checkout Isabelle + AFP repository */ 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")) hg.id(rev = version) } val isabelle_version = checkout(root, rev) val afp_repos = root + Path.explode("AFP") val afp_version = afp_rev.map(checkout(afp_repos, _)) val (afp_build_args, afp_sessions) = if (afp_rev.isEmpty) (Nil, Nil) else { val (opt, sessions) = try { val afp = AFP.init(options, afp_repos) ("-d", afp.partition(afp_partition)) } catch { case ERROR(_) => ("-D", Nil) } (List(opt, "~~/AFP/thys"), sessions) } /* main */ val other_isabelle = Other_Isabelle(root, isabelle_identifier = isabelle_identifier, user_home = user_home, progress = progress) val build_host = proper_string(hostname) getOrElse Isabelle_System.hostname() val build_history_date = Date.now() val build_group_id = build_host + ":" + build_history_date.time.ms var first_build = true for ((threads, processes) <- multicore_list) yield { /* init settings */ val component_settings = other_isabelle.init_components( component_repository = component_repository, components_base = components_base, catalogs = List("main", "optional")) other_isabelle.init_settings(component_settings ::: init_settings) other_isabelle.resolve_components(echo = verbose) val ml_platform = augment_settings(other_isabelle, threads, arch_64, heap, max_heap, more_settings) File.write(other_isabelle.etc_preferences, cat_lines(more_preferences)) val isabelle_output = other_isabelle.isabelle_home_user + Path.explode("heaps") + Path.explode(other_isabelle.getenv("ML_IDENTIFIER")) val isabelle_output_log = isabelle_output + Path.explode("log") val isabelle_base_log = isabelle_output + Path.explode("../base_log") if (first_build) { other_isabelle.resolve_components(echo = verbose) if (fresh) Isabelle_System.rm_tree(other_isabelle.isabelle_home + Path.explode("lib/classes")) other_isabelle.bash( "env PATH=\"" + File.bash_path(Path.explode("~~/lib/dummy_stty").expand) + ":$PATH\" " + "bin/isabelle jedit -b", redirect = true, echo = verbose).check for { tool <- List("ghc_setup", "ocaml_setup") if other_isabelle.getenv("ISABELLE_" + Word.uppercase(tool)) == "true" && (other_isabelle.isabelle_home + Path.explode("lib/Tools/" + tool)).is_file } other_isabelle(tool, echo = verbose) Isabelle_System.rm_tree(isabelle_base_log) } Isabelle_System.rm_tree(isabelle_output) Isabelle_System.make_directory(isabelle_output) val log_path = other_isabelle.isabelle_home_user + Build_Log.log_subdir(build_history_date) + Build_Log.log_filename(Build_History.engine, build_history_date, List(build_host, ml_platform, "M" + threads) ::: build_tags) Isabelle_System.make_directory(log_path.dir) val build_out = other_isabelle.isabelle_home_user + Path.explode("log/build.out") val build_out_progress = new File_Progress(build_out) build_out.file.delete /* build */ if (multicore_base && !first_build && isabelle_base_log.is_dir) Isabelle_System.copy_dir(isabelle_base_log, isabelle_output_log) val build_start = Date.now() val build_args1 = List("-v", "-j" + processes) ::: afp_build_args ::: build_args val build_isabelle = Other_Isabelle(root, isabelle_identifier = isabelle_identifier, user_home = user_home, progress = build_out_progress) val build_result = build_isabelle("build " + Bash.strings(build_args1 ::: afp_sessions), redirect = true, echo = true, strict = false) val build_end = Date.now() val build_info: Build_Log.Build_Info = Build_Log.Log_File(log_path.file_name, build_result.out_lines). parse_build_info(ml_statistics = true) /* output log */ val store = Sessions.store(options + "build_database_server=false") val meta_info = Properties.lines_nonempty(Build_Log.Prop.build_tags.name, build_tags) ::: Properties.lines_nonempty(Build_Log.Prop.build_args.name, build_args1) ::: List( Build_Log.Prop.build_group_id.name -> build_group_id, Build_Log.Prop.build_id.name -> (build_host + ":" + build_start.time.ms), Build_Log.Prop.build_engine.name -> Build_History.engine, Build_Log.Prop.build_host.name -> build_host, Build_Log.Prop.build_start.name -> Build_Log.print_date(build_start), Build_Log.Prop.build_end.name -> Build_Log.print_date(build_end), Build_Log.Prop.isabelle_version.name -> isabelle_version) ::: afp_version.map(Build_Log.Prop.afp_version.name -> _).toList build_out_progress.echo("Reading session build info ...") val session_build_info = build_info.finished_sessions.flatMap(session_name => { val database = isabelle_output + store.database(session_name) if (database.is_file) { using(SQLite.open_database(database))(db => { val theory_timings = try { store.read_theory_timings(db, session_name).map(ps => Protocol.Theory_Timing_Marker((Build_Log.SESSION_NAME, session_name) :: ps)) } catch { case ERROR(_) => Nil } val session_sources = store.read_build(db, session_name).map(_.sources) match { case Some(sources) if sources.length == SHA1.digest_length => List("Sources " + session_name + " " + sources) case _ => Nil } theory_timings ::: session_sources }) } else Nil }) build_out_progress.echo("Reading ML statistics ...") val ml_statistics = build_info.finished_sessions.flatMap(session_name => { val database = isabelle_output + store.database(session_name) val log_gz = isabelle_output + store.log_gz(session_name) val properties = if (database.is_file) { using(SQLite.open_database(database))(db => store.read_ml_statistics(db, session_name)) } else if (log_gz.is_file) { Build_Log.Log_File(log_gz).parse_session_info(ml_statistics = true).ml_statistics } else Nil val trimmed_properties = if (ml_statistics_step <= 0) Nil else if (ml_statistics_step == 1) properties else { (for { (ps, i) <- properties.iterator.zipWithIndex if i % ml_statistics_step == 0 } yield ps).toList } trimmed_properties.map(ps => (Build_Log.SESSION_NAME -> session_name) :: ps) }) build_out_progress.echo("Reading error messages ...") val session_errors = build_info.failed_sessions.flatMap(session_name => { val database = isabelle_output + store.database(session_name) val errors = if (database.is_file) { try { using(SQLite.open_database(database))(db => store.read_errors(db, session_name)) } // column "errors" could be missing catch { case _: java.sql.SQLException => Nil } } else Nil errors.map(msg => List(Build_Log.SESSION_NAME -> session_name, Markup.CONTENT -> msg)) }) build_out_progress.echo("Reading heap sizes ...") val heap_sizes = build_info.finished_sessions.flatMap(session_name => { val heap = isabelle_output + Path.explode(session_name) if (heap.is_file) Some("Heap " + session_name + " (" + Value.Long(heap.file.length) + " bytes)") else None }) build_out_progress.echo("Writing log file " + log_path.xz + " ...") File.write_xz(log_path.xz, terminate_lines( Protocol.Meta_Info_Marker(meta_info) :: build_result.out_lines ::: session_build_info ::: ml_statistics.map(Protocol.ML_Statistics_Marker.apply) ::: session_errors.map(Protocol.Error_Message_Marker.apply) ::: heap_sizes), XZ.options(6)) /* next build */ if (multicore_base && first_build && isabelle_output_log.is_dir) Isabelle_System.copy_dir(isabelle_output_log, isabelle_base_log) Isabelle_System.rm_tree(isabelle_output) first_build = false (build_result, log_path.xz) } } /* command line entry point */ private object Multicore { private val Pat1 = """^(\d+)$""".r private val Pat2 = """^(\d+)x(\d+)$""".r def parse(s: String): (Int, Int) = s match { case Pat1(Value.Int(x)) => (x, 1) case Pat2(Value.Int(x), Value.Int(y)) => (x, y) case _ => error("Bad multicore configuration: " + quote(s)) } } def main(args: Array[String]): Unit = { Command_Line.tool { var afp_rev: Option[String] = None var multicore_base = false var components_base: Path = Components.default_components_base var heap: Option[Int] = None var max_heap: Option[Int] = None var multicore_list = List(default_multicore) var isabelle_identifier = default_isabelle_identifier var afp_partition = 0 var component_repository = Components.default_component_repository var more_settings: List[String] = Nil var more_preferences: List[String] = Nil var fresh = false var hostname = "" var init_settings: List[String] = Nil var arch_64 = false var output_file = "" var rev = default_rev var ml_statistics_step = 1 var build_tags = List.empty[String] var user_home = default_user_home var verbose = false var exit_code = false val getopts = Getopts(""" Usage: Admin/build_history [OPTIONS] REPOSITORY [ARGS ...] Options are: -A REV include $ISABELLE_HOME/AFP repository at given revision -B first multicore build serves as base for scheduling information -C DIR base directory for Isabelle components (default: """ + Components.default_components_base + """) -H SIZE minimal ML heap in MB (default: """ + default_heap + """ for x86, """ + default_heap * 2 + """ for x86_64) -M MULTICORE multicore configurations (see below) -N NAME alternative ISABELLE_IDENTIFIER (default: """ + default_isabelle_identifier + """) -P NUMBER AFP partition number (0, 1, 2, default: 0=unrestricted) -R URL remote repository for Isabelle components (default: """ + Components.default_component_repository + """) -U SIZE maximal ML heap in MB (default: unbounded) -e TEXT additional text for generated etc/settings -f fresh build of Isabelle/Scala components (recommended) -h NAME override local hostname -i TEXT initial text for generated etc/settings -m ARCH processor architecture (32=x86, 64=x86_64, default: x86) -o FILE output file for log names (default: stdout) -p TEXT additional text for generated etc/preferences -r REV update to revision (default: """ + default_rev + """) -s NUMBER step size for ML statistics (0=none, 1=all, n=step, default: 1) -t TAG free-form build tag (multiple occurrences possible) -u DIR alternative USER_HOME directory -v verbose -x return overall exit code from build processes Build Isabelle sessions from the history of another REPOSITORY clone, passing ARGS directly to its isabelle build tool. Each MULTICORE configuration consists of one or two numbers (default 1): THREADS or THREADSxPROCESSES, e.g. -M 1,2,4 or -M 1x4,2x2,4. """, "A:" -> (arg => afp_rev = Some(arg)), "B" -> (_ => multicore_base = true), "C:" -> (arg => components_base = Path.explode(arg)), "H:" -> (arg => heap = Some(Value.Int.parse(arg))), "M:" -> (arg => multicore_list = space_explode(',', arg).map(Multicore.parse)), "N:" -> (arg => isabelle_identifier = arg), "P:" -> (arg => afp_partition = Value.Int.parse(arg)), "R:" -> (arg => component_repository = arg), "U:" -> (arg => max_heap = Some(Value.Int.parse(arg))), "e:" -> (arg => more_settings = more_settings ::: List(arg)), "f" -> (_ => fresh = true), "h:" -> (arg => hostname = arg), "i:" -> (arg => init_settings = init_settings ::: List(arg)), "m:" -> { case "32" | "x86" => arch_64 = false case "64" | "x86_64" => arch_64 = true case bad => error("Bad processor architecture: " + quote(bad)) }, "o:" -> (arg => output_file = arg), "p:" -> (arg => more_preferences = more_preferences ::: List(arg)), "r:" -> (arg => rev = arg), "s:" -> (arg => ml_statistics_step = Value.Int.parse(arg)), "t:" -> (arg => build_tags = build_tags ::: List(arg)), "u:" -> (arg => user_home = Path.explode(arg)), "v" -> (_ => verbose = true), "x" -> (_ => exit_code = true)) val more_args = getopts(args) val (root, build_args) = more_args match { case root :: build_args => (Path.explode(root), build_args) case _ => getopts.usage() } val progress = new Console_Progress(stderr = true) val results = build_history(Options.init(), root, user_home = user_home, progress = progress, rev = rev, afp_rev = afp_rev, afp_partition = afp_partition, isabelle_identifier = isabelle_identifier, ml_statistics_step = ml_statistics_step, component_repository = component_repository, components_base = components_base, fresh = fresh, hostname = hostname, multicore_base = multicore_base, multicore_list = multicore_list, arch_64 = arch_64, heap = heap.getOrElse(if (arch_64) default_heap * 2 else default_heap), max_heap = max_heap, init_settings = init_settings, more_settings = more_settings, more_preferences = more_preferences, verbose = verbose, build_tags = build_tags, build_args = build_args) if (output_file == "") { for ((_, log_path) <- results) Output.writeln(log_path.implode, stdout = true) } else { File.write(Path.explode(output_file), cat_lines(for ((_, log_path) <- results) yield log_path.implode)) } - val rc = (0 /: results) { case (rc, (res, _)) => rc max res.rc } + val rc = results.foldLeft(0) { case (rc, (res, _)) => rc max res.rc } if (rc != 0 && exit_code) sys.exit(rc) } } /** remote build_history -- via command-line **/ def remote_build_history( ssh: SSH.Session, isabelle_repos_self: Path, isabelle_repos_other: Path, isabelle_repos_source: String = Isabelle_Cronjob.isabelle_repos_source, afp_repos_source: String = AFP.repos_source, isabelle_identifier: String = "remote_build_history", self_update: Boolean = false, progress: Progress = new Progress, rev: String = "", afp_rev: Option[String] = None, options: String = "", args: String = ""): List[(String, Bytes)] = { /* Isabelle self repository */ val self_hg = Mercurial.setup_repository(isabelle_repos_source, isabelle_repos_self, ssh = ssh) def execute(cmd: String, args: String, echo: Boolean = false, strict: Boolean = true): Unit = ssh.execute( Isabelle_System.export_isabelle_identifier(isabelle_identifier) + ssh.bash_path(isabelle_repos_self + Path.explode(cmd)) + " " + args, progress_stdout = progress.echo_if(echo, _), progress_stderr = progress.echo_if(echo, _), strict = strict).check if (self_update) { val hg = Mercurial.repository(Path.explode("~~")) hg.push(self_hg.root_url, force = true) self_hg.update(rev = hg.parent(), clean = true) execute("bin/isabelle", "components -I") execute("bin/isabelle", "components -a", echo = true) execute("Admin/build", "jars_fresh") } val rev_id = self_hg.id(rev) /* Isabelle other + AFP repository */ if (Mercurial.is_repository(isabelle_repos_other, ssh = ssh)) { ssh.rm_tree(isabelle_repos_other) } Mercurial.clone_repository( ssh.bash_path(isabelle_repos_self), isabelle_repos_other, rev = rev_id, ssh = ssh) val afp_options = if (afp_rev.isEmpty) "" else { val afp_repos = isabelle_repos_other + Path.explode("AFP") Mercurial.setup_repository(afp_repos_source, afp_repos, ssh = ssh) " -A " + Bash.string(afp_rev.get) } /* Admin/build_history */ 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) try { execute("Admin/build_history", "-o " + ssh.bash_path(output_file) + rev_options + afp_options + " " + options + " " + ssh.bash_path(isabelle_repos_other) + " " + args, echo = true, strict = false) } catch { case ERROR(msg) => cat_error(msg, "The error(s) above occurred for build_bistory " + rev_options + afp_options) } for (line <- split_lines(ssh.read(output_file))) yield { val log = Path.explode(line) val bytes = ssh.read_bytes(log) ssh.rm(log) (log.file_name, bytes) } }) } } diff --git a/src/Pure/Admin/build_status.scala b/src/Pure/Admin/build_status.scala --- a/src/Pure/Admin/build_status.scala +++ b/src/Pure/Admin/build_status.scala @@ -1,621 +1,623 @@ /* Title: Pure/Admin/build_status.scala Author: Makarius Present recent build status information from database. */ package isabelle object Build_Status { /* defaults */ val default_target_dir = Path.explode("build_status") val default_image_size = (800, 600) val default_history = 30 def default_profiles: List[Profile] = Jenkins.build_status_profiles ::: Isabelle_Cronjob.build_status_profiles /* data profiles */ sealed case class Profile( 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 = { Build_Log.Data.universal_table.select(columns, distinct = true, sql = "WHERE " + Build_Log.Data.pull_date(afp) + " > " + Build_Log.Data.recent_time(days(options)) + " AND " + SQL.member(Build_Log.Data.status.ident, List( Build_Log.Session_Status.finished.toString, Build_Log.Session_Status.failed.toString)) + (if (only_sessions.isEmpty) "" else " AND " + SQL.member(Build_Log.Data.session_name.ident, only_sessions)) + " AND " + SQL.enclose(sql)) } } /* build status */ def build_status(options: Options, progress: Progress = new Progress, profiles: List[Profile] = default_profiles, only_sessions: Set[String] = Set.empty, verbose: Boolean = false, target_dir: Path = default_target_dir, ml_statistics: Boolean = false, 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 val data = read_data(options, progress = progress, profiles = profiles, only_sessions = only_sessions, verbose = verbose, ml_statistics = ml_statistics, ml_statistics_domain = ml_statistics_domain) present_data(data, progress = progress, target_dir = target_dir, image_size = image_size) } /* read data */ 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]) { 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) { require(entries.nonEmpty, "no entries") lazy val sorted_entries: List[Entry] = entries.sortBy(entry => - entry.date) def head: Entry = sorted_entries.head def order: Long = - head.timing.elapsed.ms def finished_entries: List[Entry] = sorted_entries.filter(_.finished) def finished_entries_size: Int = finished_entries.map(_.date).toSet.size def check_timing: Boolean = finished_entries_size >= 3 def check_heap: Boolean = finished_entries_size >= 3 && finished_entries.forall(entry => entry.maximum_heap > 0 || entry.average_heap > 0 || entry.stored_heap > 0) def make_csv: CSV.File = { val header = List("session_name", "chapter", "pull_date", "afp_pull_date", "isabelle_version", "afp_version", "timing_elapsed", "timing_cpu", "timing_gc", "ml_timing_elapsed", "ml_timing_cpu", "ml_timing_gc", "maximum_code", "average_code", "maximum_stack", "average_stack", "maximum_heap", "average_heap", "stored_heap", "status") val date_format = Date.Format("uuuu-MM-dd HH:mm:ss") val records = for (entry <- sorted_entries) yield { CSV.Record(name, entry.chapter, date_format(entry.pull_date), entry.afp_pull_date match { case Some(date) => date_format(date) case None => "" }, entry.isabelle_version, entry.afp_version, entry.timing.elapsed.ms, entry.timing.cpu.ms, entry.timing.gc.ms, entry.ml_timing.elapsed.ms, entry.ml_timing.cpu.ms, entry.ml_timing.gc.ms, entry.maximum_code, entry.average_code, entry.maximum_stack, entry.average_stack, entry.maximum_heap, entry.average_heap, entry.stored_heap, entry.status) } CSV.File(name, header, records) } } sealed case class Entry( chapter: String, pull_date: Date, afp_pull_date: Option[Date], isabelle_version: String, afp_version: String, timing: Timing, ml_timing: Timing, maximum_code: Long, average_code: Long, maximum_stack: Long, average_stack: Long, maximum_heap: Long, average_heap: Long, stored_heap: Long, status: Build_Log.Session_Status.Value, 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 = { if (errors.isEmpty) HTML.text(name + print_version(isabelle_version, afp_version, chapter)) else { HTML.tooltip_errors(HTML.text(name), errors.map(s => HTML.text(Symbol.decode(s)))) :: HTML.text(print_version(isabelle_version, afp_version, chapter)) } } } 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 = { val body = proper_string(isabelle_version).map("Isabelle/" + _).toList ::: (if (chapter == AFP.chapter) proper_string(afp_version).map("AFP/" + _) else None).toList if (body.isEmpty) "" else body.mkString(" (", ", ", ")") } def read_data(options: Options, progress: Progress = new Progress, profiles: List[Profile] = default_profiles, only_sessions: Set[String] = Set.empty, ml_statistics: Boolean = false, ml_statistics_domain: String => Boolean = (key: String) => true, verbose: Boolean = false): Data = { val date = Date.now() var data_hosts = Map.empty[String, Set[String]] var data_stretch = Map.empty[String, Double] var data_entries = Map.empty[String, Map[String, Session]] def get_hosts(data_name: String): Set[String] = data_hosts.getOrElse(data_name, Set.empty) val store = Build_Log.store(options) using(store.open_database())(db => { for (profile <- profiles.sortBy(_.description)) { progress.echo("input " + quote(profile.description)) val afp = profile.afp val columns = List( Build_Log.Data.pull_date(afp = false), Build_Log.Data.pull_date(afp = true), Build_Log.Prop.build_host, Build_Log.Prop.isabelle_version, Build_Log.Prop.afp_version, Build_Log.Settings.ISABELLE_BUILD_OPTIONS, Build_Log.Settings.ML_PLATFORM, Build_Log.Data.session_name, Build_Log.Data.chapter, Build_Log.Data.groups, Build_Log.Data.threads, Build_Log.Data.timing_elapsed, Build_Log.Data.timing_cpu, Build_Log.Data.timing_gc, Build_Log.Data.ml_timing_elapsed, Build_Log.Data.ml_timing_cpu, Build_Log.Data.ml_timing_gc, Build_Log.Data.heap_size, Build_Log.Data.status, Build_Log.Data.errors) ::: (if (ml_statistics) List(Build_Log.Data.ml_statistics) else Nil) val Threads_Option = """threads\s*=\s*(\d+)""".r val sql = profile.select(options, columns, only_sessions) progress.echo_if(verbose, sql) 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 threads1 = res.string(Build_Log.Settings.ISABELLE_BUILD_OPTIONS) match { case Threads_Option(Value.Int(i)) => i case _ => 1 } val threads2 = res.get_int(Build_Log.Data.threads).getOrElse(1) threads1 max threads2 } val ml_platform = res.string(Build_Log.Settings.ML_PLATFORM) val data_name = profile.description + (if (ml_platform.startsWith("x86_64-")) ", 64bit" else "") + (if (threads == 1) "" else ", " + threads + " threads") res.get_string(Build_Log.Prop.build_host).foreach(host => data_hosts += (data_name -> (get_hosts(data_name) + host))) data_stretch += (data_name -> profile.stretch(options)) val isabelle_version = res.string(Build_Log.Prop.isabelle_version) val afp_version = res.string(Build_Log.Prop.afp_version) val ml_stats = ML_Statistics( if (ml_statistics) { Properties.uncompress(res.bytes(Build_Log.Data.ml_statistics), cache = store.cache) } else Nil, domain = ml_statistics_domain, heading = session_name + print_version(isabelle_version, afp_version, chapter)) val entry = Entry( chapter = chapter, pull_date = res.date(Build_Log.Data.pull_date(afp = false)), afp_pull_date = if (afp) res.get_date(Build_Log.Data.pull_date(afp = true)) else None, isabelle_version = isabelle_version, afp_version = afp_version, timing = res.timing( Build_Log.Data.timing_elapsed, Build_Log.Data.timing_cpu, Build_Log.Data.timing_gc), ml_timing = res.timing( Build_Log.Data.ml_timing_elapsed, Build_Log.Data.ml_timing_cpu, Build_Log.Data.ml_timing_gc), maximum_code = ml_stats.maximum(ML_Statistics.CODE_SIZE).toLong, average_code = ml_stats.average(ML_Statistics.CODE_SIZE).toLong, maximum_stack = ml_stats.maximum(ML_Statistics.STACK_SIZE).toLong, average_stack = ml_stats.average(ML_Statistics.STACK_SIZE).toLong, maximum_heap = ml_stats.maximum(ML_Statistics.HEAP_SIZE).toLong, average_heap = ml_stats.average(ML_Statistics.HEAP_SIZE).toLong, stored_heap = ML_Statistics.mem_scale(res.long(Build_Log.Data.heap_size)), status = Build_Log.Session_Status.withName(res.string(Build_Log.Data.status)), errors = Build_Log.uncompress_errors( res.bytes(Build_Log.Data.errors), cache = store.cache)) val sessions = data_entries.getOrElse(data_name, Map.empty) val session = sessions.get(session_name) match { case None => Session(session_name, threads, List(entry), ml_stats, entry.date) case Some(old) => val (ml_stats1, ml_stats1_date) = if (entry.date > old.ml_statistics_date) (ml_stats, entry.date) else (old.ml_statistics, old.ml_statistics_date) Session(session_name, threads, entry :: old.entries, ml_stats1, ml_stats1_date) } if ((!afp || chapter == AFP.chapter) && (!profile.bulky || groups.exists(AFP.groups_bulky.toSet))) { data_entries += (data_name -> (sessions + (session_name -> session))) } } }) } }) val sorted_entries = (for { (name, sessions) <- data_entries.toList sorted_sessions <- proper_list(sessions.toList.map(_._2).sortBy(_.order)) } yield { val hosts = get_hosts(name).toList.sorted val stretch = data_stretch(name) Data_Entry(name, hosts, stretch, sorted_sessions) }).sortBy(_.name) Data(date, sorted_entries) } /* present data */ def present_data(data: Data, progress: Progress = new Progress, target_dir: Path = default_target_dir, 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) HTML.write_document(target_dir, "index.html", List(HTML.title("Isabelle build status")), List(HTML.chapter("Isabelle build status"), HTML.par( List(HTML.description( List(HTML.text("status date:") -> HTML.text(data.date.toString))))), HTML.par( List(HTML.itemize(data.entries.map({ case data_entry => List( HTML.link(clean_name(data_entry.name) + "/index.html", HTML.text(data_entry.name))) ::: (data_entry.failed_sessions match { case Nil => Nil case sessions => HTML.break ::: List(HTML.span(HTML.error_message, HTML.text("Failed sessions:"))) ::: List(HTML.itemize(sessions.map(s => s.head.present_errors(s.name)))) }) })))))) for (data_entry <- data.entries) { val data_name = data_entry.name val (image_width, image_height) = image_size val image_width_stretch = (image_width * data_entry.stretch).toInt progress.echo("output " + quote(data_name)) val dir = Isabelle_System.make_directory(target_dir + Path.basic(clean_name(data_name))) val data_files = (for (session <- data_entry.sessions) yield { val csv_file = session.make_csv csv_file.write(dir) session.name -> csv_file }).toMap val session_plots = Par_List.map((session: Session) => Isabelle_System.with_tmp_file(session.name, "data") { data_file => Isabelle_System.with_tmp_file(session.name, "gnuplot") { gnuplot_file => def plot_name(kind: String): String = session.name + "_" + kind + ".png" File.write(data_file, cat_lines( session.finished_entries.map(entry => List(entry.date.toString, entry.timing.elapsed.minutes.toString, entry.timing.resources.minutes.toString, entry.ml_timing.elapsed.minutes.toString, entry.ml_timing.resources.minutes.toString, entry.maximum_code.toString, entry.maximum_code.toString, entry.average_stack.toString, entry.maximum_stack.toString, entry.average_heap.toString, entry.average_heap.toString, entry.stored_heap.toString).mkString(" ")))) val max_time = - ((0.0 /: session.finished_entries){ case (m, entry) => - m.max(entry.timing.elapsed.minutes). - max(entry.timing.resources.minutes). - max(entry.ml_timing.elapsed.minutes). - max(entry.ml_timing.resources.minutes) } max 0.1) * 1.1 + (session.finished_entries.foldLeft(0.0) { + case (m, entry) => + m.max(entry.timing.elapsed.minutes). + max(entry.timing.resources.minutes). + max(entry.ml_timing.elapsed.minutes). + max(entry.ml_timing.resources.minutes) + } max 0.1) * 1.1 val timing_range = "[0:" + max_time + "]" 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, """ set terminal png size """ + image.width + "," + image.height + """ set output """ + quote(File.standard_path(dir + image.path)) + """ set xdata time set timefmt "%s" set format x "%d-%b" set xlabel """ + quote(session.name) + """ noenhanced set key left bottom plot [] """ + range + " " + plots.map(s => quote(data_file.implode) + " " + s).mkString(", ") + "\n") val result = Isabelle_System.bash("\"$ISABELLE_GNUPLOT\" " + File.bash_path(gnuplot_file)) if (!result.ok) result.error("Gnuplot failed for " + data_name + "/" + plot_name).check image } val timing_plots = { val plots1 = List( """ using 1:2 smooth sbezier title "elapsed time (smooth)" """, """ using 1:2 smooth csplines title "elapsed time" """) val plots2 = List( """ using 1:3 smooth sbezier title "cpu time (smooth)" """, """ using 1:3 smooth csplines title "cpu time" """) if (session.threads == 1) plots1 else plots1 ::: plots2 } val ml_timing_plots = List( """ using 1:4 smooth sbezier title "ML elapsed time (smooth)" """, """ using 1:4 smooth csplines title "ML elapsed time" """, """ using 1:5 smooth sbezier title "ML cpu time (smooth)" """, """ using 1:5 smooth csplines title "ML cpu time" """) val heap_plots = List( """ using 1:10 smooth sbezier title "heap maximum (smooth)" """, """ using 1:10 smooth csplines title "heap maximum" """, """ using 1:11 smooth sbezier title "heap average (smooth)" """, """ using 1:11 smooth csplines title "heap average" """, """ 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 = { val image = Image(plot_name, image_width, image_height) val chart = session.ml_statistics.chart( fields._1 + ": " + session.ml_statistics.heading, fields._2) Graphics_File.write_chart_png( (dir + image.path).file, chart, image.width, image.height) image } val images = (if (session.check_timing) List( gnuplot(plot_name("timing"), timing_plots, timing_range), gnuplot(plot_name("ml_timing"), ml_timing_plots, timing_range)) else Nil) ::: (if (session.check_heap) List(gnuplot(plot_name("heap"), heap_plots, "[0:]")) else Nil) ::: (if (session.ml_statistics.content.nonEmpty) List(jfreechart(plot_name("heap_chart"), ML_Statistics.heap_fields), jfreechart(plot_name("program_chart"), ML_Statistics.program_fields)) ::: (if (session.threads > 1) List( jfreechart(plot_name("tasks_chart"), ML_Statistics.tasks_fields), jfreechart(plot_name("workers_chart"), ML_Statistics.workers_fields)) else Nil) else Nil) session.name -> images } }, data_entry.sessions).toMap HTML.write_document(dir, "index.html", List(HTML.title("Isabelle build status for " + data_name)), HTML.chapter("Isabelle build status for " + data_name) :: HTML.par( List(HTML.description( List( HTML.text("status date:") -> HTML.text(data.date.toString), HTML.text("build host:") -> HTML.text(commas(data_entry.hosts)))))) :: HTML.par( List(HTML.itemize( data_entry.sessions.map(session => HTML.link("#session_" + session.name, HTML.text(session.name)) :: HTML.text(" (" + session.head.timing.message_resources + ")"))))) :: data_entry.sessions.flatMap(session => List( HTML.section(HTML.id("session_" + session.name), session.name), HTML.par( HTML.description( List( HTML.text("data:") -> List(HTML.link(data_files(session.name).file_name, HTML.text("CSV"))), HTML.text("timing:") -> HTML.text(session.head.timing.message_resources), HTML.text("ML timing:") -> HTML.text(session.head.ml_timing.message_resources)) ::: ML_Statistics.mem_print(session.head.maximum_code).map(s => HTML.text("code maximum:") -> HTML.text(s)).toList ::: ML_Statistics.mem_print(session.head.average_code).map(s => HTML.text("code average:") -> HTML.text(s)).toList ::: ML_Statistics.mem_print(session.head.maximum_stack).map(s => HTML.text("stack maximum:") -> HTML.text(s)).toList ::: ML_Statistics.mem_print(session.head.average_stack).map(s => HTML.text("stack average:") -> HTML.text(s)).toList ::: ML_Statistics.mem_print(session.head.maximum_heap).map(s => HTML.text("heap maximum:") -> HTML.text(s)).toList ::: ML_Statistics.mem_print(session.head.average_heap).map(s => HTML.text("heap average:") -> HTML.text(s)).toList ::: ML_Statistics.mem_print(session.head.stored_heap).map(s => HTML.text("heap stored:") -> HTML.text(s)).toList ::: proper_string(session.head.isabelle_version).map(s => HTML.text("Isabelle version:") -> HTML.text(s)).toList ::: proper_string(session.head.afp_version).map(s => HTML.text("AFP version:") -> HTML.text(s)).toList) :: session_plots.getOrElse(session.name, Nil).map(image => HTML.size(image.width / 2, image.height / 2)(HTML.image(image.name))))))) } } /* Isabelle tool wrapper */ val isabelle_tool = Isabelle_Tool("build_status", "present recent build status information from database", Scala_Project.here, args => { var target_dir = default_target_dir var ml_statistics = false var only_sessions = Set.empty[String] var options = Options.init() var image_size = default_image_size var verbose = false val getopts = Getopts(""" Usage: isabelle build_status [OPTIONS] Options are: -D DIR target directory (default """ + default_target_dir + """) -M include full ML statistics -S SESSIONS only given SESSIONS (comma separated) -l DAYS length of relevant history (default """ + options.int("build_log_history") + """) -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) -s WxH size of PNG image (default """ + image_size._1 + "x" + image_size._2 + """) -v verbose Present performance statistics from build log database, which is specified via system options build_log_database_host, build_log_database_user, build_log_history etc. """, "D:" -> (arg => target_dir = Path.explode(arg)), "M" -> (_ => ml_statistics = true), "S:" -> (arg => only_sessions = space_explode(',', arg).toSet), "l:" -> (arg => options = options + ("build_log_history=" + arg)), "o:" -> (arg => options = options + arg), "s:" -> (arg => space_explode('x', arg).map(Value.Int.parse(_)) match { case List(w, h) if w > 0 && h > 0 => image_size = (w, h) case _ => error("Error bad PNG image size: " + quote(arg)) }), "v" -> (_ => verbose = true)) val more_args = getopts(args) if (more_args.nonEmpty) getopts.usage() val progress = new Console_Progress build_status(options, progress = progress, only_sessions = only_sessions, verbose = verbose, target_dir = target_dir, ml_statistics = ml_statistics, image_size = image_size) }) } diff --git a/src/Pure/Admin/ci_profile.scala b/src/Pure/Admin/ci_profile.scala --- a/src/Pure/Admin/ci_profile.scala +++ b/src/Pure/Admin/ci_profile.scala @@ -1,163 +1,163 @@ /* Title: Pure/Admin/ci_profile.scala Author: Lars Hupel Build profile for continuous integration services. */ package isabelle import java.time.{Instant, ZoneId} import java.time.format.DateTimeFormatter import java.util.{Properties => JProperties, Map => JMap} abstract class CI_Profile extends Isabelle_Tool.Body { private def build(options: Options): (Build.Results, Time) = { val progress = new Console_Progress(verbose = true) val start_time = Time.now() val results = progress.interrupt_handler { Build.build( options + "system_heaps", selection = selection, progress = progress, clean_build = clean, verbose = true, numa_shuffling = numa, max_jobs = jobs, dirs = include, select_dirs = select) } val end_time = Time.now() (results, end_time - start_time) } private def load_properties(): JProperties = { val props = new JProperties() val file_name = Isabelle_System.getenv("ISABELLE_CI_PROPERTIES") if (file_name != "") { val file = Path.explode(file_name).file if (file.exists()) props.load(new java.io.FileReader(file)) props } else props } private def compute_timing(results: Build.Results, group: Option[String]): Timing = { val timings = results.sessions.collect { case session if group.forall(results.info(session).groups.contains(_)) => results(session).timing } - (Timing.zero /: timings)(_ + _) + timings.foldLeft(Timing.zero)(_ + _) } private def with_documents(options: Options): Options = { if (documents) options .bool.update("browser_info", true) .string.update("document", "pdf") .string.update("document_variants", "document:outline=/proof,/ML") else options } final def hg_id(path: Path): String = Mercurial.repository(path).id() final def print_section(title: String): Unit = println(s"\n=== $title ===\n") final val isabelle_home = Path.explode(Isabelle_System.getenv_strict("ISABELLE_HOME")) final val isabelle_id = hg_id(isabelle_home) final val start_time = Instant.now().atZone(ZoneId.systemDefault).format(DateTimeFormatter.RFC_1123_DATE_TIME) override final def apply(args: List[String]): Unit = { print_section("CONFIGURATION") println(Build_Log.Settings.show()) val props = load_properties() System.getProperties().asInstanceOf[JMap[AnyRef, AnyRef]].putAll(props) val options = with_documents(Options.init()) .int.update("parallel_proofs", 1) .int.update("threads", threads) println(s"jobs = $jobs, threads = $threads, numa = $numa") print_section("BUILD") println(s"Build started at $start_time") println(s"Isabelle id $isabelle_id") pre_hook(args) print_section("LOG") val (results, elapsed_time) = build(options) print_section("TIMING") val groups = results.sessions.map(results.info).flatMap(_.groups) for (group <- groups) println(s"Group $group: " + compute_timing(results, Some(group)).message_resources) val total_timing = compute_timing(results, None).copy(elapsed = elapsed_time) println("Overall: " + total_timing.message_resources) if (!results.ok) { print_section("FAILED SESSIONS") for (name <- results.sessions) { if (results.cancelled(name)) { println(s"Session $name: CANCELLED") } else { val result = results(name) if (!result.ok) println(s"Session $name: FAILED ${result.rc}") } } } post_hook(results) System.exit(results.rc) } /* profile */ def threads: Int = Isabelle_System.hostname() match { case "hpcisabelle" => 8 case "lxcisa1" => 4 case _ => 2 } def jobs: Int = Isabelle_System.hostname() match { case "hpcisabelle" => 8 case "lxcisa1" => 10 case _ => 2 } def numa: Boolean = Isabelle_System.hostname() == "hpcisabelle" def documents: Boolean = true def clean: Boolean = true def include: List[Path] def select: List[Path] def pre_hook(args: List[String]): Unit def post_hook(results: Build.Results): Unit def selection: Sessions.Selection } diff --git a/src/Pure/Admin/isabelle_cronjob.scala b/src/Pure/Admin/isabelle_cronjob.scala --- a/src/Pure/Admin/isabelle_cronjob.scala +++ b/src/Pure/Admin/isabelle_cronjob.scala @@ -1,635 +1,635 @@ /* Title: Pure/Admin/isabelle_cronjob.scala Author: Makarius Main entry point for administrative cronjob at TUM. */ package isabelle import java.nio.file.Files import scala.annotation.tailrec object Isabelle_Cronjob { /* global resources: owned by main cronjob */ val backup = "lxbroy10:cronjob" val main_dir: Path = Path.explode("~/cronjob") val main_state_file: Path = main_dir + Path.explode("run/main.state") val current_log: Path = main_dir + Path.explode("run/main.log") // owned by log service val cumulative_log: Path = main_dir + Path.explode("log/main.log") // owned by log service val isabelle_repos_source = "https://isabelle.sketis.net/repos/isabelle" val isabelle_repos: Path = main_dir + Path.explode("isabelle") val afp_repos: Path = main_dir + Path.explode("AFP") val mailman_archives_dir = Path.explode("~/cronjob/Mailman") val build_log_dirs = List(Path.explode("~/log"), Path.explode("~/afp/log"), Path.explode("~/cronjob/log")) /** logger tasks **/ sealed case class Logger_Task(name: String = "", body: Logger => Unit) /* init and exit */ def get_rev(): String = Mercurial.repository(isabelle_repos).id() def get_afp_rev(): String = Mercurial.repository(afp_repos).id() val init: Logger_Task = Logger_Task("init", logger => { Isabelle_Devel.make_index() Mercurial.setup_repository(isabelle_repos_source, isabelle_repos) Mercurial.setup_repository(AFP.repos_source, afp_repos) File.write(logger.log_dir + Build_Log.log_filename("isabelle_identify", logger.start_date), Build_Log.Identify.content(logger.start_date, Some(get_rev()), Some(get_afp_rev()))) Isabelle_System.bash( """rsync -a --include="*/" --include="plain_identify*" --exclude="*" """ + Bash.string(backup + "/log/.") + " " + File.bash_path(main_dir) + "/log/.").check if (!Isabelle_Devel.cronjob_log.is_file) Files.createSymbolicLink(Isabelle_Devel.cronjob_log.file.toPath, current_log.file.toPath) }) val exit: Logger_Task = Logger_Task("exit", logger => { Isabelle_System.bash( "rsync -a " + File.bash_path(main_dir) + "/log/." + " " + Bash.string(backup) + "/log/.") .check }) /* Mailman archives */ val mailman_archives: Logger_Task = Logger_Task("mailman_archives", logger => { Mailman.isabelle_users.download(mailman_archives_dir) Mailman.isabelle_dev.download(mailman_archives_dir) }) /* build release */ val build_release: Logger_Task = Logger_Task("build_release", logger => { Isabelle_Devel.release_snapshot(logger.options, rev = get_rev(), afp_rev = get_afp_rev()) }) /* remote build_history */ sealed case class Item( known: Boolean, isabelle_version: String, afp_version: Option[String], pull_date: Date) { def unknown: Boolean = !known def versions: (String, Option[String]) = (isabelle_version, afp_version) def known_versions(rev: String, afp_rev: Option[String]): Boolean = known && rev != "" && isabelle_version == rev && (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] = { 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 => { 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 val pull_date = res.date(Build_Log.Data.pull_date(afp)) Item(known, isabelle_version, afp_version, pull_date) }).toList) } def unknown_runs(items: List[Item]): List[List[Item]] = { val (run, rest) = Library.take_prefix[Item](_.unknown, items.dropWhile(_.known)) if (run.nonEmpty) run :: unknown_runs(rest) else Nil } sealed case class Remote_Build( description: String, host: String, actual_host: String = "", user: String = "", port: Int = 0, proxy_host: String = "", proxy_user: String = "", proxy_port: Int = 0, self_update: Boolean = false, historic: Boolean = false, history: Int = 0, history_base: String = "build_history_base", java_heap: String = "", options: String = "", args: String = "", afp: Boolean = false, bulky: Boolean = false, more_hosts: List[String] = Nil, detect: SQL.Source = "", 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, permissive = proxy_host.nonEmpty) def sql: SQL.Source = Build_Log.Prop.build_engine.toString + " = " + SQL.string(Build_History.engine) + " AND " + SQL.member(Build_Log.Prop.build_host.ident, host :: more_hosts) + (if (detect == "") "" else " AND " + SQL.enclose(detect)) def profile: Build_Status.Profile = Build_Status.Profile(description, history = history, afp = afp, bulky = bulky, sql = sql) def pick( options: Options, rev: 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])] = { val items = recent_items(db, days, rev, afp_rev, sql).filter(filter) def runs = unknown_runs(items).filter(run => run.length >= gap) if (historic || items.exists(_.known_versions(rev, afp_rev))) { val longest_run = - (List.empty[Item] /: runs)({ case (item1, item2) => - if (item1.length >= item2.length) item1 else item2 - }) + runs.foldLeft(List.empty[Item]) { + case (item1, item2) => if (item1.length >= item2.length) item1 else item2 + } if (longest_run.isEmpty) None else Some(longest_run(longest_run.length / 2).versions) } else if (rev != "") Some((rev, afp_rev)) else runs.flatten.headOption.map(_.versions) } pick_days(options.int("build_log_history") max history, 2) orElse pick_days(200, 5) orElse pick_days(2000, 1) }) } def build_history_options: String = " -h " + Bash.string(host) + " " + (java_heap match { case "" => "" case h => "-e 'ISABELLE_TOOL_JAVA_OPTIONS=\"$ISABELLE_TOOL_JAVA_OPTIONS -Xmx" + h + "\"' " }) + options } val remote_builds_old: List[Remote_Build] = List( Remote_Build("Linux A", "lxbroy9", java_heap = "2g", options = "-m32 -B -M1x2,2", args = "-N -g timing"), Remote_Build("Linux Benchmarks", "lxbroy5", historic = true, history = 90, java_heap = "2g", options = "-m32 -B -M1x2,2 -t Benchmarks" + " -e ISABELLE_GHC=ghc -e ISABELLE_MLTON=mlton -e ISABELLE_OCAML=ocaml" + " -e ISABELLE_OCAMLC=ocamlc -e ISABELLE_OCAMLFIND=ocamlfind -e ISABELLE_SMLNJ=sml" + " -e ISABELLE_SWIPL=swipl", args = "-N -a -d '~~/src/Benchmarks'", detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("Benchmarks")), Remote_Build("macOS 10.14 Mojave (Old)", "lapnipkow3", options = "-m32 -M1,2 -e ISABELLE_GHC_SETUP=true -p pide_session=false", self_update = true, args = "-a -d '~~/src/Benchmarks'"), Remote_Build("AFP old bulky", "lrzcloud1", self_update = true, proxy_host = "lxbroy10", proxy_user = "i21isatest", options = "-m64 -M6 -U30000 -s10 -t AFP", args = "-g large -g slow", afp = true, bulky = true, detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("AFP")), Remote_Build("AFP old", "lxbroy7", args = "-N -X large -X slow", afp = true, detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("AFP")), Remote_Build("Poly/ML 5.7 Linux", "lxbroy8", history_base = "37074e22e8be", options = "-m32 -B -M1x2,2 -t polyml-5.7 -i 'init_component /home/isabelle/contrib/polyml-5.7'", args = "-N -g timing", detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("polyml-5.7") + " AND " + Build_Log.Settings.ML_OPTIONS + " <> " + SQL.string("-H 500")), Remote_Build("Poly/ML 5.7.1 Linux", "lxbroy8", history_base = "a9d5b59c3e12", options = "-m32 -B -M1x2,2 -t polyml-5.7.1-pre2 -i 'init_component /home/isabelle/contrib/polyml-test-905dae2ebfda'", args = "-N -g timing", detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("polyml-5.7.1-pre1") + " OR " + Build_Log.Prop.build_tags + " = " + SQL.string("polyml-5.7.1-pre2")), Remote_Build("Poly/ML 5.7 macOS", "macbroy2", history_base = "37074e22e8be", options = "-m32 -B -M1x4,4 -t polyml-5.7 -i 'init_component /home/isabelle/contrib/polyml-5.7'", args = "-a", detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("polyml-5.7")), Remote_Build("Poly/ML 5.7.1 macOS", "macbroy2", history_base = "a9d5b59c3e12", options = "-m32 -B -M1x4,4 -t polyml-5.7.1-pre2 -i 'init_component /home/isabelle/contrib/polyml-test-905dae2ebfda'", args = "-a", detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("polyml-5.7.1-pre1") + " OR " + Build_Log.Prop.build_tags + " = " + SQL.string("polyml-5.7.1-pre2")), Remote_Build("macOS", "macbroy2", options = "-m32 -M8" + " -e ISABELLE_GHC=ghc -e ISABELLE_MLTON=mlton -e ISABELLE_OCAML=ocaml" + " -e ISABELLE_OCAMLC=ocamlc -e ISABELLE_OCAML_SETUP=true" + " -e ISABELLE_OPAM_ROOT=\"$ISABELLE_HOME/opam\"" + " -e ISABELLE_SMLNJ=/home/isabelle/smlnj/110.85/bin/sml" + " -p pide_session=false", args = "-a", detect = Build_Log.Prop.build_tags.undefined, history_base = "2c0f24e927dd"), Remote_Build("macOS, quick_and_dirty", "macbroy2", options = "-m32 -M8 -t quick_and_dirty -p pide_session=false", args = "-a -o quick_and_dirty", detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("quick_and_dirty"), history_base = "2c0f24e927dd"), Remote_Build("macOS, skip_proofs", "macbroy2", options = "-m32 -M8 -t skip_proofs -p pide_session=false", args = "-a -o skip_proofs", detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("skip_proofs"), history_base = "2c0f24e927dd"), Remote_Build("Poly/ML test", "lxbroy8", options = "-m32 -B -M1x2,2 -t polyml-test -i 'init_component /home/isabelle/contrib/polyml-5.7-20170217'", args = "-N -g timing", detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("polyml-test")), Remote_Build("macOS 10.12 Sierra", "macbroy30", options = "-m32 -M2 -p pide_session=false", args = "-a", detect = Build_Log.Prop.build_start.toString + " > date '2017-03-03'"), Remote_Build("macOS 10.10 Yosemite", "macbroy31", options = "-m32 -M2 -p pide_session=false", args = "-a"), Remote_Build("macOS 10.8 Mountain Lion", "macbroy30", options = "-m32 -M2", args = "-a", detect = Build_Log.Prop.build_start.toString + " < date '2017-03-03'")) ::: { for { (n, hosts) <- List(1 -> List("lxbroy6"), 2 -> List("lxbroy8", "lxbroy5")) } yield { Remote_Build("AFP old", host = hosts.head, more_hosts = hosts.tail, options = "-m32 -M1x2 -t AFP -P" + n + " -e ISABELLE_GHC=ghc" + " -e ISABELLE_MLTON=mlton" + " -e ISABELLE_OCAML=ocaml -e ISABELLE_OCAMLC=ocamlc -e ISABELLE_OCAMLFIND=ocamlfind" + " -e ISABELLE_SMLNJ=sml", args = "-N -X large -X slow", afp = true, detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("AFP")) } } val remote_builds1: List[List[Remote_Build]] = { List( List(Remote_Build("Linux B", "lxbroy10", historic = true, history = 90, options = "-m32 -B -M1x4,2,4,6", args = "-N -g timing")), List(Remote_Build("macOS 10.13 High Sierra", "lapbroy68", options = "-m32 -B -M1,2,4 -e ISABELLE_GHC_SETUP=true -p pide_session=false", self_update = true, args = "-a -d '~~/src/Benchmarks'")), List( Remote_Build("macOS 11.1 Big Sur", "mini1", options = "-m32 -B -M1x2,2,4 -p pide_session=false" + " -e ISABELLE_OCAML=ocaml -e ISABELLE_OCAMLC=ocamlc -e ISABELLE_OCAML_SETUP=true" + " -e ISABELLE_GHC_SETUP=true" + " -e ISABELLE_MLTON=/usr/local/bin/mlton" + " -e ISABELLE_SMLNJ=/usr/local/smlnj/bin/sml" + " -e ISABELLE_SWIPL=/usr/local/bin/swipl", self_update = true, args = "-a -d '~~/src/Benchmarks'")), List( Remote_Build("macOS 10.14 Mojave", "mini2", options = "-m32 -B -M1x2,2,4 -p pide_session=false" + " -e ISABELLE_OCAML=ocaml -e ISABELLE_OCAMLC=ocamlc -e ISABELLE_OCAML_SETUP=true" + " -e ISABELLE_GHC_SETUP=true" + " -e ISABELLE_MLTON=/usr/local/bin/mlton" + " -e ISABELLE_SMLNJ=/usr/local/smlnj/bin/sml" + " -e ISABELLE_SWIPL=/usr/local/bin/swipl", self_update = true, args = "-a -d '~~/src/Benchmarks'"), Remote_Build("macOS, quick_and_dirty", "mini2", options = "-m32 -M4 -t quick_and_dirty -p pide_session=false", self_update = true, args = "-a -o quick_and_dirty", detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("quick_and_dirty")), Remote_Build("macOS, skip_proofs", "mini2", options = "-m32 -M4 -t skip_proofs -p pide_session=false", args = "-a -o skip_proofs", detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("skip_proofs"))), List(Remote_Build("macOS 10.15 Catalina", "laramac01", user = "makarius", proxy_host = "laraserver", proxy_user = "makarius", self_update = true, options = "-m32 -M4 -e ISABELLE_GHC_SETUP=true -p pide_session=false", args = "-a -d '~~/src/Benchmarks'")), List( Remote_Build("Windows", "vmnipkow9", historic = true, history = 90, self_update = true, options = "-m32 -M4" + " -C /cygdrive/d/isatest/contrib" + " -e ISABELLE_OCAML=ocaml -e ISABELLE_OCAMLC=ocamlc -e ISABELLE_OCAML_SETUP=true" + " -e ISABELLE_GHC_SETUP=true" + " -e ISABELLE_SMLNJ=/usr/local/smlnj-110.81/bin/sml", args = "-a", detect = Build_Log.Settings.ML_PLATFORM.toString + " = " + SQL.string("x86-windows") + " OR " + Build_Log.Settings.ML_PLATFORM + " = " + SQL.string("x86_64_32-windows")), Remote_Build("Windows", "vmnipkow9", historic = true, history = 90, self_update = true, options = "-m64 -M4" + " -C /cygdrive/d/isatest/contrib" + " -e ISABELLE_OCAML=ocaml -e ISABELLE_OCAMLC=ocamlc -e ISABELLE_OCAML_SETUP=true" + " -e ISABELLE_GHC_SETUP=true" + " -e ISABELLE_SMLNJ=/usr/local/smlnj-110.81/bin/sml", args = "-a", detect = Build_Log.Settings.ML_PLATFORM.toString + " = " + SQL.string("x86_64-windows")))) } val remote_builds2: List[List[Remote_Build]] = List( List( Remote_Build("AFP", "lrzcloud2", actual_host = "10.195.4.41", self_update = true, proxy_host = "lxbroy10", proxy_user = "i21isatest", java_heap = "8g", options = "-m32 -M1x6 -t AFP" + " -e ISABELLE_GHC=ghc" + " -e ISABELLE_MLTON=mlton" + " -e ISABELLE_OCAML=ocaml -e ISABELLE_OCAMLC=ocamlc -e ISABELLE_OCAMLFIND=ocamlfind" + " -e ISABELLE_SMLNJ=sml", args = "-a -X large -X slow", afp = true, detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("AFP")), Remote_Build("AFP", "lrzcloud2", actual_host = "10.195.4.41", self_update = true, proxy_host = "lxbroy10", proxy_user = "i21isatest", java_heap = "8g", options = "-m64 -M8 -U30000 -s10 -t AFP", args = "-g large -g slow", afp = true, 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 = { val task_name = "build_history-" + r.host Logger_Task(task_name, logger => { using(r.ssh_session(logger.ssh_context))(ssh => { val results = Build_History.remote_build_history(ssh, isabelle_repos, isabelle_repos.ext(r.host), isabelle_identifier = "cronjob_build_history", self_update = r.self_update, rev = rev, afp_rev = afp_rev, options = " -N " + Bash.string(task_name) + (if (i < 0) "" else "_" + (i + 1).toString) + " -R " + Bash.string(Components.default_component_repository) + " -C '$USER_HOME/.isabelle/contrib' -f " + r.build_history_options, args = "-o timeout=10800 " + r.args) for ((log_name, bytes) <- results) { logger.log(Date.now(), log_name) Bytes.write(logger.log_dir + Path.explode(log_name), bytes) } }) }) } val build_status_profiles: List[Build_Status.Profile] = (remote_builds_old :: remote_builds1 ::: remote_builds2).flatten.map(_.profile) /** task logging **/ 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) { current_log.file.delete private val thread: Consumer_Thread[String] = Consumer_Thread.fork("cronjob: logger", daemon = true)( consume = (text: String) => { // critical File.append(current_log, text + "\n") File.append(cumulative_log, text + "\n") progress.echo(text) true }) def shutdown(): Unit = { thread.shutdown() } val hostname: String = Isabelle_System.hostname() def log(date: Date, task_name: String, msg: String): Unit = if (task_name != "") thread.send( "[" + Build_Log.print_date(date) + ", " + hostname + ", " + task_name + "]: " + msg) def start_logger(start_date: Date, task_name: String): Logger = new Logger(this, start_date, task_name) def run_task(start_date: Date, task: Logger_Task): Unit = { val logger = start_logger(start_date, task.name) val res = Exn.capture { task.body(logger) } val end_date = Date.now() val err = res match { case Exn.Res(_) => None case Exn.Exn(exn) => Output.writeln("Exception trace for " + quote(task.name) + ":") exn.printStackTrace() val first_line = split_lines(Exn.message(exn)).headOption getOrElse "exception" Some(first_line) } logger.log_end(end_date, err) } def fork_task(start_date: Date, task: Logger_Task): Task = new Task(task.name, run_task(start_date, task)) } class Logger private[Isabelle_Cronjob]( val log_service: Log_Service, val start_date: Date, val task_name: String) { def ssh_context: SSH.Context = log_service.ssh_context def options: Options = ssh_context.options def log(date: Date, msg: String): Unit = log_service.log(date, task_name, msg) def log_end(end_date: Date, err: Option[String]): Unit = { val elapsed_time = end_date.time - start_date.time val msg = (if (err.isEmpty) "finished" else "ERROR " + err.get) + (if (elapsed_time.seconds < 3.0) "" else " (" + elapsed_time.message_hms + " elapsed time)") log(end_date, msg) } val log_dir = Isabelle_System.make_directory(main_dir + Build_Log.log_subdir(start_date)) log(start_date, "started") } class Task private[Isabelle_Cronjob](name: String, body: => Unit) { private val future: Future[Unit] = Future.thread("cronjob: " + name) { body } def is_finished: Boolean = future.is_finished } /** cronjob **/ def cronjob(progress: Progress, exclude_task: Set[String]): Unit = { /* soft lock */ val still_running = try { Some(File.read(main_state_file)) } catch { case ERROR(_) => None } still_running match { case None | Some("") => case Some(running) => error("Isabelle cronjob appears to be still running: " + running) } /* log service */ val log_service = Log_Service(Options.init(), progress = progress) def run(start_date: Date, task: Logger_Task): Unit = log_service.run_task(start_date, task) def run_now(task: Logger_Task): Unit = run(Date.now(), task) /* structured tasks */ def SEQ(tasks: List[Logger_Task]): Logger_Task = Logger_Task(body = _ => for (task <- tasks.iterator if !exclude_task(task.name) || task.name == "") run_now(task)) def PAR(tasks: List[Logger_Task]): Logger_Task = Logger_Task(body = _ => { @tailrec def join(running: List[Task]): Unit = { running.partition(_.is_finished) match { case (Nil, Nil) => case (Nil, _ :: _) => Time.seconds(0.5).sleep; join(running) case (_ :: _, remaining) => join(remaining) } } val start_date = Date.now() val running = for (task <- tasks if !exclude_task(task.name)) yield log_service.fork_task(start_date, task) join(running) }) /* repository structure */ val hg = Mercurial.repository(isabelle_repos) val hg_graph = hg.graph() 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) } /* main */ val main_start_date = Date.now() File.write(main_state_file, main_start_date.toString + " " + log_service.hostname) run(main_start_date, Logger_Task("isabelle_cronjob", logger => run_now( SEQ(List( init, PAR(List(mailman_archives, build_release)), PAR( List(remote_builds1, remote_builds2).map(remote_builds => SEQ(List( PAR(remote_builds.map(_.filter(_.active)).map(seq => SEQ( for { (r, i) <- (if (seq.length <= 1) seq.map((_, -1)) else seq.zipWithIndex) (rev, afp_rev) <- r.pick(logger.options, hg.id(), history_base_filter(r)) } yield remote_build_history(rev, afp_rev, i, r)))), Logger_Task("jenkins_logs", _ => Jenkins.download_logs(logger.options, Jenkins.build_log_jobs, main_dir)), Logger_Task("build_log_database", logger => Isabelle_Devel.build_log_database(logger.options, build_log_dirs)), Logger_Task("build_status", logger => Isabelle_Devel.build_status(logger.options)))))), exit))))) log_service.shutdown() main_state_file.file.delete } /** command line entry point **/ def main(args: Array[String]): Unit = { Command_Line.tool { var force = false var verbose = false var exclude_task = Set.empty[String] val getopts = Getopts(""" Usage: Admin/cronjob/main [OPTIONS] Options are: -f apply force to do anything -v verbose -x NAME exclude tasks with this name """, "f" -> (_ => force = true), "v" -> (_ => verbose = true), "x:" -> (arg => exclude_task += arg)) val more_args = getopts(args) if (more_args.nonEmpty) getopts.usage() val progress = if (verbose) new Console_Progress() else new Progress if (force) cronjob(progress, exclude_task) else error("Need to apply force to do anything") } } } diff --git a/src/Pure/General/completion.scala b/src/Pure/General/completion.scala --- a/src/Pure/General/completion.scala +++ b/src/Pure/General/completion.scala @@ -1,516 +1,516 @@ /* Title: Pure/General/completion.scala Author: Makarius Semantic completion within the formal context (reported names). Syntactic completion of keywords and symbols, with abbreviations (based on language context). */ package isabelle import scala.collection.immutable.SortedMap import scala.util.parsing.combinator.RegexParsers import scala.util.matching.Regex import scala.math.Ordering object Completion { /** completion result **/ sealed case class Item( range: Text.Range, original: String, name: String, description: List[String], replacement: String, move: Int, immediate: Boolean) object Result { def empty(range: Text.Range): Result = Result(range, "", false, Nil) def merge(history: History, results: Option[Result]*): Option[Result] = - ((None: Option[Result]) /: results)({ + results.foldLeft(None: Option[Result]) { case (result1, None) => result1 case (None, result2) => result2 case (result1 @ Some(res1), Some(res2)) => if (res1.range != res2.range || res1.original != res2.original) result1 else { val items = (res1.items ::: res2.items).sorted(history.ordering) Some(Result(res1.range, res1.original, false, items)) } - }) + } } sealed case class Result( range: Text.Range, original: String, unique: Boolean, items: List[Item]) /** persistent history **/ private val COMPLETION_HISTORY = Path.explode("$ISABELLE_HOME_USER/etc/completion_history") object History { val empty: History = new History() def load(): History = { def ignore_error(msg: String): Unit = Output.warning("Ignoring bad content of file " + COMPLETION_HISTORY + (if (msg == "") "" else "\n" + msg)) val content = if (COMPLETION_HISTORY.is_file) { try { import XML.Decode._ list(pair(string, int))(Symbol.decode_yxml(File.read(COMPLETION_HISTORY))) } catch { case ERROR(msg) => ignore_error(msg); Nil case _: XML.Error => ignore_error(""); Nil } } else Nil - (empty /: content)(_ + _) + content.foldLeft(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 = { val (name, freq) = entry if (name == "") this else new History(rep + (name -> (frequency(name) + freq))) } def ordering: Ordering[Item] = new Ordering[Item] { def compare(item1: Item, item2: Item): Int = frequency(item2.name) compare frequency(item1.name) } def save(): Unit = { Isabelle_System.make_directory(COMPLETION_HISTORY.dir) File.write_backup(COMPLETION_HISTORY, { import XML.Encode._ Symbol.encode_yxml(list(pair(string, int))(rep.toList)) }) } } class History_Variable { private var history = History.empty def value: History = synchronized { history } def load(): Unit = { val h = History.load() synchronized { history = h } } def update(item: Item, freq: Int = 1): Unit = synchronized { history = history + (item.name -> freq) } } /** semantic completion **/ def clean_name(s: String): Option[String] = if (s == "" || s == "_") None else Some(s.reverseIterator.dropWhile(_ == '_').toList.reverse.mkString) def completed(input: String): String => Boolean = clean_name(input) match { case None => (name: String) => false case Some(prefix) => (name: String) => name.startsWith(prefix) } def report_no_completion(pos: Position.T): String = YXML.string_of_tree(Semantic.Info(pos, No_Completion)) def report_names(pos: Position.T, names: List[(String, (String, String))], total: Int = 0): String = if (names.isEmpty) "" else YXML.string_of_tree(Semantic.Info(pos, Names(if (total > 0) total else names.length, names))) 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 = { val elem = semantic match { case No_Completion => XML.Elem(Markup(Markup.NO_COMPLETION, pos), Nil) case Names(total, names) => XML.Elem(Markup(Markup.COMPLETION, pos), { import XML.Encode._ pair(int, list(pair(string, pair(string, string))))(total, names) }) } XML.Elem(Markup(Markup.REPORT, pos), List(elem)) } def unapply(info: Text.Markup): Option[Text.Info[Semantic]] = info.info match { case XML.Elem(Markup(Markup.COMPLETION, _), body) => try { val (total, names) = { import XML.Decode._ pair(int, list(pair(string, pair(string, string))))(body) } Some(Text.Info(info.range, Names(total, names))) } catch { case _: XML.Error => None } case XML.Elem(Markup(Markup.NO_COMPLETION, _), _) => Some(Text.Info(info.range, No_Completion)) case _ => None } } } sealed abstract class Semantic case object No_Completion 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] = { def decode(s: String): String = if (unicode) Symbol.decode(s) else s val items = for { (xname, (kind, name)) <- names xname1 = decode(xname) if xname1 != original (full_name, descr_name) = if (kind == "") (name, quote(decode(name))) else (Long_Name.qualify(kind, name), Word.implode(Word.explode('_', kind)) + (if (xname == name) "" else " " + quote(decode(name)))) } yield { val description = List(xname1, "(" + descr_name + ")") val replacement = List(original, xname1).map(Token.explode(Keyword.Keywords.empty, _)) match { case List(List(tok), _) if tok.kind == Token.Kind.CARTOUCHE => Symbol.cartouche_decoded(xname1) case List(_, List(tok)) if tok.is_name => xname1 case _ => quote(xname1) } Item(range, original, full_name, description, replacement, 0, true) } if (items.isEmpty) None else Some(Result(range, original, names.length == 1, items.sorted(history.ordering))) } } /** syntactic completion **/ /* 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) val ML_inner = Language_Context(Markup.Language.ML, true, false) val SML_outer = Language_Context(Markup.Language.SML, false, false) } sealed case class Language_Context(language: String, symbols: Boolean, antiquotes: Boolean) { def is_outer: Boolean = language == "" } /* make */ val empty: Completion = new Completion() def make(keywords: List[String], abbrevs: List[(String, String)]): Completion = empty.add_symbols.add_keywords(keywords).add_abbrevs( Completion.symbol_abbrevs ::: Completion.default_abbrevs ::: abbrevs) /* word parsers */ object Word_Parsers extends RegexParsers { override val whiteSpace = "".r private val symboloid_regex: Regex = """\\([A-Za-z0-9_']+|<\^?[A-Za-z0-9_']+>?)""".r def is_symboloid(s: CharSequence): Boolean = symboloid_regex.pattern.matcher(s).matches private def reverse_symbol: Parser[String] = """>[A-Za-z0-9_']+\^?<\\""".r private def reverse_symb: Parser[String] = """[A-Za-z0-9_']{2,}\^?<\\""".r private def reverse_escape: Parser[String] = """[a-zA-Z0-9_']+\\""".r private val word_regex = "[a-zA-Z0-9_'.]+".r private def word2: Parser[String] = "[a-zA-Z0-9_'.]{2,}".r private def underscores: Parser[String] = "_*".r 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] = { val reverse_in = new Library.Reverse(in) parse(reverse_symbol ^^ (_.reverse), reverse_in) match { case Success(result, _) => Some(result) case _ => None } } 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, "")) | underscores ~ word2 ~ opt("?") ^^ { case x ~ y ~ z => (z.getOrElse("") + y.reverse, x) } parse(parser, reverse_in) match { case Success(result, _) => Some(result) case _ => None } } } /* templates */ val caret_indicator = '\u0007' def split_template(s: String): (String, String) = space_explode(caret_indicator, s) match { case List(s1, s2) => (s1, s2) case _ => (s, "") } /* abbreviations */ private def symbol_abbrevs: Thy_Header.Abbrevs = for ((sym, abbr) <- Symbol.abbrevs.toList) yield (abbr, sym) private val antiquote = "@{" private val default_abbrevs = List("@{" -> "@{\u0007}", "`" -> "\\", "`" -> "\\", "`" -> "\\\u0007\\", "\"" -> "\\", "\"" -> "\\", "\"" -> "\\\u0007\\") private def default_frequency(name: String): Option[Int] = default_abbrevs.iterator.map(_._2).zipWithIndex.find(_._1 == name).map(_._2) } final class Completion private( keywords: Set[String] = Set.empty, 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) { /* keywords */ private def is_symbol(name: String): Boolean = Symbol.names.isDefinedAt(name) private def is_keyword(name: String): Boolean = !is_symbol(name) && keywords(name) private def is_keyword_template(name: String, template: Boolean): Boolean = is_keyword(name) && (words_map.getOrElse(name, name) != name) == template 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 = { - val keywords1 = (keywords /: names) { case (ks, k) => if (ks(k)) ks else ks + k } + 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) } /* symbols and abbrevs */ def add_symbols: Completion = { val words = Symbol.names.toList.flatMap({ case (sym, (name, argument)) => val word = "\\" + name val seps = if (argument == Symbol.ARGUMENT_CARTOUCHE) List("") else if (argument == Symbol.ARGUMENT_SPACE_CARTOUCHE) List(" ") else Nil List(sym -> sym, word -> sym) ::: seps.map(sep => word -> (sym + sep + "\\\u0007\\")) }) new Completion( keywords, words_lex ++ words.map(_._1), words_map ++ words, abbrevs_lex, abbrevs_map) } def add_abbrevs(abbrevs: List[(String, String)]): Completion = - (this /: abbrevs)(_.add_abbrev(_)) + abbrevs.foldLeft(this)(_.add_abbrev(_)) private def add_abbrev(abbrev: (String, String)): Completion = abbrev match { case ("", _) => this case (abbr, text) => val rev_abbr = abbr.reverse val is_word = Completion.Word_Parsers.is_word(abbr) val (words_lex1, words_map1) = if (!is_word) (words_lex, words_map) else if (text != "") (words_lex + abbr, words_map + abbrev) else (words_lex -- List(abbr), words_map - abbr) val (abbrevs_lex1, abbrevs_map1) = if (is_word) (abbrevs_lex, abbrevs_map) else if (text != "") (abbrevs_lex + rev_abbr, abbrevs_map + (rev_abbr -> abbrev)) else (abbrevs_lex -- List(rev_abbr), abbrevs_map - rev_abbr) new Completion(keywords, words_lex1, words_map1, abbrevs_lex1, abbrevs_map1) } /* complete */ def complete( history: Completion.History, unicode: Boolean, explicit: Boolean, start: Text.Offset, text: CharSequence, caret: Int, language_context: Completion.Language_Context): Option[Completion.Result] = { val length = text.length 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, _) => val abbrevs = abbrevs_map.get_list(reverse_abbr) abbrevs match { case Nil => None case (abbr, _) :: _ => val ok = if (abbr == Completion.antiquote) language_context.antiquotes else language_context.symbols || Completion.default_abbrevs.exists(_._1 == abbr) if (ok) Some((abbr, abbrevs)) else None } case _ => None } } val words_result = if (abbrevs_result.isDefined) None else { val word_context = caret < length && Completion.Word_Parsers.is_word_char(text.charAt(caret)) val result = Completion.Word_Parsers.read_symbol(text.subSequence(0, caret)) match { case Some(symbol) => Some((symbol, "")) case None => Completion.Word_Parsers.read_word(text.subSequence(0, caret)) } result.map( { case (word, underscores) => val complete_words = words_lex.completions(word) val full_word = word + underscores val completions = if (complete_words.contains(full_word) && is_keyword_template(full_word, false)) Nil else for { complete_word <- complete_words ok = if (is_keyword(complete_word)) !word_context && language_context.is_outer else language_context.symbols || Completion.Word_Parsers.is_symboloid(word) if ok completion <- words_map.get_list(complete_word) } yield (complete_word, completion) (full_word, completions) }) } (abbrevs_result orElse words_result) match { case Some((original, completions)) if completions.nonEmpty => val range = Text.Range(- original.length, 0) + caret + start val immediate = explicit || (!Completion.Word_Parsers.is_word(original) && !Completion.Word_Parsers.is_symboloid(original) && Character.codePointCount(original, 0, original.length) > 1) val unique = completions.length == 1 def decode1(s: String): String = if (unicode) Symbol.decode(s) else s def decode2(s: String): String = if (unicode) s else Symbol.decode(s) val items = for { (complete_word, name0) <- completions name1 = decode1(name0) name2 = decode2(name0) if name1 != original (s1, s2) = Completion.split_template(name1) move = - s2.length description = if (is_symbol(name0)) { if (name1 == name2) List(name1) else List(name1, "(symbol " + quote(name2) + ")") } else if (is_keyword_template(complete_word, true)) List(name1, "(template " + quote(complete_word) + ")") else if (move != 0) List(name1, "(template)") else if (is_keyword(complete_word)) List(name1, "(keyword)") else List(name1) } yield Completion.Item(range, original, name1, description, s1 + s2, move, immediate) if (items.isEmpty) None else Some(Completion.Result(range, original, unique, items.sortBy(_.name).sorted(history.ordering))) case _ => None } } } diff --git a/src/Pure/General/file_watcher.scala b/src/Pure/General/file_watcher.scala --- a/src/Pure/General/file_watcher.scala +++ b/src/Pure/General/file_watcher.scala @@ -1,135 +1,135 @@ /* Title: Pure/General/file_watcher.scala Author: Makarius Watcher for file-system events. */ package isabelle import java.io.{File => JFile} import java.nio.file.FileSystems import java.nio.file.{WatchKey, WatchEvent, Path => JPath} import java.nio.file.StandardWatchEventKinds.{ENTRY_CREATE, ENTRY_DELETE, ENTRY_MODIFY} import scala.jdk.CollectionConverters._ class File_Watcher private[File_Watcher] // dummy template { def register(dir: JFile): Unit = {} def register_parent(file: JFile): Unit = {} def deregister(dir: JFile): Unit = {} def purge(retain: Set[JFile]): Unit = {} def shutdown(): Unit = {} } object File_Watcher { val none: File_Watcher = new File_Watcher { override def toString: String = "File_Watcher.none" } def apply(handle: Set[JFile] => Unit, delay: => Time = Time.seconds(0.5)): File_Watcher = if (Platform.is_windows) none else new Impl(handle, delay) /* proper implementation */ sealed case class State( 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 { private val state = Synchronized(File_Watcher.State()) private val watcher = FileSystems.getDefault.newWatchService() override def toString: String = state.value.dirs.keySet.mkString("File_Watcher(", ", ", ")") /* registered directories */ override def register(dir: JFile): Unit = state.change(st => st.dirs.get(dir) match { case Some(key) if key.isValid => st case _ => val key = dir.toPath.register(watcher, ENTRY_CREATE, ENTRY_DELETE, ENTRY_MODIFY) st.copy(dirs = st.dirs + (dir -> key)) }) override def register_parent(file: JFile): Unit = { val dir = file.getParentFile if (dir != null && dir.isDirectory) register(dir) } override def deregister(dir: JFile): Unit = state.change(st => st.dirs.get(dir) match { case None => st case Some(key) => key.cancel st.copy(dirs = st.dirs - dir) }) override def purge(retain: Set[JFile]): Unit = state.change(st => st.copy(dirs = st.dirs -- (for ((dir, key) <- st.dirs.iterator if !retain(dir)) yield { key.cancel; dir }))) /* changed directory entries */ 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) { try { while (true) { val key = watcher.take val has_changed = state.change_result(st => { val (remove, changed) = st.dirs.collectFirst({ case (dir, key1) if key == key1 => dir }) match { case Some(dir) => val events: Iterable[WatchEvent[JPath]] = key.pollEvents.asInstanceOf[java.util.List[WatchEvent[JPath]]].asScala val remove = if (key.reset) None else Some(dir) val changed = - (Set.empty[JFile] /: events.iterator) { + events.iterator.foldLeft(Set.empty[JFile]) { case (set, event) => set + dir.toPath.resolve(event.context).toFile } (remove, changed) case None => key.pollEvents key.reset (None, Set.empty[JFile]) } (changed.nonEmpty, st.copy(dirs = st.dirs -- remove, changed = st.changed ++ changed)) }) if (has_changed) delay_changed.invoke() } } catch { case Exn.Interrupt() => } } /* shutdown */ override def shutdown(): Unit = { watcher_thread.interrupt watcher_thread.join delay_changed.revoke } } } diff --git a/src/Pure/General/graph.scala b/src/Pure/General/graph.scala --- a/src/Pure/General/graph.scala +++ b/src/Pure/General/graph.scala @@ -1,334 +1,342 @@ /* Title: Pure/General/graph.scala Author: Makarius Directed graphs. */ package isabelle import scala.collection.immutable.{SortedMap, SortedSet} import scala.annotation.tailrec 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 { override def toString: String = "Graph.Undefined(" + key.toString + ")" } 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] = { val graph1 = - (empty[Key, A](ord) /: entries) { case (graph, ((x, info), _)) => graph.new_node(x, info) } + entries.foldLeft(empty[Key, A](ord)) { + case (graph, ((x, info), _)) => graph.new_node(x, info) + } val graph2 = - (graph1 /: entries) { case (graph, ((x, _), ys)) => - (graph /: ys)({ case (g, y) => if (symmetric) g.add_edge(y, x) else g.add_edge(x, y) }) } + entries.foldLeft(graph1) { + case (graph, ((x, _), ys)) => + ys.foldLeft(graph) { + case (g, y) => if (symmetric) g.add_edge(y, x) else g.add_edge(x, y) + } + } graph2 } def string[A]: Graph[String, A] = empty(Ordering.String) def int[A]: Graph[Int, A] = empty(Ordering.Int) def long[A]: Graph[Long, A] = empty(Ordering.Long) /* XML data representation */ def encode[Key, A](key: XML.Encode.T[Key], info: XML.Encode.T[A]): XML.Encode.T[Graph[Key, A]] = (graph: Graph[Key, A]) => { import XML.Encode._ list(pair(pair(key, info), list(key)))(graph.dest) } def decode[Key, A](key: XML.Decode.T[Key], info: XML.Decode.T[A])( implicit ord: Ordering[Key]): XML.Decode.T[Graph[Key, A]] = (body: XML.Body) => { import XML.Decode._ make(list(pair(pair(key, info), list(key)))(body))(ord) } } final class Graph[Key, A] private(rep: SortedMap[Key, (A, (SortedSet[Key], SortedSet[Key]))]) { type Keys = SortedSet[Key] type Entry = (A, (Keys, Keys)) def ordering: Ordering[Key] = rep.ordering def empty_keys: Keys = SortedSet.empty[Key](ordering) /* graphs */ def is_empty: Boolean = rep.isEmpty def defined(x: Key): Boolean = rep.isDefinedAt(x) def domain: Set[Key] = rep.keySet def size: Int = rep.size def iterator: Iterator[(Key, Entry)] = rep.iterator def keys_iterator: Iterator[Key] = iterator.map(_._1) def keys: List[Key] = keys_iterator.toList def dest: List[((Key, A), List[Key])] = (for ((x, (i, (_, succs))) <- iterator) yield ((x, i), succs.toList)).toList override def toString: String = dest.map({ case ((x, _), ys) => x.toString + " -> " + ys.iterator.map(_.toString).mkString("{", ", ", "}") }) .mkString("Graph(", ", ", ")") private def get_entry(x: Key): Entry = rep.get(x) match { case Some(entry) => entry case None => throw new Graph.Undefined(x) } private def map_entry(x: Key, f: Entry => Entry): Graph[Key, A] = new Graph[Key, A](rep + (x -> f(get_entry(x)))) /* nodes */ def get_node(x: Key): A = get_entry(x)._1 def map_node(x: Key, f: A => A): Graph[Key, A] = map_entry(x, { case (i, ps) => (f(i), ps) }) /* reachability */ /*reachable nodes with length of longest path*/ def reachable_length( count: Key => Long, next: Key => Keys, 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 ((rs + (x -> length)) /: next(x))(reach(length + count(x))) - (Map.empty[Key, Long] /: xs)(reach(0)) + else next(x).foldLeft(rs + (x -> length))(reach(length + count(x))) + xs.foldLeft(Map.empty[Key, Long])(reach(0)) } def node_height(count: Key => Long): Map[Key, Long] = reachable_length(count, imm_preds, maximals) def node_depth(count: Key => Long): Map[Key, Long] = reachable_length(count, imm_succs, minimals) /*reachable nodes with size limit*/ def reachable_limit( limit: Long, count: Key => Long, next: Key => Keys, xs: List[Key]): Keys = { def reach(res: (Long, Keys), x: Key): (Long, Keys) = { val (n, rs) = res if (rs.contains(x)) res - else ((n + count(x), rs + x) /: next(x))(reach) + else next(x).foldLeft((n + count(x), rs + x))(reach) } @tailrec def reachs(size: Long, rs: Keys, rest: List[Key]): Keys = rest match { case Nil => rs case head :: tail => val (size1, rs1) = reach((size, rs), head) if (size > 0 && size1 > limit) rs else reachs(size1, rs1, tail) } reachs(0, empty_keys, xs) } /*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) = { val (rs, r_set) = reached if (r_set(x)) reached else { val (rs1, r_set1) = - if (rev) ((rs, r_set + x) /: next(x))({ case (b, a) => reach(a, b) }) + if (rev) next(x).foldLeft((rs, r_set + x)) { case (b, a) => reach(a, b) } else (next(x) :\ (rs, r_set + x))(reach) (x :: rs1, r_set1) } } 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) } - ((List.empty[List[Key]], empty_keys) /: xs)(reachs) + xs.foldLeft((List.empty[List[Key]], empty_keys))(reachs) } /*immediate*/ def imm_preds(x: Key): Keys = get_entry(x)._2._1 def imm_succs(x: Key): Keys = get_entry(x)._2._2 /*transitive*/ def all_preds_rev(xs: List[Key]): List[Key] = reachable(imm_preds, xs, rev = true)._1.flatten.reverse def all_preds(xs: List[Key]): List[Key] = reachable(imm_preds, xs)._1.flatten def all_succs(xs: List[Key]): List[Key] = reachable(imm_succs, xs)._1.flatten /*strongly connected components; see: David King and John Launchbury, "Structuring Depth First Search Algorithms in Haskell"*/ def strong_conn: List[List[Key]] = reachable(imm_preds, all_succs(keys))._1.filterNot(_.isEmpty).reverse /* minimal and maximal elements */ def minimals: List[Key] = - (List.empty[Key] /: rep) { - case (ms, (m, (_, (preds, _)))) => if (preds.isEmpty) m :: ms else ms } + rep.foldLeft(List.empty[Key]) { + case (ms, (m, (_, (preds, _)))) => if (preds.isEmpty) m :: ms else ms + } def maximals: List[Key] = - (List.empty[Key] /: rep) { - case (ms, (m, (_, (_, succs)))) => if (succs.isEmpty) m :: ms else ms } + rep.foldLeft(List.empty[Key]) { + case (ms, (m, (_, (_, succs)))) => if (succs.isEmpty) m :: ms else ms + } def is_minimal(x: Key): Boolean = imm_preds(x).isEmpty def is_maximal(x: Key): Boolean = imm_succs(x).isEmpty def is_isolated(x: Key): Boolean = is_minimal(x) && is_maximal(x) /* node operations */ 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)))) } def default_node(x: Key, info: A): Graph[Key, A] = if (defined(x)) this else new_node(x, info) private def del_adjacent(fst: Boolean, x: Key)(map: SortedMap[Key, Entry], y: Key) : SortedMap[Key, Entry] = map.get(y) match { case None => map case Some((i, (preds, succs))) => map + (y -> (i, if (fst) (preds - x, succs) else (preds, succs - x))) } def del_node(x: Key): Graph[Key, A] = { val (preds, succs) = get_entry(x)._2 new Graph[Key, A]( - (((rep - x) /: preds)(del_adjacent(false, x)) /: succs)(del_adjacent(true, x))) + succs.foldLeft(preds.foldLeft(rep - x)(del_adjacent(false, x)))(del_adjacent(true, x))) } def restrict(pred: Key => Boolean): Graph[Key, A] = - (this /: iterator){ case (graph, (x, _)) => if (!pred(x)) graph.del_node(x) else graph } + iterator.foldLeft(this) { case (graph, (x, _)) => if (!pred(x)) graph.del_node(x) else graph } def exclude(pred: Key => Boolean): Graph[Key, A] = restrict(name => !pred(name)) /* edge operations */ def edges_iterator: Iterator[(Key, Key)] = for { x <- keys_iterator; y <- imm_succs(x).iterator } yield (x, y) def is_edge(x: Key, y: Key): Boolean = defined(x) && defined(y) && imm_succs(x)(y) def add_edge(x: Key, y: Key): Graph[Key, A] = if (is_edge(x, y)) this else map_entry(y, { case (i, (preds, succs)) => (i, (preds + x, succs)) }). map_entry(x, { case (i, (preds, succs)) => (i, (preds, succs + y)) }) def del_edge(x: Key, y: Key): Graph[Key, A] = if (is_edge(x, y)) map_entry(y, { case (i, (preds, succs)) => (i, (preds - x, succs)) }). map_entry(x, { case (i, (preds, succs)) => (i, (preds, succs - y)) }) else this /* irreducible paths -- Hasse diagram */ 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 { case Nil => xs1 case x :: xs => if (!x_set(x) || x == z || path.contains(x) || xs.exists(red(x)) || xs1.exists(red(x))) irreds(xs, xs1) else irreds(xs, x :: xs1) } irreds(imm_preds(z).toList, Nil) } 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 - else (ps /: irreducible_preds(x_set, path, z))(paths(z :: path)) + else irreducible_preds(x_set, path, z).foldLeft(ps)(paths(z :: path)) if ((x == y) && !is_edge(x, x)) List(Nil) else paths(Nil)(Nil, y) } /* transitive closure and reduction */ 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) graph } - def transitive_closure: Graph[Key, A] = (this /: keys_iterator)(_.transitive_step(_)) + def transitive_closure: Graph[Key, A] = keys_iterator.foldLeft(this)(_.transitive_step(_)) 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") var graph = this for { (x, (_, (_, succs))) <- iterator y <- succs if trans.imm_preds(y).exists(z => trans.is_edge(x, z)) } graph = graph.del_edge(x, y) graph } /* maintain acyclic graphs */ def add_edge_acyclic(x: Key, y: Key): Graph[Key, A] = if (is_edge(x, y)) this else { irreducible_paths(y, x) match { case Nil => add_edge(x, y) case cycles => throw new Graph.Cycles(cycles.map(x :: _)) } } def add_deps_acyclic(y: Key, xs: List[Key]): Graph[Key, A] = - (this /: xs)(_.add_edge_acyclic(_, y)) + xs.foldLeft(this)(_.add_edge_acyclic(_, y)) def topological_order: List[Key] = all_succs(minimals) } diff --git a/src/Pure/General/graph_display.scala b/src/Pure/General/graph_display.scala --- a/src/Pure/General/graph_display.scala +++ b/src/Pure/General/graph_display.scala @@ -1,74 +1,75 @@ /* Title: Pure/General/graph_display.scala Author: Makarius Support for graph display. */ package isabelle object Graph_Display { /* graph entries */ type Entry = ((String, (String, XML.Body)), List[String]) // ident, name, content, parents /* graph structure */ object Node { val dummy: Node = 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 case ord => ord } } } sealed case class Node(name: String, ident: String) { def is_dummy: Boolean = this == Node.dummy override def toString: String = name } type Edge = (Node, Node) type Graph = isabelle.Graph[Node, XML.Body] val empty_graph: Graph = isabelle.Graph.empty(Node.Ordering) def build_graph(entries: List[Entry]): Graph = { val node = - (Map.empty[String, Node] /: entries) { + entries.foldLeft(Map.empty[String, Node]) { case (m, ((ident, (name, _)), _)) => m + (ident -> Node(name, ident)) } - (((empty_graph /: entries) { + entries.foldLeft( + entries.foldLeft(empty_graph) { case (g, ((ident, (_, content)), _)) => g.new_node(node(ident), content) - }) /: entries) { + }) { case (g1, ((ident, _), parents)) => - (g1 /: parents) { case (g2, parent) => g2.add_edge(node(parent), node(ident)) } + parents.foldLeft(g1) { case (g2, parent) => g2.add_edge(node(parent), node(ident)) } } } def decode_graph(body: XML.Body): Graph = build_graph( { import XML.Decode._ list(pair(pair(string, pair(string, x => x)), list(string)))(body) }) def make_graph[A]( graph: isabelle.Graph[String, A], isolated: Boolean = false, 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 build_graph(entries) } } diff --git a/src/Pure/General/linear_set.scala b/src/Pure/General/linear_set.scala --- a/src/Pure/General/linear_set.scala +++ b/src/Pure/General/linear_set.scala @@ -1,164 +1,164 @@ /* Title: Pure/General/linear_set.scala Author: Makarius Author: Fabian Immler, TU Munich Sets with canonical linear order, or immutable linked-lists. */ package isabelle import scala.collection.mutable import scala.collection.immutable.SetOps import scala.collection.{IterableFactory, IterableFactoryDefaults} 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]] - def from[A](entries: IterableOnce[A]): Linear_Set[A] = (empty[A] /: entries)(_.incl(_)) + def from[A](entries: IterableOnce[A]): Linear_Set[A] = entries.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 var res = empty[A] override def clear(): Unit = { res = empty[A] } override def addOne(elem: A): this.type = { res = res.incl(elem); this } override def result(): Linear_Set[A] = res } class Duplicate[A](x: A) extends Exception class Undefined[A](x: A) extends Exception class Next_Undefined[A](x: Option[A]) extends Exception } final class Linear_Set[A] private( 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] { /* relative addressing */ def next(elem: A): Option[A] = if (contains(elem)) nexts.get(elem) else throw new Linear_Set.Undefined(elem) def prev(elem: A): Option[A] = if (contains(elem)) prevs.get(elem) else throw new Linear_Set.Undefined(elem) def get_after(hook: Option[A]): Option[A] = hook match { case None => start case Some(elem) => next(elem) } def insert_after(hook: Option[A], elem: A): Linear_Set[A] = if (contains(elem)) throw new Linear_Set.Duplicate(elem) else hook match { case None => start match { case None => new Linear_Set[A](Some(elem), Some(elem), Map(), Map()) case Some(elem1) => new Linear_Set[A](Some(elem), end, nexts + (elem -> elem1), prevs + (elem1 -> elem)) } case Some(elem1) => if (!contains(elem1)) throw new Linear_Set.Undefined(elem1) else nexts.get(elem1) match { case None => new Linear_Set[A](start, Some(elem), nexts + (elem1 -> elem), prevs + (elem -> elem1)) case Some(elem2) => new Linear_Set[A](start, end, nexts + (elem1 -> elem) + (elem -> elem2), prevs + (elem2 -> elem) + (elem -> elem1)) } } def append_after(hook: Option[A], elems: Iterable[A]): Linear_Set[A] = // FIXME reverse fold - ((hook, this) /: elems) { + elems.foldLeft((hook, this)) { case ((last, set), elem) => (Some(elem), set.insert_after(last, elem)) }._2 def delete_after(hook: Option[A]): Linear_Set[A] = hook match { case None => start match { case None => throw new Linear_Set.Next_Undefined[A](None) case Some(elem1) => nexts.get(elem1) match { case None => empty case Some(elem2) => new Linear_Set[A](Some(elem2), end, nexts - elem1, prevs - elem2) } } case Some(elem1) => if (!contains(elem1)) throw new Linear_Set.Undefined(elem1) else nexts.get(elem1) match { case None => throw new Linear_Set.Next_Undefined(Some(elem1)) case Some(elem2) => nexts.get(elem2) match { case None => new Linear_Set[A](start, Some(elem1), nexts - elem1, prevs - elem2) case Some(elem3) => new Linear_Set[A](start, end, nexts - elem2 + (elem1 -> elem3), prevs - elem2 + (elem3 -> elem1)) } } } /* Set methods */ override def isEmpty: Boolean = start.isEmpty override def size: Int = if (isEmpty) 0 else nexts.size + 1 def contains(elem: A): Boolean = nonEmpty && (end.get == elem || nexts.isDefinedAt(elem)) private def make_iterator(from: Option[A]): Iterator[A] = new Iterator[A] { private var next_elem = from def hasNext: Boolean = next_elem.isDefined def next(): A = next_elem match { case Some(elem) => next_elem = nexts.get(elem) elem case None => Iterator.empty.next() } } override def iterator: Iterator[A] = make_iterator(start) def iterator(elem: A): Iterator[A] = if (contains(elem)) make_iterator(Some(elem)) else throw new Linear_Set.Undefined(elem) def iterator(from: A, to: A): Iterator[A] = if (contains(to)) nexts.get(to) match { case None => iterator(from) case Some(stop) => iterator(from).takeWhile(_ != stop) } else throw new Linear_Set.Undefined(to) def reverse: Linear_Set[A] = new Linear_Set(end, start, prevs, nexts) override def last: A = reverse.head def incl(elem: A): Linear_Set[A] = insert_after(end, elem) def excl(elem: A): Linear_Set[A] = delete_after(prev(elem)) override def iterableFactory: IterableFactory[Linear_Set] = Linear_Set override def toString: String = mkString("Linear_Set(", ", ", ")") } diff --git a/src/Pure/General/mercurial.scala b/src/Pure/General/mercurial.scala --- a/src/Pure/General/mercurial.scala +++ b/src/Pure/General/mercurial.scala @@ -1,337 +1,337 @@ /* Title: Pure/General/mercurial.scala Author: Makarius Support for Mercurial repositories, with local or remote repository clone and working directory (via ssh connection). */ package isabelle import java.io.{File => JFile} import scala.annotation.tailrec import scala.collection.mutable object Mercurial { type Graph = isabelle.Graph[String, Unit] /* command-line syntax */ def optional(s: String, prefix: String = ""): String = if (s == "") "" else " " + prefix + " " + Bash.string(s) def opt_flag(flag: String, b: Boolean): String = if (b) " " + flag else "" def opt_rev(s: String): String = optional(s, "--rev") def opt_template(s: String): String = optional(s, "--template") /* repository access */ def is_repository(root: Path, ssh: SSH.System = SSH.Local): Boolean = ssh.is_dir(root + Path.explode(".hg")) && new Repository(root, ssh).command("root").ok 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] = { @tailrec def find(root: Path): Option[Repository] = if (is_repository(root, ssh)) Some(repository(root, ssh = ssh)) else if (root.is_root) None else find(root + Path.parent) find(ssh.expand_path(start)) } 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 hg } def init_repository(root: Path, ssh: SSH.System = SSH.Local): Repository = make_repository(root, "init", ssh.bash_path(root), ssh = ssh) def clone_repository(source: String, root: Path, rev: String = "", options: String = "", ssh: SSH.System = SSH.Local): Repository = 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 = { 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) { hg => val root: Path = ssh.expand_path(root_path) def root_url: String = ssh.hg_url + root.implode override def toString: String = ssh.hg_url + root.implode def command(name: String, args: String = "", options: String = "", repository: Boolean = true): Process_Result = { val cmdline = "export HGPLAIN=\n\"${HG:-hg}\" --config " + Bash.string("defaults." + name + "=") + (if (repository) " --repository " + ssh.bash_path(root) else "") + " --noninteractive " + name + " " + options + " " + args ssh.execute(cmdline) } def add(files: List[Path]): Unit = hg.command("add", files.map(ssh.bash_path).mkString(" ")) def archive(target: String, rev: String = "", options: String = ""): Unit = hg.command("archive", opt_rev(rev) + " " + Bash.string(target), options).check def heads(template: String = "{node|short}\n", options: String = ""): List[String] = hg.command("heads", opt_template(template), options).check.out_lines def identify(rev: String = "tip", options: String = ""): String = hg.command("id", opt_rev(rev), options).check.out_lines.headOption getOrElse "" def id(rev: String = "tip"): String = identify(rev, options = "-i") def paths(args: String = "", options: String = ""): List[String] = hg.command("paths", args = args, options = options).check.out_lines def manifest(rev: String = "", options: String = ""): List[String] = hg.command("manifest", opt_rev(rev), options).check.out_lines def log(rev: String = "", template: String = "", options: String = ""): String = hg.command("log", opt_rev(rev) + opt_template(template), options).check.out def parent(): String = log(rev = "p1()", template = "{node|short}") def push( 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) } def pull(remote: String = "", rev: String = "", options: String = ""): Unit = hg.command("pull", opt_rev(rev) + optional(remote), options).check def update( 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 } def known_files(): List[String] = hg.command("status", options = "--modified --added --clean --no-status").check.out_lines 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""") - (Graph.string[Unit] /: split_lines(log_result)) { + split_lines(log_result).foldLeft(Graph.string[Unit]) { case (graph, Node(x, y, z)) => val deps = List(y, z).filterNot(s => s.forall(_ == '0')) - val graph1 = (graph /: (x :: deps))(_.default_node(_, ())) - (graph1 /: deps)({ case (g, dep) => g.add_edge(dep, x) }) + val graph1 = (x :: deps).foldLeft(graph)(_.default_node(_, ())) + deps.foldLeft(graph1)({ case (g, dep) => g.add_edge(dep, x) }) case (graph, _) => graph } } } /* check files */ 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 = { paths match { case path :: rest => find_repository(path, ssh) match { case None => outside += path; check(rest) case Some(hg) => val known = hg.known_files().iterator.map(name => (hg.root + Path.explode(name)).canonical_file).toSet if (!known(path.canonical_file)) unknown += path check(rest.filterNot(p => known(p.canonical_file))) } case Nil => } } check(files) (outside.toList, unknown.toList) } /* setup remote vs. local repository */ 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 = { File.write(hgrc, cat_lines(lines)) true } val edited = hgrc.is_file && { val lines = split_lines(File.read(hgrc)) lines.count(header) == 1 && { if (local_hg.paths(options = "-q").contains(path_name)) { val old_source = local_hg.paths(args = path_name).head val old_entry = path_name + ".old = " + old_source val lines1 = lines.map { case Entry(a, b) if a == path_name && b == old_source => new_entry + "\n" + old_entry case line => line } lines != lines1 && commit(lines1) } else { val prefix = lines.takeWhile(line => !header(line)) val n = prefix.length commit(prefix ::: List(lines(n), new_entry) ::: lines.drop(n + 1)) } } } if (!edited) File.append(hgrc, "\n[paths]\n" + new_entry + "\n") } val default_path_name = "default" def hg_setup( remote: String, local_path: Path, remote_name: String = "", path_name: String = default_path_name, remote_exists: Boolean = false, progress: Progress = new Progress): Unit = { /* local repository */ Isabelle_System.make_directory(local_path) val repos_name = proper_string(remote_name) getOrElse local_path.absolute.base.implode val local_hg = if (is_repository(local_path)) repository(local_path) else init_repository(local_path) progress.echo("Local repository " + local_hg.root.absolute) /* remote repository */ val remote_url = remote match { case _ if remote.startsWith("ssh://") => val ssh_url = remote + "/" + repos_name if (!remote_exists) { try { local_hg.command("init", ssh_url, repository = false).check } catch { case ERROR(msg) => progress.echo_warning(msg) } } ssh_url case SSH.Target(user, host) => val phabricator = Phabricator.API(user, host) var repos = phabricator.get_repositories().find(r => r.short_name == repos_name) getOrElse { if (remote_exists) { error("Remote repository " + quote(phabricator.hg_url + "/source/" + repos_name) + " expected to exist") } else phabricator.create_repository(repos_name, short_name = repos_name) } while (repos.importing) { progress.echo("Awaiting remote repository ...") Time.seconds(0.5).sleep repos = phabricator.the_repository(repos.phid) } repos.ssh_url case _ => error("Malformed remote specification " + quote(remote)) } progress.echo("Remote repository " + quote(remote_url)) /* synchronize local and remote state */ progress.echo("Synchronizing ...") edit_hgrc(local_hg, path_name, remote_url) local_hg.pull(options = "-u") local_hg.push(remote = remote_url) } val isabelle_tool = Isabelle_Tool("hg_setup", "setup remote vs. local Mercurial repository", Scala_Project.here, args => { var remote_name = "" var path_name = default_path_name var remote_exists = false val getopts = Getopts(""" Usage: isabelle hg_setup [OPTIONS] REMOTE LOCAL_DIR Options are: -n NAME remote repository name (default: base name of LOCAL_DIR) -p PATH Mercurial path name (default: """ + quote(default_path_name) + """) -r assume that remote repository already exists Setup a remote vs. local Mercurial repository: REMOTE either refers to a Phabricator server "user@host" or SSH file server "ssh://user@host/path". """, "n:" -> (arg => remote_name = arg), "p:" -> (arg => path_name = arg), "r" -> (_ => remote_exists = true)) val more_args = getopts(args) val (remote, local_path) = more_args match { case List(arg1, arg2) => (arg1, Path.explode(arg2)) case _ => getopts.usage() } val progress = new Console_Progress hg_setup(remote, local_path, remote_name = remote_name, path_name = path_name, remote_exists = remote_exists, progress = progress) }) } diff --git a/src/Pure/General/multi_map.scala b/src/Pure/General/multi_map.scala --- a/src/Pure/General/multi_map.scala +++ b/src/Pure/General/multi_map.scala @@ -1,91 +1,91 @@ /* Title: Pure/General/multi_map.scala Author: Makarius Maps with multiple entries per key. */ package isabelle import scala.collection.mutable import scala.collection.{IterableFactory, MapFactory, MapFactoryDefaults} import scala.collection.immutable.{Iterable, MapOps} 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]] def from[A, B](entries: IterableOnce[(A, B)]): Multi_Map[A, B] = - (empty[A, B] /: entries)({ case (m, (a, b)) => m.insert(a, b) }) + entries.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 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 } override def result(): Multi_Map[A, B] = res } } 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] { /* 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] = { 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] = { val bs = get_list(a) if (bs.contains(b)) { bs.filterNot(_ == b) match { case Nil => new Multi_Map(rep - a) case bs1 => new Multi_Map(rep + (a -> bs1)) } } else this } def ++[B1 >: B] (other: Multi_Map[A, B1]): Multi_Map[A, B1] = if (this eq other) this else if (isEmpty) other else - (this.asInstanceOf[Multi_Map[A, B1]] /: other.rep.iterator) { + other.rep.iterator.foldLeft(this.asInstanceOf[Multi_Map[A, B1]]) { case (m1, (a, bs)) => (bs :\ m1) { case (b, m2) => m2.insert(a, b) } } /* Map operations */ override def empty: Multi_Map[A, Nothing] = Multi_Map.empty override def isEmpty: Boolean = rep.isEmpty override def keySet: Set[A] = rep.keySet override def iterator: Iterator[(A, B)] = for ((a, bs) <- rep.iterator; b <- bs.iterator) yield (a, b) def get(a: A): Option[B] = get_list(a).headOption override def updated[B1 >: B](a: A, b: B1): Multi_Map[A, B1] = insert(a, b) override def removed(a: A): Multi_Map[A, B] = if (rep.isDefinedAt(a)) new Multi_Map(rep - a) else this override def mapFactory: MapFactory[Multi_Map] = Multi_Map override def toString: String = mkString("Multi_Map(", ", ", ")") } diff --git a/src/Pure/General/path.scala b/src/Pure/General/path.scala --- a/src/Pure/General/path.scala +++ b/src/Pure/General/path.scala @@ -1,306 +1,306 @@ /* Title: Pure/General/path.scala Author: Makarius Algebra of file-system paths: basic POSIX notation, extended by named roots (e.g. //foo) and variables (e.g. $BAR). */ package isabelle import java.io.{File => JFile} import scala.util.matching.Regex object Path { /* path elements */ sealed abstract class Elem private case class Root(name: String) extends Elem private case class Basic(name: String) extends Elem private case class Variable(name: String) extends Elem private case object Parent extends Elem private def err_elem(msg: String, s: String): Nothing = error(msg + " path element " + quote(s)) private val illegal_elem = Set("", "~", "~~", ".", "..") private val illegal_char = "/\\$:\"'<>|?*" private def check_elem(s: String): String = if (illegal_elem.contains(s)) err_elem("Illegal", s) else { for (c <- s) { if (c.toInt < 32) err_elem("Illegal control character " + c.toInt + " in", s) if (illegal_char.contains(c)) err_elem("Illegal character " + quote(c.toString) + " in", s) } s } private def root_elem(s: String): Elem = Root(check_elem(s)) private def basic_elem(s: String): Elem = Basic(check_elem(s)) private def variable_elem(s: String): Elem = Variable(check_elem(s)) private def apply_elem(y: Elem, xs: List[Elem]): List[Elem] = (y, xs) match { case (Root(_), _) => List(y) case (Parent, Root(_) :: _) => xs case (Parent, Basic(_) :: rest) => rest case _ => y :: xs } private def norm_elems(elems: List[Elem]): List[Elem] = (elems :\ (Nil: List[Elem]))(apply_elem) private def implode_elem(elem: Elem, short: Boolean): String = elem match { case Root("") => "" case Root(s) => "//" + s case Basic(s) => s case Variable("USER_HOME") if short => "~" case Variable("ISABELLE_HOME") if short => "~~" case Variable(s) => "$" + s case Parent => ".." } private def squash_elem(elem: Elem): String = elem match { case Root("") => "ROOT" case Root(s) => "SERVER_" + s case Basic(s) => s case Variable(s) => s case Parent => "PARENT" } /* path constructors */ val current: Path = new Path(Nil) val root: Path = new Path(List(Root(""))) def named_root(s: String): Path = new Path(List(root_elem(s))) def make(elems: List[String]): Path = new Path(elems.reverse.map(basic_elem)) def basic(s: String): Path = new Path(List(basic_elem(s))) def variable(s: String): Path = new Path(List(variable_elem(s))) val parent: Path = new Path(List(Parent)) /* explode */ def explode(str: String): Path = { def explode_elem(s: String): Elem = try { if (s == "..") Parent else if (s == "~") Variable("USER_HOME") else if (s == "~~") Variable("ISABELLE_HOME") else if (s.startsWith("$")) variable_elem(s.substring(1)) else basic_elem(s) } catch { case ERROR(msg) => cat_error(msg, "The error(s) above occurred in " + quote(str)) } val ss = space_explode('/', str) val r = ss.takeWhile(_.isEmpty).length val es = ss.dropWhile(_.isEmpty) val (roots, raw_elems) = if (r == 0) (Nil, es) else if (r == 1) (List(Root("")), es) else if (es.isEmpty) (List(Root("")), Nil) else (List(root_elem(es.head)), es.tail) val elems = raw_elems.filterNot(s => s.isEmpty || s == ".").map(explode_elem) new Path(norm_elems(elems reverse_::: roots)) } def is_wellformed(str: String): Boolean = try { explode(str); true } catch { case ERROR(_) => false } def is_valid(str: String): Boolean = try { explode(str).expand; true } catch { case ERROR(_) => false } def split(str: String): List[Path] = space_explode(':', str).filterNot(_.isEmpty).map(explode) /* encode */ val encode: XML.Encode.T[Path] = (path => XML.Encode.string(path.implode)) /* reserved names */ private val reserved_windows: Set[String] = Set("CON", "PRN", "AUX", "NUL", "COM1", "COM2", "COM3", "COM4", "COM5", "COM6", "COM7", "COM8", "COM9", "LPT1", "LPT2", "LPT3", "LPT4", "LPT5", "LPT6", "LPT7", "LPT8", "LPT9") def is_reserved(name: String): Boolean = Long_Name.explode(name).exists(a => reserved_windows.contains(Word.uppercase(a))) /* case-insensitive names */ def check_case_insensitive(paths: List[Path]): Unit = { val table = - (Multi_Map.empty[String, String] /: paths)({ case (tab, path) => + paths.foldLeft(Multi_Map.empty[String, String]) { case (tab, path) => val name = path.expand.implode tab.insert(Word.lowercase(name), name) - }) + } val collisions = (for { (_, coll) <- table.iterator_list if coll.length > 1 } yield coll).toList.flatten if (collisions.nonEmpty) { error(("Collision of file names due case-insensitivity:" :: collisions).mkString("\n ")) } } } 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 { case other: Path => elems == other.elems case _ => false } def is_current: Boolean = elems.isEmpty def is_absolute: Boolean = elems.nonEmpty && elems.last.isInstanceOf[Path.Root] def is_root: Boolean = elems match { case List(Path.Root(_)) => true case _ => false } def is_basic: Boolean = elems match { case List(Path.Basic(_)) => true case _ => false } def starts_basic: Boolean = elems.nonEmpty && elems.last.isInstanceOf[Path.Basic] def +(other: Path): Path = new Path((other.elems :\ elems)(Path.apply_elem)) /* implode */ private def gen_implode(short: Boolean): String = elems match { case Nil => "." case List(Path.Root("")) => "/" case _ => elems.map(Path.implode_elem(_, short)).reverse.mkString("/") } def implode: String = gen_implode(false) def implode_short: String = gen_implode(true) override def toString: String = quote(implode) /* base element */ private def split_path: (Path, String) = elems match { case Path.Basic(s) :: xs => (new Path(xs), s) case _ => error("Cannot split path into dir/base: " + toString) } def dir: Path = split_path._1 def base: Path = new Path(List(Path.Basic(split_path._2))) def ext(e: String): Path = if (e == "") this else { val (prfx, s) = split_path prfx + Path.basic(s + "." + e) } def xz: Path = ext("xz") def html: Path = ext("html") def tex: Path = ext("tex") def pdf: Path = ext("pdf") def thy: Path = ext("thy") def backup: Path = { val (prfx, s) = split_path prfx + Path.basic(s + "~") } def backup2: Path = { val (prfx, s) = split_path prfx + Path.basic(s + "~~") } def platform_exe: Path = if (Platform.is_windows) ext("exe") else this private val Ext = new Regex("(.*)\\.([^.]*)") def split_ext: (Path, String) = { val (prefix, base) = split_path base match { case Ext(b, e) => (prefix + Path.basic(b), e) case _ => (prefix + Path.basic(base), "") } } def drop_ext: Path = split_ext._1 def get_ext: String = split_ext._2 def squash: Path = new Path(elems.map(elem => Path.Basic(Path.squash_elem(elem)))) /* expand */ def expand_env(env: Map[String, String]): Path = { def eval(elem: Path.Elem): List[Path.Elem] = elem match { case Path.Variable(s) => val path = Path.explode(Isabelle_System.getenv_strict(s, env)) if (path.elems.exists(_.isInstanceOf[Path.Variable])) error("Illegal path variable nesting: " + s + "=" + path.toString) else path.elems case x => List(x) } new Path(Path.norm_elems(elems.flatMap(eval))) } def expand: Path = expand_env(Isabelle_System.settings()) def file_name: String = expand.base.implode /* implode wrt. given directories */ def implode_symbolic: String = { val directories = Library.space_explode(':', Isabelle_System.getenv("ISABELLE_DIRECTORIES")).reverse val full_name = expand.implode directories.view.flatMap(a => try { val b = Path.explode(a).expand.implode if (full_name == b) Some(a) else { Library.try_unprefix(b + "/", full_name) match { case Some(name) => Some(a + "/" + name) case None => None } } } catch { case ERROR(_) => None }).headOption.getOrElse(implode) } def position: Position.T = Position.File(implode_symbolic) /* platform files */ def file: JFile = File.platform_file(this) def is_file: Boolean = file.isFile def is_dir: Boolean = file.isDirectory def absolute_file: JFile = File.absolute(file) def canonical_file: JFile = File.canonical(file) def absolute: Path = File.path(absolute_file) def canonical: Path = File.path(canonical_file) } diff --git a/src/Pure/General/pretty.scala b/src/Pure/General/pretty.scala --- a/src/Pure/General/pretty.scala +++ b/src/Pure/General/pretty.scala @@ -1,189 +1,189 @@ /* Title: Pure/General/pretty.scala Author: Makarius Generic pretty printing module. */ package isabelle import scala.annotation.tailrec object Pretty { /* XML constructors */ val space: XML.Body = List(XML.Text(Symbol.space)) def spaces(n: Int): XML.Body = if (n == 0) Nil else if (n == 1) space else List(XML.Text(Symbol.spaces(n))) def block(consistent: Boolean, indent: Int, body: XML.Body): XML.Tree = XML.Elem(Markup.Block(consistent, indent), body) def block(indent: Int, body: XML.Body): XML.Tree = block(false, indent, body) def block(body: XML.Body): XML.Tree = block(2, body) def brk(width: Int, indent: Int = 0): XML.Tree = XML.Elem(Markup.Break(width, indent), spaces(width)) val fbrk: XML.Tree = XML.newline def fbreaks(ts: List[XML.Tree]): XML.Body = Library.separate(fbrk, ts) val Separator: XML.Body = List(XML.elem(Markup.SEPARATOR, space), fbrk) def separate(ts: List[XML.Tree]): XML.Body = Library.separate(Separator, ts.map(List(_))).flatten /* text metric -- standardized to width of space */ abstract class Metric { val unit: Double def apply(s: String): Double } object Default_Metric extends Metric { val unit = 1.0 def apply(s: String): Double = s.length.toDouble } /* markup trees with physical blocks and breaks */ private def force_nat(i: Int): Int = i max 0 private sealed abstract class Tree { def length: Double } 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 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 = { val indent1 = force_nat(indent) @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 = ((0.0 /: line) { case (l, t) => l + t.length }) max len + val len1 = (line.foldLeft(0.0) { case (l, t) => l + t.length }) max len (rest: @unchecked) match { case Break(true, _, ind) :: rest1 => body_length(Break(false, indent1 + ind, 0) :: rest1, len1) case Nil => len1 } } Block(markup, consistent, indent1, body, body_length(body, 0.0)) } /* formatted output */ 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) def blanks(wd: Int): Text = string(Symbol.spaces(wd), wd.toDouble) def content: XML.Body = tx.reverse } private def break_dist(trees: List[Tree], after: Double): Double = trees match { case (_: Break) :: _ => 0.0 case t :: ts => t.length + break_dist(ts, after) case Nil => after } private def force_break(tree: Tree): Tree = tree match { case Break(false, wd, ind) => Break(true, wd, ind) case _ => tree } private def force_all(trees: List[Tree]): List[Tree] = trees.map(force_break) private def force_next(trees: List[Tree]): List[Tree] = trees match { case Nil => Nil case (t: Break) :: ts => force_break(t) :: ts case t :: ts => t :: force_next(ts) } val default_margin: Double = 76.0 val default_breakgain: Double = default_margin / 20 def formatted(input: XML.Body, margin: Double = default_margin, breakgain: Double = default_breakgain, metric: Metric = Default_Metric): XML.Body = { val emergencypos = (margin / 2).round.toInt def make_tree(inp: XML.Body): List[Tree] = inp flatMap { case XML.Wrapped_Elem(markup, body1, body2) => List(make_block(Some(markup, Some(body1)), false, 0, make_tree(body2))) case XML.Elem(markup, body) => markup match { case Markup.Block(consistent, indent) => List(make_block(None, consistent, indent, make_tree(body))) case Markup.Break(width, indent) => List(Break(false, force_nat(width), force_nat(indent))) case Markup(Markup.ITEM, _) => List(make_block(None, false, 2, make_tree(XML.elem(Markup.BULLET, space) :: space ::: body))) case _ => List(make_block(Some((markup, None)), false, 0, make_tree(body))) } case XML.Text(text) => Library.separate(FBreak, split_lines(text).map(s => Str(s, metric(s)))) } def format(trees: List[Tree], blockin: Int, after: Double, text: Text): Text = trees match { case Nil => text case Block(markup, consistent, indent, body, blen) :: ts => val pos1 = (text.pos + indent).ceil.toInt val pos2 = pos1 % emergencypos val blockin1 = if (pos1 < emergencypos) pos1 else pos2 val d = break_dist(ts, after) val body1 = if (consistent && text.pos + blen > margin - d) force_all(body) else body val btext = markup match { case None => format(body1, blockin1, d, text) case Some((m, markup_body)) => val btext0 = format(body1, blockin1, d, text.copy(tx = Nil)) val elem = markup_body match { case None => XML.Elem(m, btext0.content) case Some(b) => XML.Wrapped_Elem(m, b, btext0.content) } btext0.copy(tx = elem :: text.tx) } val ts1 = if (text.nl < btext.nl) force_next(ts) else ts format(ts1, blockin, after, btext) case Break(force, wd, ind) :: ts => if (!force && text.pos + wd <= ((margin - break_dist(ts, after)) max (blockin + breakgain))) format(ts, blockin, after, text.blanks(wd)) else format(ts, blockin, after, text.newline.blanks(blockin + ind)) case Str(s, len) :: ts => format(ts, blockin, after, text.string(s, len)) } format(make_tree(input), 0, 0.0, Text()).content } def string_of(input: XML.Body, margin: Double = default_margin, breakgain: Double = default_breakgain, metric: Metric = Default_Metric): String = XML.content(formatted(input, margin = margin, breakgain = breakgain, metric = metric)) } diff --git a/src/Pure/General/scan.scala b/src/Pure/General/scan.scala --- a/src/Pure/General/scan.scala +++ b/src/Pure/General/scan.scala @@ -1,511 +1,511 @@ /* Title: Pure/General/scan.scala Author: Makarius Efficient scanning of keywords and tokens. */ package isabelle import scala.annotation.tailrec import scala.collection.IndexedSeq import scala.util.matching.Regex import scala.util.parsing.input.{OffsetPosition, Position => InputPosition, Reader, CharSequenceReader, PagedSeq} import scala.util.parsing.combinator.RegexParsers import java.io.{File => JFile, BufferedInputStream, FileInputStream, InputStream} import java.net.URL object Scan { /** context of partial line-oriented scans **/ abstract class Line_Context case object Finished extends Line_Context case class Quoted(quote: String) extends Line_Context case object Verbatim extends Line_Context case class Cartouche(depth: Int) extends Line_Context case class Comment_Prefix(symbol: Symbol.Symbol) extends Line_Context case class Cartouche_Comment(depth: Int) extends Line_Context case class Comment(depth: Int) extends Line_Context /** parser combinators **/ object Parsers extends Parsers trait Parsers extends RegexParsers { override val whiteSpace: Regex = "".r /* optional termination */ def opt_term[T](p: => Parser[T]): Parser[Option[T]] = p ^^ (x => Some(x)) | """\z""".r ^^ (_ => None) /* repeated symbols */ def repeated(pred: Symbol.Symbol => Boolean, min_count: Int, max_count: Int): Parser[String] = new Parser[String] { def apply(in: Input) = { val start = in.offset val end = in.source.length val matcher = new Symbol.Matcher(in.source) var i = start var count = 0 var finished = false while (!finished && i < end && count < max_count) { val n = matcher(i, end) val sym = in.source.subSequence(i, i + n).toString if (pred(sym)) { i += n; count += 1 } else finished = true } if (count < min_count) Failure("bad input", in) else Success(in.source.subSequence(start, i).toString, in.drop(i - start)) } }.named("repeated") def one(pred: Symbol.Symbol => Boolean): Parser[String] = repeated(pred, 1, 1) def maybe(pred: Symbol.Symbol => Boolean): Parser[String] = repeated(pred, 0, 1) def many(pred: Symbol.Symbol => Boolean): Parser[String] = repeated(pred, 0, Integer.MAX_VALUE) def many1(pred: Symbol.Symbol => Boolean): Parser[String] = repeated(pred, 1, Integer.MAX_VALUE) /* character */ def character(pred: Char => Boolean): Symbol.Symbol => Boolean = (s: Symbol. Symbol) => s.length == 1 && pred(s.charAt(0)) /* quoted strings */ 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] = { quote ~ quoted_body(quote) ~ quote ^^ { case x ~ y ~ z => x + y + z } }.named("quoted") 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(_ == '\\')) { val content = rep(many1(sym => sym != quote && sym != "\\") | "\\" ~> (quote | "\\" | """\d\d\d""".r ^^ { case x => x.toInt.toChar.toString })) parseAll(content ^^ (_.mkString), body).get } else body } def quoted_line(quote: Symbol.Symbol, ctxt: Line_Context): Parser[(String, Line_Context)] = { ctxt match { case Finished => quote ~ quoted_body(quote) ~ opt_term(quote) ^^ { case x ~ y ~ Some(z) => (x + y + z, Finished) case x ~ y ~ None => (x + y, Quoted(quote)) } case Quoted(q) if q == quote => quoted_body(quote) ~ opt_term(quote) ^^ { case x ~ Some(y) => (x + y, Finished) case x ~ None => (x, ctxt) } case _ => failure("") } }.named("quoted_line") def recover_quoted(quote: Symbol.Symbol): Parser[String] = quote ~ quoted_body(quote) ^^ { case x ~ y => x + y } /* verbatim text */ private def verbatim_body: Parser[String] = rep(many1(sym => sym != "*") | """\*(?!\})""".r) ^^ (_.mkString) def verbatim: Parser[String] = { "{*" ~ verbatim_body ~ "*}" ^^ { case x ~ y ~ z => x + y + z } }.named("verbatim") def verbatim_content(source: String): String = { require(parseAll(verbatim, source).successful, "no verbatim text") source.substring(2, source.length - 2) } def verbatim_line(ctxt: Line_Context): Parser[(String, Line_Context)] = { ctxt match { case Finished => "{*" ~ verbatim_body ~ opt_term("*}") ^^ { case x ~ y ~ Some(z) => (x + y + z, Finished) case x ~ y ~ None => (x + y, Verbatim) } case Verbatim => verbatim_body ~ opt_term("*}") ^^ { case x ~ Some(y) => (x + y, Finished) case x ~ None => (x, Verbatim) } case _ => failure("") } }.named("verbatim_line") val recover_verbatim: Parser[String] = "{*" ~ verbatim_body ^^ { case x ~ y => x + y } /* nested text cartouches */ def cartouche_depth(depth: Int): Parser[(String, Int)] = new Parser[(String, Int)] { require(depth >= 0, "bad cartouche depth") def apply(in: Input) = { val start = in.offset val end = in.source.length val matcher = new Symbol.Matcher(in.source) var i = start var d = depth var finished = false while (!finished && i < end) { val n = matcher(i, end) val sym = in.source.subSequence(i, i + n).toString if (Symbol.is_open(sym)) { i += n; d += 1 } else if (Symbol.is_close(sym) && d > 0) { i += n; d -= 1; if (d == 0) finished = true } else if (d > 0) i += n else finished = true } if (i == start) Failure("bad input", in) else Success((in.source.subSequence(start, i).toString, d), in.drop(i - start)) } }.named("cartouche_depth") 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_context(d: Int): Line_Context = if (d == 0) Finished else Cartouche(d) ctxt match { case Finished => cartouche_depth(0) ^^ { case (c, d) => (c, cartouche_context(d)) } case Cartouche(depth) => cartouche_depth(depth) ^^ { case (c, d) => (c, cartouche_context(d)) } case _ => failure("") } } val recover_cartouche: Parser[String] = cartouche_depth(0) ^^ (_._1) 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 Library.try_unprefix(Symbol.open, source) getOrElse err() Library.try_unsuffix(Symbol.close_decoded, source1) orElse Library.try_unsuffix(Symbol.close, source1) getOrElse err() } /* nested comments */ 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) = { var rest = in def try_parse[A](p: Parser[A]): Boolean = { parse(p ^^^ (()), rest) match { case Success(_, next) => { rest = next; true } case _ => false } } var d = depth var finished = false while (!finished) { if (try_parse("(*")) d += 1 else if (d > 0 && try_parse("*)")) { d -= 1; if (d == 0) finished = true } else if (d == 0 || !try_parse(comment_text)) finished = true } if (in.offset < rest.offset) Success((in.source.subSequence(in.offset, rest.offset).toString, d), rest) else Failure("comment expected", in) } }.named("comment_depth") def comment: Parser[String] = comment_depth(0) ^? { case (x, d) if d == 0 => x } def comment_line(ctxt: Line_Context): Parser[(String, Line_Context)] = { val depth = ctxt match { case Finished => 0 case Comment(d) => d case _ => -1 } if (depth >= 0) comment_depth(depth) ^^ { case (x, 0) => (x, Finished) case (x, d) => (x, Comment(d)) } else failure("") } val recover_comment: Parser[String] = comment_depth(0) ^^ (_._1) def comment_content(source: String): String = { require(parseAll(comment, source).successful, "no comment") source.substring(2, source.length - 2) } /* keyword */ 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)) } }.named("keyword") } /** Lexicon -- position tree **/ object Lexicon { /* representation */ private sealed case class Tree(branches: Map[Char, (String, Tree)]) private val empty_tree = Tree(Map()) val empty: Lexicon = new Lexicon(empty_tree) def apply(elems: String*): Lexicon = empty ++ elems } final class Lexicon private(rep: Lexicon.Tree) { /* auxiliary operations */ private def dest(tree: Lexicon.Tree, result: List[String]): List[String] = - (result /: tree.branches.toList) ((res, entry) => - entry match { case (_, (s, tr)) => - if (s.isEmpty) dest(tr, res) else dest(tr, s :: res) }) + tree.branches.toList.foldLeft(result) { + case (res, (_, (s, tr))) => if (s.isEmpty) dest(tr, res) else dest(tr, s :: res) + } 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)] = { if (i < len) { tree.branches.get(str.charAt(i)) match { case Some((s, tr)) => look(tr, s.nonEmpty, i + 1) case None => None } } else Some(tip, tree) } look(rep, false, 0) } def completions(str: CharSequence): List[String] = lookup(str) match { case Some((true, tree)) => dest(tree, List(str.toString)) case Some((false, tree)) => dest(tree, Nil) case None => Nil } /* pseudo Set methods */ def raw_iterator: Iterator[String] = dest(rep, Nil).iterator def iterator: Iterator[String] = dest(rep, Nil).sorted.iterator override def toString: String = iterator.mkString("Lexicon(", ", ", ")") def is_empty: Boolean = rep.branches.isEmpty def contains(elem: String): Boolean = lookup(elem) match { case Some((tip, _)) => tip case _ => false } /* build lexicon */ def + (elem: String): Lexicon = if (contains(elem)) this else { val len = elem.length def extend(tree: Lexicon.Tree, i: Int): Lexicon.Tree = if (i < len) { val c = elem.charAt(i) val end = (i + 1 == len) tree.branches.get(c) match { case Some((s, tr)) => Lexicon.Tree(tree.branches + (c -> (if (end) elem else s, extend(tr, i + 1)))) case None => Lexicon.Tree(tree.branches + (c -> (if (end) elem else "", extend(Lexicon.empty_tree, i + 1)))) } } else tree new Lexicon(extend(rep, 0)) } - def ++ (elems: IterableOnce[String]): Lexicon = (this /: elems)(_ + _) + def ++ (elems: IterableOnce[String]): Lexicon = elems.foldLeft(this)(_ + _) def ++ (other: Lexicon): Lexicon = if (this eq other) this else if (is_empty) other else this ++ other.raw_iterator def -- (remove: Iterable[String]): Lexicon = if (remove.exists(contains)) Lexicon.empty ++ iterator.filterNot(a => remove.exists(b => a == b)) else this /* scan */ 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 = { 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) case None => result } } else result } scan_tree(rep, "", 0) } } /** read stream without decoding: efficient length operation **/ 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 def length: Int = end - start // avoid expensive seq.length def subSequence(i: Int, j: Int): CharSequence = if (0 <= i && i <= j && j <= length) new Restricted_Seq(seq, start + i, start + j) else throw new IndexOutOfBoundsException override def toString: String = { val buf = new StringBuilder(length) for (offset <- start until end) buf.append(seq(offset)) buf.toString } } abstract class Byte_Reader extends Reader[Char] with AutoCloseable private def make_byte_reader(stream: InputStream, stream_length: Int): Byte_Reader = { val buffered_stream = new BufferedInputStream(stream) val seq = new PagedSeq( (buf: Array[Char], offset: Int, length: Int) => { var i = 0 var c = 0 var eof = false while (!eof && i < length) { c = buffered_stream.read if (c == -1) eof = true else { buf(offset + i) = c.toChar; i += 1 } } if (i > 0) i else -1 }) val restricted_seq = new Restricted_Seq(seq, 0, stream_length) 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 def pos: InputPosition = new OffsetPosition(source, offset) def atEnd: Boolean = !seq.isDefinedAt(offset) override def drop(n: Int): Paged_Reader = new Paged_Reader(offset + n) def close: Unit = buffered_stream.close } new Paged_Reader(0) } def byte_reader(file: JFile): Byte_Reader = make_byte_reader(new FileInputStream(file), file.length.toInt) def byte_reader(url: URL): Byte_Reader = { val connection = url.openConnection val stream = connection.getInputStream val stream_length = connection.getContentLength make_byte_reader(stream, stream_length) } def reader_is_utf8(reader: Reader[Char]): Boolean = reader.isInstanceOf[Byte_Reader] def reader_decode_utf8(is_utf8: Boolean, s: String): String = if (is_utf8) UTF8.decode_permissive(s) else s def reader_decode_utf8(reader: Reader[Char], s: String): String = reader_decode_utf8(reader_is_utf8(reader), s) /* plain text reader */ def char_reader(text: CharSequence): CharSequenceReader = new CharSequenceReader(text) } diff --git a/src/Pure/General/symbol.scala b/src/Pure/General/symbol.scala --- a/src/Pure/General/symbol.scala +++ b/src/Pure/General/symbol.scala @@ -1,669 +1,669 @@ /* Title: Pure/General/symbol.scala Author: Makarius Isabelle text symbols. */ package isabelle import scala.collection.mutable import scala.util.matching.Regex import scala.annotation.tailrec object Symbol { type Symbol = String // counting Isabelle symbols, starting from 1 type Offset = Text.Offset type Range = Text.Range /* spaces */ val space_char = ' ' val space = " " private val static_spaces = space * 4000 def spaces(n: Int): String = { require(n >= 0, "negative spaces") if (n < static_spaces.length) static_spaces.substring(0, n) else space * n } /* ASCII characters */ def is_ascii_printable(c: Char): Boolean = space_char <= c && c <= '~' def is_ascii_letter(c: Char): Boolean = 'A' <= c && c <= 'Z' || 'a' <= c && c <= 'z' def is_ascii_digit(c: Char): Boolean = '0' <= c && c <= '9' def is_ascii_hex(c: Char): Boolean = '0' <= c && c <= '9' || 'A' <= c && c <= 'F' || 'a' <= c && c <= 'f' def is_ascii_quasi(c: Char): Boolean = c == '_' || c == '\'' def is_ascii_blank(c: Char): Boolean = " \t\n\u000b\f\r".contains(c) def is_ascii_line_terminator(c: Char): Boolean = "\r\n".contains(c) def is_ascii_letdig(c: Char): Boolean = is_ascii_letter(c) || is_ascii_digit(c) || is_ascii_quasi(c) def is_ascii_identifier(s: String): Boolean = s.length > 0 && is_ascii_letter(s(0)) && s.forall(is_ascii_letdig) def ascii(c: Char): Symbol = { if (c > 127) error("Non-ASCII character: " + quote(c.toString)) else char_symbols(c.toInt) } def is_ascii(s: Symbol): Boolean = s.length == 1 && s(0) < 128 /* symbol matching */ private val symbol_total = new Regex("(?xs) [\ud800-\udbff][\udc00-\udfff] " + """ | \r\n | \\ < \^? ([A-Za-z][A-Za-z0-9_']*)? >? | .""") private def is_plain(c: Char): Boolean = !(c == '\r' || c == '\\' || Character.isHighSurrogate(c)) def is_malformed(s: Symbol): Boolean = s.length match { case 1 => val c = s(0) Character.isHighSurrogate(c) || Character.isLowSurrogate(c) || c == '\ufffd' case 2 => val c1 = s(0) val c2 = s(1) !(c1 == '\r' && c2 == '\n' || Character.isSurrogatePair(c1, c2)) case _ => !s.endsWith(">") || s == "\\<>" || s == "\\<^>" } def is_newline(s: Symbol): Boolean = s == "\n" || s == "\r" || s == "\r\n" class Matcher(text: CharSequence) { private val matcher = symbol_total.pattern.matcher(text) def apply(start: Int, end: Int): Int = { require(0 <= start && start < end && end <= text.length, "bad matcher range") if (is_plain(text.charAt(start))) 1 else { matcher.region(start, end).lookingAt matcher.group.length } } } /* iterator */ private val char_symbols: Array[Symbol] = (0 until 256).iterator.map(i => new String(Array(i.toChar))).toArray def iterator(text: CharSequence): Iterator[Symbol] = new Iterator[Symbol] { private val matcher = new Matcher(text) private var i = 0 def hasNext: Boolean = i < text.length def next(): Symbol = { val n = matcher(i, text.length) val s = if (n == 0) "" else if (n == 1) { val c = text.charAt(i) if (c < char_symbols.length) char_symbols(c) else text.subSequence(i, i + n).toString } else text.subSequence(i, i + n).toString i += n s } } def explode(text: CharSequence): List[Symbol] = iterator(text).toList def length(text: CharSequence): Int = iterator(text).length def trim_blanks(text: CharSequence): String = Library.trim(is_blank, explode(text)).mkString def all_blank(str: String): Boolean = iterator(str).forall(is_blank) def trim_blank_lines(text: String): String = cat_lines(split_lines(text).dropWhile(all_blank).reverse.dropWhile(all_blank).reverse) /* decoding offsets */ object Index { private sealed case class Entry(chr: Int, sym: Int) val empty: Index = new Index(Nil) def apply(text: CharSequence): Index = { val matcher = new Matcher(text) val buf = new mutable.ListBuffer[Entry] var chr = 0 var sym = 0 while (chr < text.length) { val n = matcher(chr, text.length) chr += n sym += 1 if (n > 1) buf += Entry(chr, sym) } if (buf.isEmpty) empty else new Index(buf.toList) } } 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 = { val sym = symbol_offset - 1 val end = index.length @tailrec def bisect(a: Int, b: Int): Int = { if (a < b) { val c = (a + b) / 2 if (sym < index(c).sym) bisect(a, c) else if (c + 1 == end || sym < index(c + 1).sym) c else bisect(c + 1, b) } else -1 } val i = bisect(0, end) if (i < 0) sym else index(i).chr + sym - index(i).sym } def decode(symbol_range: Range): Text.Range = symbol_range.map(decode) override def hashCode: Int = hash override def equals(that: Any): Boolean = that match { case other: Index => index.sameElements(other.index) case _ => false } } /* symbolic text chunks -- without actual text */ object Text_Chunk { sealed abstract class Name case object Default extends Name case class Id(id: Document_ID.Generic) extends Name case class File(name: String) extends Name def apply(text: CharSequence): Text_Chunk = new Text_Chunk(Text.Range(0, text.length), Index(text)) } 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 { case other: Text_Chunk => range == other.range && index == other.index case _ => false } override def toString: String = "Text_Chunk" + range.toString 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 in(r: Range): Option[Text.Range] = range.try_restrict(decode(r)) match { case Some(r1) if !r1.is_singularity => Some(r1) case _ => None } in(symbol_range) orElse in(symbol_range - 1) } } /* recoding text */ private class Recoder(list: List[(String, String)]) { private val (min, max) = { var min = '\uffff' var max = '\u0000' for ((x, _) <- list) { val c = x(0) if (c < min) min = c if (c > max) max = c } (min, max) } private val table = { var tab = Map[String, String]() for ((x, y) <- list) { tab.get(x) match { case None => tab += (x -> y) case Some(z) => error("Duplicate symbol mapping of " + quote(x) + " to " + quote(y) + " vs. " + quote(z)) } } tab } def recode(text: String): String = { val len = text.length val matcher = symbol_total.pattern.matcher(text) val result = new StringBuilder(len) var i = 0 while (i < len) { val c = text(i) if (min <= c && c <= max) { matcher.region(i, len).lookingAt val x = matcher.group result.append(table.getOrElse(x, x)) i = matcher.end } else { result.append(c); i += 1 } } result.toString } } /** symbol interpretation **/ val ARGUMENT_CARTOUCHE = "cartouche" val ARGUMENT_SPACE_CARTOUCHE = "space_cartouche" private lazy val symbols = { val contents = for (path <- Path.split(Isabelle_System.getenv("ISABELLE_SYMBOLS")) if path.is_file) yield (File.read(path)) new Interpretation(cat_lines(contents)) } private class Interpretation(symbols_spec: String) { /* read symbols */ private val No_Decl = new Regex("""(?xs) ^\s* (?: \#.* )? $ """) private val Key = new Regex("""(?xs) (.+): """) private def read_decl(decl: String): (Symbol, Properties.T) = { def err() = error("Bad symbol declaration: " + decl) def read_props(props: List[String]): Properties.T = { props match { case Nil => Nil case _ :: Nil => err() case Key(x) :: y :: rest => (x -> y.replace('\u2423', ' ')) :: read_props(rest) case _ => err() } } decl.split("\\s+").toList match { case sym :: props if sym.length > 1 && !is_malformed(sym) => (sym, read_props(props)) case _ => err() } } private val symbols: List[(Symbol, Properties.T)] = - (((List.empty[(Symbol, Properties.T)], Set.empty[Symbol]) /: - split_lines(symbols_spec).reverse) - { case (res, No_Decl()) => res + split_lines(symbols_spec).reverse. + foldLeft((List.empty[(Symbol, Properties.T)], Set.empty[Symbol])) { + case (res, No_Decl()) => res case ((list, known), decl) => val (sym, props) = read_decl(decl) if (known(sym)) (list, known) else ((sym, props) :: list, known + sym) - })._1 + }._1 /* basic properties */ val properties: Map[Symbol, Properties.T] = Map(symbols: _*) val names: Map[Symbol, (String, String)] = { val Name = new Regex("""\\<\^?([A-Za-z][A-Za-z0-9_']*)>""") val Argument = new Properties.String("argument") def argument(sym: Symbol, props: Properties.T): String = props match { case Argument(arg) => if (arg == ARGUMENT_CARTOUCHE || arg == ARGUMENT_SPACE_CARTOUCHE) arg else error("Bad argument: " + quote(arg) + " for symbol " + quote(sym)) case _ => "" } Map((for ((sym @ Name(a), props) <- symbols) yield sym -> (a, argument(sym, props))): _*) } val groups: List[(String, List[Symbol])] = symbols.flatMap({ case (sym, props) => val gs = for (("group", g) <- props) yield g if (gs.isEmpty) List(sym -> "unsorted") else gs.map(sym -> _) }).groupBy(_._2).toList.map({ case (group, list) => (group, list.map(_._1)) }) .sortBy(_._1) val abbrevs: Multi_Map[Symbol, String] = Multi_Map(( for { (sym, props) <- symbols ("abbrev", a) <- props.reverse } yield sym -> a): _*) val codes: List[(Symbol, Int)] = { val Code = new Properties.String("code") for { (sym, props) <- symbols code <- props match { case Code(s) => try { Some(Integer.decode(s).intValue) } catch { case _: NumberFormatException => error("Bad code for symbol " + sym) } case _ => None } } yield { if (code < 128) error("Illegal ASCII code for symbol " + sym) else (sym, code) } } /* recoding */ private val (decoder, encoder) = { val mapping = for ((sym, code) <- codes) yield (sym, new String(Character.toChars(code))) (new Recoder(mapping), new Recoder(for ((x, y) <- mapping) yield (y, x))) } def decode(text: String): String = decoder.recode(text) def encode(text: String): String = encoder.recode(text) private def recode_set(elems: String*): Set[String] = { val content = elems.toList Set((content ::: content.map(decode)): _*) } private def recode_map[A](elems: (String, A)*): Map[String, A] = { val content = elems.toList Map((content ::: content.map({ case (sym, a) => (decode(sym), a) })): _*) } /* user fonts */ private val Font = new Properties.String("font") val fonts: Map[Symbol, String] = recode_map((for ((sym, Font(font)) <- symbols) yield sym -> font): _*) val font_names: List[String] = Set(fonts.toList.map(_._2): _*).toList val font_index: Map[String, Int] = Map((font_names zip font_names.indices.toList): _*) /* classification */ val letters: Set[String] = recode_set( "A", "B", "C", "D", "E", "F", "G", "H", "I", "J", "K", "L", "M", "N", "O", "P", "Q", "R", "S", "T", "U", "V", "W", "X", "Y", "Z", "a", "b", "c", "d", "e", "f", "g", "h", "i", "j", "k", "l", "m", "n", "o", "p", "q", "r", "s", "t", "u", "v", "w", "x", "y", "z", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\

", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\

", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\
", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\") val blanks: Set[String] = recode_set(space, "\t", "\n", "\u000B", "\f", "\r", "\r\n") val sym_chars = Set("!", "#", "$", "%", "&", "*", "+", "-", "/", "<", "=", ">", "?", "@", "^", "_", "|", "~") val symbolic: Set[String] = recode_set((for {(sym, _) <- symbols; if raw_symbolic(sym)} yield sym): _*) /* misc symbols */ val newline_decoded = decode(newline) val comment_decoded = decode(comment) val cancel_decoded = decode(cancel) val latex_decoded = decode(latex) val marker_decoded = decode(marker) val open_decoded = decode(open) val close_decoded = decode(close) /* control symbols */ val control_decoded: Set[Symbol] = Set((for ((sym, _) <- symbols if sym.startsWith("\\<^")) yield decode(sym)): _*) val sub_decoded = decode(sub) val sup_decoded = decode(sup) val bold_decoded = decode(bold) val emph_decoded = decode(emph) val bsub_decoded = decode(bsub) val esub_decoded = decode(esub) val bsup_decoded = decode(bsup) val esup_decoded = decode(esup) } /* tables */ def properties: Map[Symbol, Properties.T] = symbols.properties def names: Map[Symbol, (String, String)] = symbols.names def groups: List[(String, List[Symbol])] = symbols.groups def abbrevs: Multi_Map[Symbol, String] = symbols.abbrevs def codes: List[(Symbol, Int)] = symbols.codes def groups_code: List[(String, List[Symbol])] = { val has_code = codes.iterator.map(_._1).toSet groups.flatMap({ case (group, symbols) => val symbols1 = symbols.filter(has_code) if (symbols1.isEmpty) None else Some((group, symbols1)) }) } lazy val is_code: Int => Boolean = codes.map(_._2).toSet def decode(text: String): String = symbols.decode(text) def encode(text: String): String = symbols.encode(text) def decode_yxml(text: String, cache: XML.Cache = XML.Cache.none): XML.Body = YXML.parse_body(decode(text), cache = cache) def decode_yxml_failsafe(text: String, cache: XML.Cache = XML.Cache.none): XML.Body = YXML.parse_body_failsafe(decode(text), cache = cache) def encode_yxml(body: XML.Body): String = encode(YXML.string_of_body(body)) def decode_strict(text: String): String = { val decoded = decode(text) if (encode(decoded) == text) decoded else { val bad = new mutable.ListBuffer[Symbol] for (s <- iterator(text) if encode(decode(s)) != s && !bad.contains(s)) bad += s error("Bad Unicode symbols in text: " + commas_quote(bad)) } } def output(unicode_symbols: Boolean, text: String): String = if (unicode_symbols) Symbol.decode(text) else Symbol.encode(text) def fonts: Map[Symbol, String] = symbols.fonts def font_names: List[String] = symbols.font_names def font_index: Map[String, Int] = symbols.font_index def lookup_font(sym: Symbol): Option[Int] = symbols.fonts.get(sym).map(font_index(_)) /* classification */ def is_letter(sym: Symbol): Boolean = symbols.letters.contains(sym) def is_digit(sym: Symbol): Boolean = sym.length == 1 && '0' <= sym(0) && sym(0) <= '9' def is_quasi(sym: Symbol): Boolean = sym == "_" || sym == "'" def is_letdig(sym: Symbol): Boolean = is_letter(sym) || is_digit(sym) || is_quasi(sym) def is_blank(sym: Symbol): Boolean = symbols.blanks.contains(sym) /* symbolic newline */ val newline: Symbol = "\\" def newline_decoded: Symbol = symbols.newline_decoded def print_newlines(str: String): String = if (str.contains('\n')) (for (s <- iterator(str)) yield { if (s == "\n") newline_decoded else s }).mkString else str /* formal comments */ val comment: Symbol = "\\" val cancel: Symbol = "\\<^cancel>" val latex: Symbol = "\\<^latex>" val marker: Symbol = "\\<^marker>" def comment_decoded: Symbol = symbols.comment_decoded def cancel_decoded: Symbol = symbols.cancel_decoded def latex_decoded: Symbol = symbols.latex_decoded def marker_decoded: Symbol = symbols.marker_decoded /* cartouches */ val open: Symbol = "\\" val close: Symbol = "\\" def open_decoded: Symbol = symbols.open_decoded def close_decoded: Symbol = symbols.close_decoded def is_open(sym: Symbol): Boolean = sym == open_decoded || sym == open def is_close(sym: Symbol): Boolean = sym == close_decoded || sym == close def cartouche(s: String): String = open + s + close def cartouche_decoded(s: String): String = open_decoded + s + close_decoded /* symbols for symbolic identifiers */ private def raw_symbolic(sym: Symbol): Boolean = sym.startsWith("\\<") && sym.endsWith(">") && !sym.startsWith("\\<^") def is_symbolic(sym: Symbol): Boolean = !is_open(sym) && !is_close(sym) && (raw_symbolic(sym) || symbols.symbolic.contains(sym)) def is_symbolic_char(sym: Symbol): Boolean = symbols.sym_chars.contains(sym) /* control symbols */ val control_prefix = "\\<^" val control_suffix = ">" def control_name(sym: Symbol): Option[String] = if (is_control_encoded(sym)) Some(sym.substring(control_prefix.length, sym.length - control_suffix.length)) else None def is_control_encoded(sym: Symbol): Boolean = sym.startsWith(control_prefix) && sym.endsWith(control_suffix) def is_control(sym: Symbol): Boolean = is_control_encoded(sym) || symbols.control_decoded.contains(sym) def is_controllable(sym: Symbol): Boolean = !is_blank(sym) && !is_control(sym) && !is_open(sym) && !is_close(sym) && !is_malformed(sym) && sym != "\"" val sub: Symbol = "\\<^sub>" val sup: Symbol = "\\<^sup>" val bold: Symbol = "\\<^bold>" val emph: Symbol = "\\<^emph>" val bsub: Symbol = "\\<^bsub>" val esub: Symbol = "\\<^esub>" val bsup: Symbol = "\\<^bsup>" val esup: Symbol = "\\<^esup>" def sub_decoded: Symbol = symbols.sub_decoded def sup_decoded: Symbol = symbols.sup_decoded def bold_decoded: Symbol = symbols.bold_decoded def emph_decoded: Symbol = symbols.emph_decoded def bsub_decoded: Symbol = symbols.bsub_decoded def esub_decoded: Symbol = symbols.esub_decoded def bsup_decoded: Symbol = symbols.bsup_decoded def esup_decoded: Symbol = symbols.esup_decoded /* metric */ def is_printable(sym: Symbol): Boolean = if (is_ascii(sym)) is_ascii_printable(sym(0)) else !is_control(sym) object Metric extends Pretty.Metric { val unit = 1.0 def apply(str: String): Double = (for (s <- iterator(str)) yield { val sym = encode(s) if (sym.startsWith("\\ Keyword.theory(kind) && !Keyword.theory_end(kind), false) def is_document_command(keywords: Keyword.Keywords, command: Command): Boolean = command.span.is_kind(keywords, Keyword.document, false) def is_diag_command(keywords: Keyword.Keywords, command: Command): Boolean = command.span.is_kind(keywords, Keyword.diag, false) def is_heading_command(command: Command): Boolean = proper_heading_level(command).isDefined def proper_heading_level(command: Command): Option[Int] = command.span.name match { case Thy_Header.CHAPTER => Some(0) case Thy_Header.SECTION => Some(1) case Thy_Header.SUBSECTION => Some(2) case Thy_Header.SUBSUBSECTION => Some(3) case Thy_Header.PARAGRAPH => Some(4) case Thy_Header.SUBPARAGRAPH => Some(5) case _ => None } def heading_level(keywords: Keyword.Keywords, command: Command): Option[Int] = proper_heading_level(command) orElse (if (is_theory_command(keywords, command)) Some(6) else None) /** context blocks **/ def parse_blocks( syntax: Outer_Syntax, node_name: Document.Node.Name, 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 /* stack operations */ def buffer(): mutable.ListBuffer[Document] = new mutable.ListBuffer[Document] var stack: List[(Command, mutable.ListBuffer[Document])] = List((Command.empty, buffer())) def open(command: Command): Unit = { stack = (command, buffer()) :: stack } def close(): Boolean = stack match { case (command, body) :: (_, body2) :: _ => body2 += Block(command.span.name, command.source, body.toList) stack = stack.tail true case _ => false } def flush(): Unit = if (is_plain_theory(stack.head._1)) close() def result(): List[Document] = { while (close()) { } stack.head._2.toList } 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() } stack.head._2 += Atom(command.source.length) } /* result structure */ val spans = syntax.parse_spans(text) spans.foreach(span => add(Command(Document_ID.none, node_name, Command.Blobs_Info.none, span))) result() } /** section headings **/ trait Item { def name: String = "" def source: String = "" def heading_level: Option[Int] = None } object No_Item extends Item 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 = { stack match { case (lev, item, body) :: (_, _, body2) :: _ if level(lev) => body2 += Block(item.name, item.source, body.toList) stack = stack.tail close(level) case _ => } } def result(): List[Document] = { close(_ => true) stack.head._3.toList } def add(item: Item): Unit = { item.heading_level match { case Some(i) => close(_ > i) stack = (i + 1, item, buffer()) :: stack case None => } stack.head._3 += Atom(item.source.length) } } /* outer syntax sections */ 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) } def parse_sections( syntax: Outer_Syntax, node_name: Document.Node.Name, text: CharSequence): List[Document] = { val sections = new Sections(syntax.keywords) for { span <- syntax.parse_spans(text) } { sections.add( new Command_Item(syntax.keywords, Command(Document_ID.none, node_name, Command.Blobs_Info.none, span))) } sections.result() } /* ML sections */ 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] = { val sections = new Sections(Keyword.Keywords.empty) val nl = new ML_Item(ML_Lex.Token(ML_Lex.Kind.SPACE, "\n"), None) var context: Scan.Line_Context = Scan.Finished for (line <- Library.separated_chunks(_ == '\n', text)) { val (toks, next_context) = ML_Lex.tokenize_line(SML, line, context) val level = 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 (s.startsWith("(**** ") && s.endsWith(" ****)")) Some(0) else if (s.startsWith("(*** ") && s.endsWith(" ***)")) Some(1) else if (s.startsWith("(** ") && s.endsWith(" **)")) Some(2) else if (s.startsWith("(* ") && s.endsWith(" *)")) Some(3) else None } else None case _ => None } if (level.isDefined && context == Scan.Finished && next_context == Scan.Finished) toks.foreach(tok => sections.add(new ML_Item(tok, if (tok.is_comment) level else None))) else toks.foreach(tok => sections.add(new ML_Item(tok, None))) sections.add(nl) context = next_context } sections.result() } } diff --git a/src/Pure/Isar/keyword.scala b/src/Pure/Isar/keyword.scala --- a/src/Pure/Isar/keyword.scala +++ b/src/Pure/Isar/keyword.scala @@ -1,231 +1,232 @@ /* Title: Pure/Isar/keyword.scala Author: Makarius Isar keyword classification. */ package isabelle object Keyword { /** keyword classification **/ /* kinds */ val DIAG = "diag" val DOCUMENT_HEADING = "document_heading" val DOCUMENT_BODY = "document_body" val DOCUMENT_RAW = "document_raw" val THY_BEGIN = "thy_begin" val THY_END = "thy_end" val THY_DECL = "thy_decl" val THY_DECL_BLOCK = "thy_decl_block" val THY_DEFN = "thy_defn" val THY_STMT = "thy_stmt" val THY_LOAD = "thy_load" val THY_GOAL = "thy_goal" val THY_GOAL_DEFN = "thy_goal_defn" val THY_GOAL_STMT = "thy_goal_stmt" val QED = "qed" val QED_SCRIPT = "qed_script" val QED_BLOCK = "qed_block" val QED_GLOBAL = "qed_global" val PRF_GOAL = "prf_goal" val PRF_BLOCK = "prf_block" val NEXT_BLOCK = "next_block" val PRF_OPEN = "prf_open" val PRF_CLOSE = "prf_close" val PRF_CHAIN = "prf_chain" val PRF_DECL = "prf_decl" val PRF_ASM = "prf_asm" val PRF_ASM_GOAL = "prf_asm_goal" val PRF_SCRIPT = "prf_script" val PRF_SCRIPT_GOAL = "prf_script_goal" val PRF_SCRIPT_ASM_GOAL = "prf_script_asm_goal" val BEFORE_COMMAND = "before_command" val QUASI_COMMAND = "quasi_command" /* command categories */ val vacuous = Set(DIAG, DOCUMENT_HEADING, DOCUMENT_BODY, DOCUMENT_RAW) val diag = Set(DIAG) val document_heading = Set(DOCUMENT_HEADING) val document_body = Set(DOCUMENT_BODY) val document_raw = Set(DOCUMENT_RAW) val document = Set(DOCUMENT_HEADING, DOCUMENT_BODY, DOCUMENT_RAW) val theory_begin = Set(THY_BEGIN) val theory_end = Set(THY_END) val theory_load = Set(THY_LOAD) val theory = Set(THY_BEGIN, THY_END, THY_LOAD, THY_DECL, THY_DECL_BLOCK, THY_DEFN, THY_STMT, THY_GOAL, THY_GOAL_DEFN, THY_GOAL_STMT) val theory_block = Set(THY_BEGIN, THY_DECL_BLOCK) val theory_body = Set(THY_LOAD, THY_DECL, THY_DECL_BLOCK, THY_DEFN, THY_STMT, THY_GOAL, THY_GOAL_DEFN, THY_GOAL_STMT) val theory_defn = Set(THY_DEFN, THY_GOAL_DEFN) val prf_script = Set(PRF_SCRIPT) val proof = Set(QED, QED_SCRIPT, QED_BLOCK, QED_GLOBAL, PRF_GOAL, PRF_BLOCK, NEXT_BLOCK, PRF_OPEN, PRF_CLOSE, PRF_CHAIN, PRF_DECL, PRF_ASM, PRF_ASM_GOAL, PRF_SCRIPT, PRF_SCRIPT_GOAL, PRF_SCRIPT_ASM_GOAL) val proof_body = Set(DIAG, DOCUMENT_HEADING, DOCUMENT_BODY, DOCUMENT_RAW, PRF_BLOCK, NEXT_BLOCK, PRF_OPEN, PRF_CLOSE, PRF_CHAIN, PRF_DECL, PRF_ASM, PRF_ASM_GOAL, PRF_SCRIPT, PRF_SCRIPT_GOAL, PRF_SCRIPT_ASM_GOAL) val proof_asm = Set(PRF_ASM, PRF_ASM_GOAL) val improper = Set(QED_SCRIPT, PRF_SCRIPT, PRF_SCRIPT_GOAL, PRF_SCRIPT_ASM_GOAL) val theory_goal = Set(THY_GOAL, THY_GOAL_DEFN, THY_GOAL_STMT) val proof_goal = Set(PRF_GOAL, PRF_ASM_GOAL, PRF_SCRIPT_GOAL, PRF_SCRIPT_ASM_GOAL) val qed = Set(QED, QED_SCRIPT, QED_BLOCK) val qed_global = Set(QED_GLOBAL) val proof_open = proof_goal + PRF_OPEN val proof_close = qed + PRF_CLOSE val proof_enclose = Set(PRF_BLOCK, NEXT_BLOCK, QED_BLOCK, PRF_CLOSE) val close_structure = Set(NEXT_BLOCK, QED_BLOCK, PRF_CLOSE, THY_END) /** keyword tables **/ sealed case class Spec( kind: String = "", kind_pos: Position.T = Position.none, load_command: String = "", load_command_pos: Position.T = Position.none, tags: List[String] = Nil) { override def toString: String = kind + (if (load_command.isEmpty) "" else " (" + quote(load_command) + ")") + (if (tags.isEmpty) "" else tags.map(quote).mkString(" % ", " % ", "")) def map(f: String => String): Spec = copy(kind = f(kind), load_command = f(load_command), tags = tags.map(f)) } 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 entries = for ((name, kind) <- kinds.toList.sortBy(_._1)) yield { val load_decl = load_commands.get(name) match { case Some(load_command) => " (" + quote(load_command) + ")" case None => "" } val kind_decl = if (kind == "") "" else " :: " + quote(kind) + load_decl quote(name) + kind_decl } entries.mkString("keywords\n ", " and\n ", "") } /* merge */ def is_empty: Boolean = kinds.isEmpty def ++ (other: Keywords): Keywords = if (this eq other) this else if (is_empty) other else { val kinds1 = if (kinds eq other.kinds) kinds else if (kinds.isEmpty) other.kinds - else (kinds /: other.kinds) { case (m, e) => if (m.isDefinedAt(e._1)) m else m + e } + else other.kinds.foldLeft(kinds) { case (m, e) => if (m.isDefinedAt(e._1)) m else m + e } val load_commands1 = if (load_commands eq other.load_commands) load_commands else if (load_commands.isEmpty) other.load_commands - else - (load_commands /: other.load_commands) { - case (m, e) => if (m.isDefinedAt(e._1)) m else m + e } + else { + other.load_commands.foldLeft(load_commands) { + case (m, e) => if (m.isDefinedAt(e._1)) m else m + e + } + } new Keywords(kinds1, load_commands1) } /* add keywords */ def + (name: String, kind: String = "", load_command: String = ""): Keywords = { val kinds1 = kinds + (name -> kind) val load_commands1 = if (kind == THY_LOAD) { if (!Symbol.iterator(name).forall(Symbol.is_ascii)) error("Bad theory load command " + quote(name)) load_commands + (name -> load_command) } else load_commands new Keywords(kinds1, load_commands1) } def add_keywords(header: Thy_Header.Keywords): Keywords = - (this /: header) { + header.foldLeft(this) { case (keywords, (name, spec)) => if (spec.kind.isEmpty) keywords + Symbol.decode(name) + Symbol.encode(name) else keywords + (Symbol.decode(name), spec.kind, spec.load_command) + (Symbol.encode(name), spec.kind, spec.load_command) } /* command kind */ def is_command(token: Token, check_kind: String => Boolean): Boolean = token.is_command && (kinds.get(token.source) match { case Some(k) => check_kind(k) case None => false }) def is_before_command(token: Token): Boolean = token.is_keyword && kinds.get(token.source) == Some(BEFORE_COMMAND) def is_quasi_command(token: Token): Boolean = token.is_keyword && kinds.get(token.source) == Some(QUASI_COMMAND) def is_indent_command(token: Token): Boolean = token.is_begin_or_command || is_quasi_command(token) /* lexicons */ private def make_lexicon(is_minor: Boolean): Scan.Lexicon = - (Scan.Lexicon.empty /: kinds)( - { - case (lex, (name, kind)) => - if ((kind == "" || kind == BEFORE_COMMAND || kind == QUASI_COMMAND) == is_minor) - lex + name - else lex - }) + kinds.foldLeft(Scan.Lexicon.empty) { + case (lex, (name, kind)) => + if ((kind == "" || kind == BEFORE_COMMAND || kind == QUASI_COMMAND) == is_minor) + lex + name + else lex + } lazy val minor: Scan.Lexicon = make_lexicon(true) lazy val major: Scan.Lexicon = make_lexicon(false) } } diff --git a/src/Pure/Isar/line_structure.scala b/src/Pure/Isar/line_structure.scala --- a/src/Pure/Isar/line_structure.scala +++ b/src/Pure/Isar/line_structure.scala @@ -1,70 +1,70 @@ /* Title: Pure/Isar/line_structure.scala Author: Makarius Line-oriented document structure, e.g. for fold handling. */ package isabelle object Line_Structure { val init: Line_Structure = Line_Structure() } sealed case class Line_Structure( improper: Boolean = true, blank: Boolean = true, command: Boolean = false, 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 = { val improper1 = tokens.forall(tok => !tok.is_proper) val blank1 = tokens.forall(_.is_space) val command1 = tokens.exists(_.is_begin_or_command) val command_depth = tokens.iterator.filter(_.is_proper).nextOption() match { case Some(tok) => if (keywords.is_command(tok, Keyword.close_structure)) Some(after_span_depth - 1) else None case None => None } val depth1 = if (tokens.exists(tok => keywords.is_before_command(tok) || !tok.is_end && keywords.is_command(tok, Keyword.theory))) element_depth else if (command_depth.isDefined) command_depth.get else if (command1) after_span_depth else span_depth val (span_depth1, after_span_depth1, element_depth1) = - ((span_depth, after_span_depth, element_depth) /: tokens) { + tokens.foldLeft((span_depth, after_span_depth, element_depth)) { case (depths @ (x, y, z), tok) => if (tok.is_begin) (z + 2, z + 1, z + 1) else if (tok.is_end) (z + 1, z - 1, z - 1) else if (tok.is_command) { val depth0 = element_depth if (keywords.is_command(tok, Keyword.theory_goal)) (depth0 + 2, depth0 + 1, z) else if (keywords.is_command(tok, Keyword.theory)) (depth0 + 1, depth0, z) else if (keywords.is_command(tok, Keyword.proof_open)) (y + 2, y + 1, z) else if (keywords.is_command(tok, Set(Keyword.PRF_BLOCK))) (y + 2, y + 1, z) else if (keywords.is_command(tok, Set(Keyword.QED_BLOCK))) (y - 1, y - 2, z) else if (keywords.is_command(tok, Set(Keyword.PRF_CLOSE))) (y, y - 1, z) else if (keywords.is_command(tok, Keyword.proof_close)) (y + 1, y - 1, z) else if (keywords.is_command(tok, Keyword.qed_global)) (depth0 + 1, depth0, z) else depths } else depths } Line_Structure( improper1, blank1, command1, depth1, span_depth1, after_span_depth1, element_depth1) } } diff --git a/src/Pure/Isar/outer_syntax.scala b/src/Pure/Isar/outer_syntax.scala --- a/src/Pure/Isar/outer_syntax.scala +++ b/src/Pure/Isar/outer_syntax.scala @@ -1,210 +1,211 @@ /* Title: Pure/Isar/outer_syntax.scala Author: Makarius Isabelle/Isar outer syntax. */ package isabelle import scala.collection.mutable object Outer_Syntax { /* syntax */ val empty: Outer_Syntax = new Outer_Syntax() - def merge(syns: List[Outer_Syntax]): Outer_Syntax = (empty /: syns)(_ ++ _) + def merge(syns: List[Outer_Syntax]): Outer_Syntax = syns.foldLeft(empty)(_ ++ _) /* string literals */ def quote_string(str: String): String = { val result = new StringBuilder(str.length + 10) result += '"' for (s <- Symbol.iterator(str)) { if (s.length == 1) { val c = s(0) if (c < 32 && c != YXML.X && c != YXML.Y || c == '\\' || c == '"') { result += '\\' if (c < 10) result += '0' if (c < 100) result += '0' result ++= c.asInstanceOf[Int].toString } else result += c } else result ++= s } result += '"' result.toString } } final class Outer_Syntax private( 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) { /** syntax content **/ override def toString: String = keywords.toString /* keywords */ def + (name: String, kind: String = "", load_command: String = ""): Outer_Syntax = new Outer_Syntax( keywords + (name, kind, load_command), rev_abbrevs, language_context, has_tokens = true) def add_keywords(keywords: Thy_Header.Keywords): Outer_Syntax = - (this /: keywords) { + keywords.foldLeft(this) { case (syntax, (name, spec)) => syntax + (Symbol.decode(name), spec.kind, spec.load_command) + (Symbol.encode(name), spec.kind, spec.load_command) } /* abbrevs */ def abbrevs: Thy_Header.Abbrevs = rev_abbrevs.reverse def add_abbrevs(new_abbrevs: Thy_Header.Abbrevs): Outer_Syntax = if (new_abbrevs.isEmpty) this else { val rev_abbrevs1 = Library.distinct(new_abbrevs) reverse_::: rev_abbrevs new Outer_Syntax(keywords, rev_abbrevs1, language_context, has_tokens) } /* completion */ private lazy val completion: Completion = { val completion_keywords = (keywords.minor.iterator ++ keywords.major.iterator).toList val completion_abbrevs = completion_keywords.flatMap(name => (if (Keyword.theory_block.contains(keywords.kinds(name))) List((name, name + "\nbegin\n\u0007\nend")) else Nil) ::: (if (Completion.Word_Parsers.is_word(name)) List((name, name)) else Nil)) ::: abbrevs.flatMap( { case (a, b) => val a1 = Symbol.decode(a) val a2 = Symbol.encode(a) val b1 = Symbol.decode(b) List((a1, b1), (a2, b1)) }) Completion.make(completion_keywords, completion_abbrevs) } def complete( history: Completion.History, unicode: Boolean, explicit: Boolean, start: Text.Offset, text: CharSequence, caret: Int, context: Completion.Language_Context): Option[Completion.Result] = { completion.complete(history, unicode, explicit, start, text, caret, context) } /* build */ def + (header: Document.Node.Header): Outer_Syntax = add_keywords(header.keywords).add_abbrevs(header.abbrevs) def ++ (other: Outer_Syntax): Outer_Syntax = if (this eq other) this else if (this eq Outer_Syntax.empty) other else { val keywords1 = keywords ++ other.keywords val rev_abbrevs1 = Library.merge(rev_abbrevs, other.rev_abbrevs) if ((keywords eq keywords1) && (rev_abbrevs eq rev_abbrevs1)) this else new Outer_Syntax(keywords1, rev_abbrevs1, language_context, has_tokens) } /* load commands */ def load_command(name: String): Option[String] = keywords.load_commands.get(name) def has_load_commands(text: String): Boolean = keywords.load_commands.exists({ case (cmd, _) => text.containsSlice(cmd) }) /* language context */ def set_language_context(context: Completion.Language_Context): Outer_Syntax = new Outer_Syntax(keywords, rev_abbrevs, context, has_tokens) def no_tokens: Outer_Syntax = { require(keywords.is_empty, "bad syntax keywords") new Outer_Syntax( rev_abbrevs = rev_abbrevs, language_context = language_context, has_tokens = false) } /** parsing **/ /* command spans */ 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 = { val kind = if (content.forall(_.is_ignored)) Command_Span.Ignored_Span else if (content.exists(_.is_error)) Command_Span.Malformed_Span else content.find(_.is_command) match { case None => Command_Span.Malformed_Span case Some(cmd) => val name = cmd.source val offset = - (0 /: content.takeWhile(_ != cmd)) { - case (i, tok) => i + Symbol.length(tok.source) } + content.takeWhile(_ != cmd).foldLeft(0) { + case (i, tok) => i + Symbol.length(tok.source) + } val end_offset = offset + Symbol.length(name) val range = Text.Range(offset, end_offset) + 1 Command_Span.Command_Span(name, Position.Range(range)) } result += Command_Span.Span(kind, content) } def flush(): Unit = { if (content.nonEmpty) { ship(content.toList); content.clear } if (ignored.nonEmpty) { ship(ignored.toList); ignored.clear } } for (tok <- toks) { 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 } else { content ++= ignored; ignored.clear; content += tok } } flush() result.toList } def parse_spans(input: CharSequence): List[Command_Span.Span] = parse_spans(Token.explode(keywords, input)) } diff --git a/src/Pure/ML/ml_statistics.scala b/src/Pure/ML/ml_statistics.scala --- a/src/Pure/ML/ml_statistics.scala +++ b/src/Pure/ML/ml_statistics.scala @@ -1,324 +1,324 @@ /* Title: Pure/ML/ml_statistics.scala Author: Makarius ML runtime statistics. */ package isabelle import scala.annotation.tailrec import scala.collection.mutable import scala.collection.immutable.{SortedSet, SortedMap} import scala.swing.{Frame, Component} import org.jfree.data.xy.{XYSeries, XYSeriesCollection} import org.jfree.chart.{JFreeChart, ChartPanel, ChartFactory} import org.jfree.chart.plot.PlotOrientation object ML_Statistics { /* properties */ val Now = new Properties.Double("now") def now(props: Properties.T): Double = Now.unapply(props).get /* memory status */ val Heap_Size = new Properties.Long("size_heap") 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) { 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 gc_progress: Option[Double] = if (1 <= gc_percent && gc_percent <= 100) Some((gc_percent - 1) * 0.01) else None } 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) Memory_Status(heap_size, heap_free, gc_percent) } /* monitor process */ 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 = { val props = Library.space_explode(',', line).flatMap((entry: String) => Library.space_explode('=', entry) match { case List(a, b) => Some((a, b)) case _ => None }) if (props.nonEmpty) consume(props) } val env_prefix = if (stats_dir.isEmpty) "" else "export POLYSTATSDIR=" + Bash.string(stats_dir) + "\n" Bash.process(env_prefix + "exec \"$POLYML_EXE\" -q --use src/Pure/ML/ml_statistics.ML --eval " + Bash.string("ML_Statistics.monitor " + ML_Syntax.print_long(pid) + " " + ML_Syntax.print_double(delay.seconds)), cwd = Path.explode("~~").file) .result(progress_stdout = progress_stdout, strict = false).check } /* protocol handler */ class Handler extends Session.Protocol_Handler { private var session: Session = null private var monitoring: Future[Unit] = Future.value(()) override def init(session: Session): Unit = synchronized { this.session = session } override def exit(): Unit = synchronized { session = null monitoring.cancel } private def consume(props: Properties.T): Unit = synchronized { if (session != null) { val props1 = (session.cache.props(props ::: Java_Statistics.jvm_statistics())) session.runtime_statistics.post(Session.Runtime_Statistics(props1)) } } private def ml_statistics(msg: Prover.Protocol_Output): Boolean = synchronized { msg.properties match { case Markup.ML_Statistics(pid, stats_dir) => monitoring = Future.thread("ML_statistics") { monitor(pid, stats_dir = stats_dir, consume = consume) } true case _ => false } } override val functions = List(Markup.ML_Statistics.name -> ml_statistics) } /* memory fields (mega bytes) */ def mem_print(x: Long): Option[String] = if (x == 0L) None else Some(x.toString + " M") def mem_scale(x: Long): Long = x / 1024 / 1024 def mem_field_scale(name: String, x: Double): Double = if (heap_fields._2.contains(name) || program_fields._2.contains(name)) mem_scale(x.toLong).toDouble else x val CODE_SIZE = "size_code" val STACK_SIZE = "size_stacks" val HEAP_SIZE = "size_heap" /* standard fields */ type Fields = (String, List[String]) val tasks_fields: Fields = ("Future tasks", List("tasks_ready", "tasks_pending", "tasks_running", "tasks_passive", "tasks_urgent", "tasks_total")) val workers_fields: Fields = ("Worker threads", List("workers_total", "workers_active", "workers_waiting")) val GC_fields: Fields = ("GCs", List("partial_GCs", "full_GCs", "share_passes")) val heap_fields: Fields = ("Heap", List(HEAP_SIZE, "size_allocation", "size_allocation_free", "size_heap_free_last_full_GC", "size_heap_free_last_GC")) val program_fields: Fields = ("Program", List("size_code", "size_stacks")) val threads_fields: Fields = ("Threads", List("threads_total", "threads_in_ML", "threads_wait_condvar", "threads_wait_IO", "threads_wait_mutex", "threads_wait_signal")) val time_fields: Fields = ("Time", List("time_elapsed", "time_elapsed_GC", "time_CPU", "time_GC")) val speed_fields: Fields = ("Speed", List("speed_CPU", "speed_GC")) private val time_speed = Map("time_CPU" -> "speed_CPU", "time_GC" -> "speed_GC") val java_heap_fields: Fields = ("Java heap", List("java_heap_size", "java_heap_used")) val java_thread_fields: Fields = ("Java threads", List("java_threads_total", "java_workers_total", "java_workers_active")) val main_fields: List[Fields] = List(heap_fields, tasks_fields, workers_fields) val other_fields: List[Fields] = List(threads_fields, GC_fields, program_fields, time_fields, speed_fields, java_heap_fields, java_thread_fields) val all_fields: List[Fields] = main_fields ::: other_fields /* content interpretation */ 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 = (key: String) => true): ML_Statistics = { require(ml_statistics0.forall(props => Now.unapply(props).isDefined), "missing \"now\" field") val ml_statistics = ml_statistics0.sortBy(now) val time_start = if (ml_statistics.isEmpty) 0.0 else now(ml_statistics.head) val duration = if (ml_statistics.isEmpty) 0.0 else now(ml_statistics.last) - time_start val fields = SortedSet.empty[String] ++ (for { props <- ml_statistics.iterator (x, _) <- props.iterator if x != Now.name && domain(x) } yield x) val content = { var last_edge = Map.empty[String, (Double, Double, Double)] val result = new mutable.ListBuffer[ML_Statistics.Entry] for (props <- ml_statistics) { val time = now(props) - time_start // rising edges -- relative speed val speeds = (for { (key, value) <- props.iterator key1 <- time_speed.get(key) if domain(key1) } yield { val (x0, y0, s0) = last_edge.getOrElse(key, (0.0, 0.0, 0.0)) val x1 = time val y1 = java.lang.Double.parseDouble(value) val s1 = if (x1 == x0) 0.0 else (y1 - y0) / (x1 - x0) if (y1 > y0) { last_edge += (key -> (x1, y1, s1)) (key1, s1.toString) } else (key1, s0.toString) }).toList val data = SortedMap.empty[String, Double] ++ (for { (x, y) <- props.iterator ++ speeds.iterator if x != Now.name && domain(x) z = java.lang.Double.parseDouble(y) if z != 0.0 } yield { (x.intern, mem_field_scale(x, z)) }) result += ML_Statistics.Entry(time, data) } result.toList } new ML_Statistics(heading, fields, content, time_start, duration) } } final class ML_Statistics private( val heading: String, val fields: Set[String], val content: List[ML_Statistics.Entry], val time_start: Double, val duration: Double) { /* content */ def maximum(field: String): Double = - (0.0 /: content)({ case (m, e) => m max e.get(field) }) + content.foldLeft(0.0) { case (m, e) => m max e.get(field) } def average(field: String): Double = { @tailrec def sum(t0: Double, list: List[ML_Statistics.Entry], acc: Double): Double = list match { case Nil => acc case e :: es => val t = e.time sum(t, es, (t - t0) * e.get(field) + acc) } content match { case Nil => 0.0 case List(e) => e.get(field) case e :: es => sum(e.time, es, 0.0) / duration } } /* charts */ def update_data(data: XYSeriesCollection, selected_fields: List[String]): Unit = { data.removeAllSeries for (field <- selected_fields) { val series = new XYSeries(field) content.foreach(entry => series.add(entry.time, entry.get(field))) data.addSeries(series) } } def chart(title: String, selected_fields: List[String]): JFreeChart = { val data = new XYSeriesCollection update_data(data, selected_fields) ChartFactory.createXYLineChart(title, "time", "value", data, PlotOrientation.VERTICAL, true, true, true) } def chart(fields: ML_Statistics.Fields): JFreeChart = chart(fields._1, fields._2) def show_frames(fields: List[ML_Statistics.Fields] = ML_Statistics.main_fields): Unit = fields.map(chart).foreach(c => GUI_Thread.later { new Frame { iconImage = GUI.isabelle_image() title = heading contents = Component.wrap(new ChartPanel(c)) visible = true } }) } diff --git a/src/Pure/PIDE/command.scala b/src/Pure/PIDE/command.scala --- a/src/Pure/PIDE/command.scala +++ b/src/Pure/PIDE/command.scala @@ -1,651 +1,653 @@ /* Title: Pure/PIDE/command.scala Author: Fabian Immler, TU Munich Author: Makarius Prover commands with accumulated results from execution. */ package isabelle import scala.collection.immutable.SortedMap object Command { /* blobs */ 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))) } } sealed case class Blob( name: Document.Node.Name, src_path: Path, 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 { val none: Blobs_Info = Blobs_Info(Nil) def errors(msgs: List[String]): Blobs_Info = Blobs_Info(msgs.map(msg => Exn.Exn[Blob](ERROR(msg)))) } sealed case class Blobs_Info(blobs: List[Exn.Result[Blob]], index: Int = -1) /** accumulated results from prover **/ /* results */ object Results { type Entry = (Long, XML.Elem) val empty: Results = new Results(SortedMap.empty) - def make(args: IterableOnce[Results.Entry]): Results = (empty /: args)(_ + _) - def merge(args: IterableOnce[Results]): Results = (empty /: args)(_ ++ _) + def make(args: IterableOnce[Results.Entry]): Results = args.foldLeft(empty)(_ + _) + def merge(args: IterableOnce[Results]): Results = args.foldLeft(empty)(_ ++ _) } 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) def iterator: Iterator[Results.Entry] = rep.iterator def + (entry: Results.Entry): Results = if (defined(entry._1)) this else new Results(rep + entry) def ++ (other: Results): Results = if (this eq other) this else if (rep.isEmpty) other - else (this /: other.iterator)(_ + _) + else other.iterator.foldLeft(this)(_ + _) override def hashCode: Int = rep.hashCode override def equals(that: Any): Boolean = that match { case other: Results => rep == other.rep case _ => false } override def toString: String = iterator.mkString("Results(", ", ", ")") } /* exports */ object Exports { type Entry = (Long, Export.Entry) val empty: Exports = new Exports(SortedMap.empty) - def merge(args: IterableOnce[Exports]): Exports = (empty /: args)(_ ++ _) + def merge(args: IterableOnce[Exports]): Exports = args.foldLeft(empty)(_ ++ _) } final class Exports private(private val rep: SortedMap[Long, Export.Entry]) { def is_empty: Boolean = rep.isEmpty def iterator: Iterator[Exports.Entry] = rep.iterator def + (entry: Exports.Entry): Exports = if (rep.isDefinedAt(entry._1)) this else new Exports(rep + entry) def ++ (other: Exports): Exports = if (this eq other) this else if (rep.isEmpty) other - else (this /: other.iterator)(_ + _) + else other.iterator.foldLeft(this)(_ + _) override def hashCode: Int = rep.hashCode override def equals(that: Any): Boolean = that match { case other: Exports => rep == other.rep case _ => false } override def toString: String = iterator.mkString("Exports(", ", ", ")") } /* markups */ 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 { 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)) - def make(args: IterableOnce[Entry]): Markups = (empty /: args)(_ + _) - def merge(args: IterableOnce[Markups]): Markups = (empty /: args)(_ ++ _) + def make(args: IterableOnce[Entry]): Markups = args.foldLeft(empty)(_ + _) + def merge(args: IterableOnce[Markups]): Markups = args.foldLeft(empty)(_ ++ _) } 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 = rep.getOrElse(index, Markup_Tree.empty) def add(index: Markup_Index, markup: Text.Markup): Markups = new Markups(rep + (index -> (this(index) + markup))) def + (entry: Markups.Entry): Markups = { val (index, tree) = entry new Markups(rep + (index -> (this(index).merge(tree, Text.Range.full, Markup.Elements.full)))) } def ++ (other: Markups): Markups = if (this eq other) this else if (rep.isEmpty) other - else (this /: other.rep.iterator)(_ + _) + else other.rep.iterator.foldLeft(this)(_ + _) def redirection_iterator: Iterator[Document_ID.Generic] = for (Markup_Index(_, Symbol.Text_Chunk.Id(id)) <- rep.keysIterator) yield id def redirect(other_id: Document_ID.Generic): Markups = { val rep1 = (for { (Markup_Index(status, Symbol.Text_Chunk.Id(id)), markup) <- rep.iterator if other_id == id } yield (Markup_Index(status, Symbol.Text_Chunk.Default), markup)).toMap if (rep1.isEmpty) Markups.empty else new Markups(rep1) } override def hashCode: Int = rep.hashCode override def equals(that: Any): Boolean = that match { case other: Markups => rep == other.rep case _ => false } override def toString: String = rep.iterator.mkString("Markups(", ", ", ")") } /* 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) def get_result_proper(states: List[State], props: Properties.T): Option[Results.Entry] = for { serial <- Markup.Serial.unapply(props) elem <- get_result(states, serial) if elem.body.nonEmpty } yield serial -> elem def merge_results(states: List[State]): Results = Results.merge(states.map(_.results)) def merge_exports(states: List[State]): Exports = Exports.merge(states.map(_.exports)) def merge_markups(states: List[State]): Markups = Markups.merge(states.map(_.markups)) def merge_markup(states: List[State], index: Markup_Index, range: Text.Range, elements: Markup.Elements): Markup_Tree = Markup_Tree.merge(states.map(_.markup(index)), range, elements) def merge(command: Command, states: List[State]): State = State(command, states.flatMap(_.status), merge_results(states), merge_exports(states), merge_markups(states)) } sealed case class State( command: Command, status: List[Markup] = Nil, results: Results = Results.empty, exports: Exports = Exports.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 = { var touched = false var forks = 0 var runs = 0 for (markup <- status) { markup.name match { case Markup.FORKED => touched = true; forks += 1 case Markup.JOINED => forks -= 1 case Markup.RUNNING => touched = true; runs += 1 case Markup.FINISHED => runs -= 1 case _ => } } touched && forks == 0 && runs == 0 } 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)) else Nil val errors = if (results.iterator.exists(p => Protocol.is_error(p._2))) List(Markup(Markup.ERROR, Nil)) else Nil Document_Status.Command_Status.make((warnings ::: errors ::: status).iterator) } def markup(index: Markup_Index): Markup_Tree = markups(index) 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)) } private def add_status(st: Markup): State = copy(status = st :: status) private def add_result(entry: Results.Entry): State = copy(results = results + entry) def add_export(entry: Exports.Entry): Option[State] = if (command.node_name.theory == entry._2.theory_name) Some(copy(exports = exports + entry)) else None private def add_markup( 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) else markups copy(markups = markups1.add(Markup_Index(false, chunk_name), m)) } def accumulate( self_id: Document_ID.Generic => Boolean, other_id: (Document.Node.Name, Document_ID.Generic) => Option[(Symbol.Text_Chunk.Id, Symbol.Text_Chunk)], message: XML.Elem, cache: XML.Cache): State = message match { case XML.Elem(Markup(Markup.STATUS, _), msgs) => if (command.span.is_theory) this else { - (this /: msgs)((state, msg) => - msg match { - case elem @ XML.Elem(markup, Nil) => - state. - add_status(markup). - add_markup(true, Symbol.Text_Chunk.Default, Text.Info(command.core_range, elem)) - case _ => - Output.warning("Ignored status message: " + msg) - state - }) + msgs.foldLeft(this) { + case (state, msg) => + msg match { + case elem @ XML.Elem(markup, Nil) => + state. + add_status(markup). + add_markup(true, Symbol.Text_Chunk.Default, Text.Info(command.core_range, elem)) + case _ => + Output.warning("Ignored status message: " + msg) + state + } + } } case XML.Elem(Markup(Markup.REPORT, atts0), msgs) => - (this /: msgs)((state, msg) => - { + msgs.foldLeft(this) { + case (state, msg) => def bad(): Unit = Output.warning("Ignored report message: " + msg) msg match { case XML.Elem(Markup(name, atts), args) => command.reported_position(atts) orElse command.reported_position(atts0) match { case Some((id, chunk_name, target_range)) => val target = if (self_id(id) && command.chunks.isDefinedAt(chunk_name)) Some((chunk_name, command.chunks(chunk_name))) else if (chunk_name == Symbol.Text_Chunk.Default) other_id(command.node_name, id) else None (target, target_range) match { case (Some((target_name, target_chunk)), Some(symbol_range)) if !symbol_range.is_singularity => target_chunk.incorporate(symbol_range) match { case Some(range) => val props = atts.filterNot(Markup.position_property) val elem = cache.elem(XML.Elem(Markup(name, props), args)) state.add_markup(false, target_name, Text.Info(range, elem)) case None => bad(); state } case _ => // silently ignore excessive reports state } case _ => bad(); state } case _ => bad(); state } - }) + } case XML.Elem(Markup(name, props), body) => props match { case Markup.Serial(i) => val markup_message = cache.elem(XML.Elem(Markup(Markup.message(name), props), body)) val message_markup = cache.elem(XML.elem(Markup(name, props.filter(p => p._1 == Markup.SERIAL)))) var st = add_result(i -> markup_message) if (Protocol.is_inlined(message)) { for { (chunk_name, chunk) <- command.chunks.iterator range <- command.message_positions(self_id, chunk_name, chunk, message) } st = st.add_markup(false, chunk_name, Text.Info(range, message_markup)) } st case _ => Output.warning("Ignored message without serial number: " + message) this } } } /** static content **/ /* make commands */ def apply( id: Document_ID.Command, node_name: Document.Node.Name, blobs_info: Blobs_Info, span: Command_Span.Span): Command = { val (source, span1) = span.compact_source new Command(id, node_name, blobs_info, span1, source, Results.empty, Markups.empty) } val empty: Command = Command(Document_ID.none, Document.Node.Name.empty, Blobs_Info.none, Command_Span.empty) def unparsed( source: String, theory: Boolean = false, id: Document_ID.Command = Document_ID.none, 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 = { val (source1, span1) = Command_Span.unparsed(source, theory).compact_source new Command(id, node_name, blobs_info, span1, source1, results, markups) } def text(source: String): Command = unparsed(source) def rich_text(id: Document_ID.Command, results: Results, body: XML.Body): Command = unparsed(XML.content(body), id = id, results = results, markups = Markups.init(Markup_Tree.from_XML(body))) /* edits and perspective */ type Edit = (Option[Command], Option[Command]) object Perspective { val empty: Perspective = Perspective(Nil) } sealed case class Perspective(commands: List[Command]) // visible commands in canonical order { def is_empty: Boolean = commands.isEmpty def same(that: Perspective): Boolean = { val cmds1 = this.commands val cmds2 = that.commands require(!cmds1.exists(_.is_undefined), "cmds1 not defined") require(!cmds2.exists(_.is_undefined), "cmds2 not defined") cmds1.length == cmds2.length && (cmds1.iterator zip cmds2.iterator).forall({ case (c1, c2) => c1.id == c2.id }) } } /* blobs: inlined errors and auxiliary files */ def blobs_info( resources: Resources, syntax: Outer_Syntax, 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.name match { // inlined errors case Thy_Header.THEORY => val reader = span.content_reader val header = resources.check_thy(node_name, span.content_reader) val imports_pos = header.imports_pos val raw_imports = try { val read_imports = Thy_Header.read(node_name, reader).imports.map(_._1) if (imports_pos.length == read_imports.length) read_imports else error("") } catch { case _: Throwable => List.fill(header.imports.length)("") } val errors = for { ((import_name, pos), s) <- imports_pos zip raw_imports if !can_import(import_name) } yield { val completion = if (Thy_Header.is_base_name(s)) resources.complete_import_name(node_name, s) else Nil "Bad theory import " + Markup.Path(import_name.node).markup(quote(import_name.toString)) + Position.here(pos) + Completion.report_theories(pos, completion) } Blobs_Info.errors(errors) // auxiliary files case _ => val loaded_files = span.loaded_files(syntax) val blobs = loaded_files.files.map(file => (Exn.capture { val src_path = Path.explode(file) val name = Document.Node.Name(resources.append(node_name, src_path)) val content = get_blob(name).map(blob => (blob.bytes.sha1_digest, blob.chunk)) Blob(name, src_path, content) }).user_error) Blobs_Info(blobs, index = loaded_files.index) } } def build_blobs_info( syntax: Outer_Syntax, node_name: Document.Node.Name, load_commands: List[Command_Span.Span]): Blobs_Info = { val blobs = for { span <- load_commands file <- span.loaded_files(syntax).files } yield { (Exn.capture { val dir = node_name.master_dir_path val src_path = Path.explode(file) val name = Document.Node.Name((dir + src_path).expand.implode_symbolic) Blob.read_file(name, src_path) }).user_error } Blobs_Info(blobs) } } 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) { override def toString: String = id + "/" + span.kind.toString /* classification */ def is_proper: Boolean = span.kind.isInstanceOf[Command_Span.Command_Span] def is_ignored: Boolean = span.kind == Command_Span.Ignored_Span def is_undefined: Boolean = id == Document_ID.none val is_unparsed: Boolean = span.content.exists(_.is_unparsed) val is_unfinished: Boolean = span.content.exists(_.is_unfinished) def potentially_initialized: Boolean = span.name == Thy_Header.THEORY /* blobs */ def blobs: List[Exn.Result[Command.Blob]] = blobs_info.blobs def blobs_index: Int = blobs_info.index def blobs_ok: Boolean = blobs.forall({ case Exn.Res(_) => true case _ => false }) def blobs_names: List[Document.Node.Name] = for (Exn.Res(blob) <- blobs) yield blob.name def blobs_undefined: List[Document.Node.Name] = for (Exn.Res(blob) <- blobs if blob.content.isEmpty) yield blob.name def blobs_defined: List[(Document.Node.Name, SHA1.Digest)] = for (Exn.Res(blob) <- blobs; (digest, _) <- blob.content) yield (blob.name, digest) def blobs_changed(doc_blobs: Document.Blobs): Boolean = blobs.exists({ case Exn.Res(blob) => doc_blobs.changed(blob.name) case _ => false }) /* source chunks */ val chunk: Symbol.Text_Chunk = Symbol.Text_Chunk(source) val chunks: Map[Symbol.Text_Chunk.Name, Symbol.Text_Chunk] = ((Symbol.Text_Chunk.Default -> chunk) :: (for (Exn.Res(blob) <- blobs; (_, file) <- blob.content) yield blob.chunk_file -> file)).toMap def length: Int = source.length def range: Text.Range = chunk.range val core_range: Text.Range = Text.Range(0, - (length /: span.content.reverse.iterator.takeWhile(_.is_ignored))(_ - _.source.length)) + span.content.reverse.iterator.takeWhile(_.is_ignored).foldLeft(length)(_ - _.source.length)) def source(range: Text.Range): String = range.substring(source) /* theory parents */ def theory_parents(resources: Resources): List[Document.Node.Name] = if (span.name == Thy_Header.THEORY) { try { val header = Thy_Header.read(node_name, span.content_reader) for ((s, _) <- header.imports) yield { try { resources.import_name(node_name, s) } catch { case ERROR(_) => Document.Node.Name.empty } } } catch { case ERROR(_) => Nil } } else Nil /* reported positions */ 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 = pos match { case Position.File(name) if name != node_name.node => Symbol.Text_Chunk.File(name) case _ => Symbol.Text_Chunk.Default } Some((id, chunk_name, Position.Range.unapply(pos))) case _ => None } } def message_positions( self_id: Document_ID.Generic => Boolean, chunk_name: Symbol.Text_Chunk.Name, chunk: Symbol.Text_Chunk, 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 => val opt_range = reported_range orElse { if (name == Symbol.Text_Chunk.Default) Position.Range.unapply(span.position) else None } opt_range match { case Some(symbol_range) => chunk.incorporate(symbol_range) match { case Some(range) => set + range case _ => set } case None => set } case _ => set } def tree(set: Set[Text.Range], t: XML.Tree): Set[Text.Range] = t match { case XML.Wrapped_Elem(Markup(name, props), _, body) => body.foldLeft(if (Rendering.position_elements(name)) elem(props, set) else set)(tree) case XML.Elem(Markup(name, props), body) => body.foldLeft(if (Rendering.position_elements(name)) elem(props, set) else set)(tree) case XML.Text(_) => set } val set = tree(Set.empty, message) if (set.isEmpty) elem(message.markup.properties, set) else set } /* accumulated results */ val init_state: Command.State = Command.State(this, results = init_results, markups = init_markups) val empty_state: Command.State = Command.State(this) } diff --git a/src/Pure/PIDE/command_span.scala b/src/Pure/PIDE/command_span.scala --- a/src/Pure/PIDE/command_span.scala +++ b/src/Pure/PIDE/command_span.scala @@ -1,146 +1,146 @@ /* Title: Pure/PIDE/command_span.scala Author: Makarius Syntactic representation of command spans. */ package isabelle import scala.collection.mutable import scala.util.parsing.input.CharSequenceReader object Command_Span { /* loaded files */ object Loaded_Files { val none: Loaded_Files = Loaded_Files(Nil, -1) } sealed case class Loaded_Files(files: List[String], index: Int) class Load_Command(val name: String) extends Isabelle_System.Service { override def toString: String = name def extensions: List[String] = Nil def loaded_files(tokens: List[(Token, Int)]): Loaded_Files = tokens.collectFirst({ case (t, i) if t.is_embedded => (t.content, i) }) match { case Some((file, i)) => extensions match { case Nil => Loaded_Files(List(file), i) case exts => Loaded_Files(exts.map(ext => file + "." + ext), i) } case None => Loaded_Files.none } } lazy val load_commands: List[Load_Command] = new Load_Command("") :: Isabelle_System.make_services(classOf[Load_Command]) /* span kind */ sealed abstract class Kind { override def toString: String = this match { case Command_Span(name, _) => proper_string(name) getOrElse "" case Ignored_Span => "" case Malformed_Span => "" case Theory_Span => "" } } case class Command_Span(name: String, pos: Position.T) extends Kind case object Ignored_Span extends Kind case object Malformed_Span extends Kind case object Theory_Span extends Kind /* span */ sealed case class Span(kind: Kind, content: List[Token]) { def is_theory: Boolean = kind == Theory_Span def name: String = kind match { case k: Command_Span => k.name case _ => "" } def position: Position.T = kind match { case k: Command_Span => k.pos case _ => Position.none } def keyword_pos(start: Token.Pos): Token.Pos = kind match { case _: Command_Span => - (start /: content.iterator.takeWhile(tok => !tok.is_command))(_.advance(_)) + content.iterator.takeWhile(tok => !tok.is_command).foldLeft(start)(_.advance(_)) case _ => start } def is_kind(keywords: Keyword.Keywords, pred: String => Boolean, other: Boolean): Boolean = keywords.kinds.get(name) match { case Some(k) => pred(k) case None => other } def is_begin: Boolean = content.exists(_.is_begin) def is_end: Boolean = content.exists(_.is_end) def content_reader: CharSequenceReader = Scan.char_reader(Token.implode(content)) - def length: Int = (0 /: content)(_ + _.source.length) + def length: Int = content.foldLeft(0)(_ + _.source.length) def compact_source: (String, Span) = { val source = Token.implode(content) val content1 = new mutable.ListBuffer[Token] var i = 0 for (Token(kind, s) <- content) { val n = s.length val s1 = source.substring(i, i + n) content1 += Token(kind, s1) i += n } (source, Span(kind, content1.toList)) } def clean_arguments: List[(Token, Int)] = { if (name.nonEmpty) { def clean(toks: List[(Token, Int)]): List[(Token, Int)] = toks match { case (t1, i1) :: (t2, i2) :: rest => if (t1.is_keyword && t1.source == "%" && t2.is_name) clean(rest) else (t1, i1) :: clean((t2, i2) :: rest) case _ => toks } clean(content.zipWithIndex.filter({ case (t, _) => t.is_proper })) .dropWhile({ case (t, _) => !t.is_command }) .dropWhile({ case (t, _) => t.is_command }) } else Nil } def is_load_command(syntax: Outer_Syntax): Boolean = syntax.load_command(name).isDefined def loaded_files(syntax: Outer_Syntax): Loaded_Files = syntax.load_command(name) match { case None => Loaded_Files.none case Some(a) => load_commands.find(_.name == a) match { case Some(load_command) => load_command.loaded_files(clean_arguments) case None => error("Undefined load command function: " + a) } } } val empty: Span = Span(Ignored_Span, Nil) def unparsed(source: String, theory: Boolean): Span = { val kind = if (theory) Theory_Span else Malformed_Span Span(kind, List(Token(Token.Kind.UNPARSED, source))) } } diff --git a/src/Pure/PIDE/document.scala b/src/Pure/PIDE/document.scala --- a/src/Pure/PIDE/document.scala +++ b/src/Pure/PIDE/document.scala @@ -1,1264 +1,1269 @@ /* Title: Pure/PIDE/document.scala Author: Makarius Document as collection of named nodes, each consisting of an editable list of commands, associated with asynchronous execution process. */ package isabelle import scala.collection.mutable object Document { /** document structure **/ /* overlays -- print functions with arguments */ object Overlays { val empty = new Overlays(Map.empty) } 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 = { val node_overlays = f(apply(name)) new Overlays(if (node_overlays.is_empty) rep - name else rep + (name -> node_overlays)) } def insert(command: Command, fn: String, args: List[String]): Overlays = update(command.node_name, _.insert(command, fn, args)) def remove(command: Command, fn: String, args: List[String]): Overlays = update(command.node_name, _.remove(command, fn, args)) override def toString: String = rep.mkString("Overlays(", ",", ")") } /* document blobs: auxiliary files */ sealed case class Blob(bytes: Bytes, source: String, chunk: Symbol.Text_Chunk, changed: Boolean) { def unchanged: Blob = copy(changed = false) } 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]) { def get(name: Node.Name): Option[Blob] = blobs.get(name) def changed(name: Node.Name): Boolean = get(name) match { case Some(blob) => blob.changed case None => false } override def toString: String = blobs.mkString("Blobs(", ",", ")") } /* document nodes: theories and auxiliary files */ type Edit[A, B] = (Node.Name, Node.Edit[A, B]) type Edit_Text = Edit[Text.Edit, Text.Perspective] type Edit_Command = Edit[Command.Edit, Command.Perspective] 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) { def imports_offset: Map[Int, Name] = (for { (name, Position.Offset(i)) <- imports_pos } yield i -> name).toMap def imports: List[Name] = imports_pos.map(_._1) def append_errors(msgs: List[String]): Header = copy(errors = errors ::: msgs) def cat_errors(msg2: String): Header = copy(errors = errors.map(msg1 => Exn.cat_message(msg1, msg2))) } val no_header: Header = Header() def bad_header(msg: String): Header = Header(errors = List(msg)) object Name { val empty: Name = Name("") object Ordering extends scala.math.Ordering[Name] { def compare(name1: Name, name2: Name): Int = name1.node compare name2.node } type Graph[A] = isabelle.Graph[Node.Name, A] def make_graph[A](entries: List[((Name, A), List[Name])]): Graph[A] = Graph.make(entries, symmetric = true)(Ordering) } 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 { case other: Name => node == other.node case _ => false } def path: Path = Path.explode(File.standard_path(node)) def master_dir_path: Path = Path.explode(File.standard_path(master_dir)) def expand: Name = Name(path.expand.implode, master_dir_path.expand.implode, theory) def symbolic: Name = Name(path.implode_symbolic, master_dir_path.implode_symbolic, theory) def is_theory: Boolean = theory.nonEmpty def theory_base_name: String = Long_Name.base_name(theory) override def toString: String = if (is_theory) theory else node def map(f: String => String): Name = copy(f(node), f(master_dir), theory) def map_theory(f: String => String): Name = copy(node, master_dir, f(theory)) def json: JSON.Object.T = JSON.Object("node_name" -> node, "theory_name" -> theory) } 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 } /* node overlays */ object Overlays { val empty = new Overlays(Multi_Map.empty) } 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 def insert(cmd: Command, fn: String, args: List[String]): Overlays = new Overlays(rep.insert(cmd, (fn, args))) def remove(cmd: Command, fn: String, args: List[String]): Overlays = new Overlays(rep.remove(cmd, (fn, args))) override def toString: String = rep.mkString("Node.Overlays(", ",", ")") } /* edits */ sealed abstract class Edit[A, B] { def foreach(f: A => Unit): Unit = { this match { case Edits(es) => es.foreach(f) case _ => } } def is_void: Boolean = this match { case Edits(Nil) => true case _ => false } } case class Blob[A, B](blob: Document.Blob) extends Edit[A, B] case class Edits[A, B](edits: List[A]) extends Edit[A, B] case class Deps[A, B](header: Header) extends Edit[A, B] case class Perspective[A, B](required: Boolean, visible: B, overlays: Overlays) extends Edit[A, B] /* perspective */ type Perspective_Text = Perspective[Text.Edit, Text.Perspective] type Perspective_Command = Perspective[Command.Edit, Command.Perspective] val no_perspective_text: Perspective_Text = Perspective(false, Text.Perspective.empty, Overlays.empty) val no_perspective_command: Perspective_Command = Perspective(false, Command.Perspective.empty, Overlays.empty) def is_no_perspective_text(perspective: Perspective_Text): Boolean = !perspective.required && perspective.visible.is_empty && perspective.overlays.is_empty def is_no_perspective_command(perspective: Perspective_Command): Boolean = !perspective.required && perspective.visible.is_empty && perspective.overlays.is_empty /* 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)] = { var i = offset for (command <- commands) yield { val start = i i += command.length (command, start) } } 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 - p = (p /: command.span.content)(_.advance(_)) + p = command.span.content.foldLeft(p)(_.advance(_)) (command, start) } } private val block_size = 256 } 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) = { val blocks = new mutable.ListBuffer[(Command, Text.Offset)] var next_block = 0 var last_stop = 0 for ((command, start) <- Commands.starts(commands.iterator)) { last_stop = start + command.length while (last_stop + 1 > next_block) { blocks += (command -> start) next_block += Commands.block_size } } (blocks.toArray, Text.Range(0, last_stop)) } private def full_range: Text.Range = full_index._2 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 { case (cmd, start) => start + cmd.length <= i } } else Iterator.empty } } val empty: Node = new Node() } final class Node private( val get_blob: Option[Document.Blob] = None, val header: Node.Header = Node.no_header, 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) { def is_empty: Boolean = get_blob.isEmpty && header == Node.no_header && text_perspective.is_empty && Node.is_no_perspective_command(perspective) && commands.isEmpty def has_header: Boolean = header != Node.no_header override def toString: String = if (is_empty) "empty" else if (get_blob.isDefined) "blob" else "node" def commands: Linear_Set[Command] = _commands.commands def load_commands: List[Command] = _commands.load_commands def load_commands_changed(doc_blobs: Blobs): Boolean = load_commands.exists(_.blobs_changed(doc_blobs)) def init_blob(blob: Blob): Node = new Node(get_blob = Some(blob.unchanged)) def update_header(new_header: Node.Header): Node = new Node(get_blob, new_header, syntax, text_perspective, perspective, _commands) def update_syntax(new_syntax: Option[Outer_Syntax]): Node = new Node(get_blob, header, new_syntax, text_perspective, perspective, _commands) def update_perspective( new_text_perspective: Text.Perspective, new_perspective: Node.Perspective_Command): Node = new Node(get_blob, header, syntax, new_text_perspective, new_perspective, _commands) def edit_perspective: Node.Edit[Text.Edit, Text.Perspective] = Node.Perspective(perspective.required, text_perspective, perspective.overlays) def same_perspective( other_text_perspective: Text.Perspective, other_perspective: Node.Perspective_Command): Boolean = text_perspective == other_text_perspective && perspective.required == other_perspective.required && perspective.visible.same(other_perspective.visible) && perspective.overlays == other_perspective.overlays def update_commands(new_commands: Linear_Set[Command]): Node = if (new_commands eq _commands.commands) this else new Node(get_blob, header, syntax, text_perspective, perspective, Node.Commands(new_commands)) def command_iterator(i: Text.Offset = 0): Iterator[(Command, Text.Offset)] = _commands.iterator(i) def command_iterator(range: Text.Range): Iterator[(Command, Text.Offset)] = command_iterator(range.start) takeWhile { case (_, start) => start < range.stop } def command_start(cmd: Command): Option[Text.Offset] = Node.Commands.starts(commands.iterator).find(_._1 == cmd).map(_._2) def source: String = get_blob match { case Some(blob) => blob.source case None => command_iterator(0).map({ case (cmd, _) => cmd.source }).mkString } } /* development graph */ object Nodes { val empty: Nodes = new Nodes(Graph.empty(Node.Name.Ordering)) } 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 = { val graph1 = graph.default_node(name, Node.empty) graph1.is_maximal(name) && graph1.get_node(name).is_empty } def purge_suppressed: Option[Nodes] = graph.keys_iterator.filter(is_suppressed).toList match { case Nil => None - case del => Some(new Nodes((graph /: del)(_.del_node(_)))) + case del => Some(new Nodes(del.foldLeft(graph)(_.del_node(_)))) } def + (entry: (Node.Name, Node)): Nodes = { val (name, node) = entry val imports = node.header.imports val graph1 = - (graph.default_node(name, Node.empty) /: imports)((g, p) => g.default_node(p, Node.empty)) - val graph2 = (graph1 /: graph1.imm_preds(name))((g, dep) => g.del_edge(dep, name)) - val graph3 = (graph2 /: imports)((g, dep) => g.add_edge(dep, name)) + imports.foldLeft(graph.default_node(name, Node.empty)) { + case (g, p) => g.default_node(p, Node.empty) + } + val graph2 = + graph1.imm_preds(name).foldLeft(graph1) { case (g, dep) => g.del_edge(dep, name) } + val graph3 = imports.foldLeft(graph2) { case (g, dep) => g.add_edge(dep, name) } new Nodes(graph3.map_node(name, _ => node)) } def domain: Set[Node.Name] = graph.domain def iterator: Iterator[(Node.Name, Node)] = graph.iterator.map({ case (name, (node, _)) => (name, node) }) def theory_name(theory: String): Option[Node.Name] = graph.keys_iterator.find(name => name.theory == theory) def commands_loading(file_name: Node.Name): List[Command] = (for { (_, node) <- iterator cmd <- node.load_commands.iterator name <- cmd.blobs_names.iterator if name == file_name } yield cmd).toList def descendants(names: List[Node.Name]): List[Node.Name] = graph.all_succs(names) def requirements(names: List[Node.Name]): List[Node.Name] = graph.all_preds_rev(names) def topological_order: List[Node.Name] = graph.topological_order override def toString: String = topological_order.mkString("Nodes(", ",", ")") } /** versioning **/ /* particular document versions */ 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] = { if (future.is_finished) { val version = future.join versions.get(version.id) match { case Some(version1) if !(version eq version1) => Future.value(version1) case _ => future } } else future } def purge_suppressed( versions: Map[Document_ID.Version, Version]): Map[Document_ID.Version, Version] = { - (versions /: - (for ((id, v) <- versions.iterator; v1 <- v.purge_suppressed) yield (id, v1)))(_ + _) + (for ((id, v) <- versions.iterator; v1 <- v.purge_suppressed) yield (id, v1)). + foldLeft(versions)(_ + _) } } final class Version private( val id: Document_ID.Version = Document_ID.none, val nodes: Nodes = Nodes.empty) { override def toString: String = "Version(" + id + ")" def purge_suppressed: Option[Version] = nodes.purge_suppressed.map(new Version(id, _)) } /* changes of plain text, eventually resulting in document edits */ object Change { val init: Change = new Change() def make(previous: Future[Version], edits: List[Edit_Text], version: Future[Version]): Change = new Change(Some(previous), edits.reverse, version) } 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)) { 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] = { val previous1 = previous.map(Version.purge_future(versions, _)) val version1 = Version.purge_future(versions, version) if ((previous eq previous1) && (version eq version1)) None else Some(new Change(previous1, rev_edits, version1)) } } /* history navigation */ object History { val init: History = new History() } final class History private( 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)] = { val n = undo_list.iterator.zipWithIndex.find(p => check(p._1)).get._2 + 1 val (retained, dropped) = undo_list.splitAt(n max retain) retained.splitAt(retained.length - 1) match { case (prefix, List(last)) => Some(dropped, new History(prefix ::: List(last.truncate))) case _ => None } } 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)) } } /* snapshot: persistent user-view of document state */ object Snapshot { val init: Snapshot = State.init.snapshot() } class Snapshot private[Document]( val state: State, val version: Version, val node_name: Node.Name, edits: List[Text.Edit], val snippet_command: Option[Command]) { override def toString: String = "Snapshot(node = " + node_name.node + ", version = " + version.id + (if (is_outdated) ", outdated" else "") + ")" /* nodes */ def get_node(name: Node.Name): Node = version.nodes(name) val node: Node = get_node(node_name) def node_files: List[Node.Name] = node_name :: node.load_commands.flatMap(_.blobs_names) /* edits */ def is_outdated: Boolean = edits.nonEmpty private lazy val reverse_edits = edits.reverse def convert(offset: Text.Offset): Text.Offset = - (offset /: edits)((i, edit) => edit.convert(i)) + edits.foldLeft(offset) { case (i, edit) => edit.convert(i) } def revert(offset: Text.Offset): Text.Offset = - (offset /: reverse_edits)((i, edit) => edit.revert(i)) + reverse_edits.foldLeft(offset) { case (i, edit) => edit.revert(i) } def convert(range: Text.Range): Text.Range = range.map(convert) def revert(range: Text.Range): Text.Range = range.map(revert) /* theory load commands */ val commands_loading: List[Command] = if (node_name.is_theory) Nil else version.nodes.commands_loading(node_name) def commands_loading_ranges(pred: Node.Name => Boolean): List[Text.Range] = (for { cmd <- node.load_commands.iterator blob_name <- cmd.blobs_names.iterator if pred(blob_name) start <- node.command_start(cmd) } yield convert(cmd.core_range + start)).toList /* command as add-on snippet */ def snippet(command: Command): Snapshot = { val node_name = command.node_name val nodes0 = version.nodes val nodes1 = nodes0 + (node_name -> nodes0(node_name).update_commands(Linear_Set(command))) val version1 = Document.Version.make(nodes1) val edits: List[Edit_Text] = List(node_name -> Node.Edits(List(Text.Edit.insert(0, command.source)))) val state0 = state.define_command(command) val state1 = state0.continue_history(Future.value(version), edits, Future.value(version1)) .define_version(version1, state0.the_assignment(version)) .assign(version1.id, Nil, List(command.id -> List(Document_ID.make())))._2 state1.snapshot(node_name = node_name, snippet_command = Some(command)) } /* XML markup */ def xml_markup( range: Text.Range = Text.Range.full, 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)] = { snippet_command match { case None => Nil case Some(command) => for (Exn.Res(blob) <- command.blobs) yield { val bytes = blob.read_file val text = bytes.text val xml = if (Bytes(text) == bytes) { val markup = command.init_markups(Command.Markup_Index.blob(blob)) markup.to_XML(Text.Range(0, text.length), text, elements) } else Nil blob.src_path -> xml } } } /* messages */ lazy val messages: List[(XML.Elem, Position.T)] = (for { (command, start) <- Document.Node.Commands.starts_pos( node.commands.iterator, Token.Pos.file(node_name.node)) pos = command.span.keyword_pos(start).position(command.span.name) (_, elem) <- state.command_results(version, command).iterator } yield (elem, pos)).toList /* exports */ lazy val exports: List[Export.Entry] = state.node_exports(version, node_name).iterator.map(_._2).toList lazy val exports_map: Map[String, Export.Entry] = (for (entry <- exports.iterator) yield (entry.name, entry)).toMap /* find command */ def find_command(id: Document_ID.Generic): Option[(Node, Command)] = state.lookup_id(id) match { case None => None case Some(st) => val command = st.command val command_node = get_node(command.node_name) if (command_node.commands.contains(command)) Some((command_node, command)) else None } def find_command_position(id: Document_ID.Generic, offset: Symbol.Offset) : Option[Line.Node_Position] = for ((node, command) <- find_command(id)) yield { val name = command.node_name.node val sources_iterator = node.commands.iterator.takeWhile(_ != command).map(_.source) ++ (if (offset == 0) Iterator.empty else Iterator.single(command.source(Text.Range(0, command.chunk.decode(offset))))) - val pos = (Line.Position.zero /: sources_iterator)(_.advance(_)) + val pos = sources_iterator.foldLeft(Line.Position.zero)(_.advance(_)) Line.Node_Position(name, pos) } def current_command(other_node_name: Node.Name, offset: Text.Offset): Option[Command] = if (other_node_name.is_theory) { val other_node = get_node(other_node_name) val iterator = other_node.command_iterator(revert(offset) max 0) if (iterator.hasNext) { val (command0, _) = iterator.next() other_node.commands.reverse.iterator(command0).find(command => !command.is_ignored) } else other_node.commands.reverse.iterator.find(command => !command.is_ignored) } else version.nodes.commands_loading(other_node_name).headOption /* command results */ 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)) def command_results(command: Command): Command.Results = state.command_results(version, command) /* command ids: static and dynamic */ def command_id_map: Map[Document_ID.Generic, Command] = state.command_id_map(version, get_node(node_name).commands) /* cumulate markup */ def cumulate[A]( range: Text.Range, info: A, elements: Markup.Elements, result: List[Command.State] => (A, Text.Markup) => Option[A], status: Boolean = false): List[Text.Info[A]] = { val former_range = revert(range).inflate_singularity val (chunk_name, command_iterator) = commands_loading.headOption match { case None => (Symbol.Text_Chunk.Default, node.command_iterator(former_range)) case Some(command) => (Symbol.Text_Chunk.File(node_name.node), Iterator((command, 0))) } val markup_index = Command.Markup_Index(status, chunk_name) (for { (command, command_start) <- command_iterator chunk <- command.chunks.get(chunk_name).iterator states = state.command_states(version, command) res = result(states) markup_range <- (former_range - command_start).try_restrict(chunk.range).iterator markup = Command.State.merge_markup(states, markup_index, markup_range, elements) Text.Info(r0, a) <- markup.cumulate[A](markup_range, info, elements, { case (a, Text.Info(r0, b)) => res(a, Text.Info(convert(r0 + command_start), b)) }).iterator r1 <- convert(r0 + command_start).try_restrict(range).iterator } yield Text.Info(r1, a)).toList } def select[A]( 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]] = { val res = result(states) (_: Option[A], x: Text.Markup) => res(x) match { case None => None case some => Some(some) } } for (Text.Info(r, Some(x)) <- cumulate(range, None, elements, result1, status)) yield Text.Info(r, x) } } /* model */ trait Session { def resources: Resources } trait Model { def session: Session def is_stable: Boolean def snapshot(): Snapshot def node_name: Node.Name def is_theory: Boolean = node_name.is_theory override def toString: String = node_name.toString def get_text(range: Text.Range): Option[String] def node_required: Boolean def get_blob: Option[Blob] def bibtex_entries: List[Text.Info[String]] def node_edits( node_header: Node.Header, text_edits: List[Text.Edit], perspective: Node.Perspective_Text): List[Edit_Text] = { val edits: List[Node.Edit[Text.Edit, Text.Perspective]] = get_blob match { case None => List( Node.Deps( if (session.resources.session_base.loaded_theory(node_name)) { node_header.append_errors( List("Cannot update finished theory " + quote(node_name.theory))) } else node_header), Node.Edits(text_edits), perspective) case Some(blob) => List(Node.Blob(blob), Node.Edits(text_edits)) } edits.flatMap(edit => if (edit.is_void) None else Some(node_name -> edit)) } } /** global state -- document structure, execution process, editing history **/ type Assign_Update = List[(Document_ID.Command, List[Document_ID.Exec])] // update of exec state assignment object State { class Fail(state: State) extends Exception 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) { 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 = { require(!is_finished, "assignment already finished") val command_execs1 = - (command_execs /: update) { + update.foldLeft(command_execs) { case (res, (command_id, exec_ids)) => if (exec_ids.isEmpty) res - command_id else res + (command_id -> exec_ids) } new Assignment(command_execs1, true) } } val init: State = State().define_version(Version.init, Assignment.init).assign(Version.init.id, Nil, Nil)._2 } final case class State private( /*reachable versions*/ versions: Map[Document_ID.Version, Version] = Map.empty, /*inlined auxiliary files*/ blobs: Set[SHA1.Digest] = Set.empty, /*loaded theories in batch builds*/ theories: Map[Document_ID.Exec, Command.State] = Map.empty, /*static markup from define_command*/ commands: Map[Document_ID.Command, Command.State] = Map.empty, /*dynamic markup from execution*/ execs: Map[Document_ID.Exec, Command.State] = Map.empty, /*command-exec assignment for each version*/ assignments: Map[Document_ID.Version, State.Assignment] = Map.empty, /*commands with markup produced by other commands (imm_succs)*/ commands_redirection: Graph[Document_ID.Command, Unit] = Graph.long, /*explicit (linear) history*/ history: History = History.init, /*intermediate state between remove_versions/removed_versions*/ removing_versions: Boolean = false) { override def toString: String = "State(versions = " + versions.size + ", blobs = " + blobs.size + ", commands = " + commands.size + ", execs = " + execs.size + ", assignments = " + assignments.size + ", commands_redirection = " + commands_redirection.size + ", history = " + history.undo_list.size + ", removing_versions = " + removing_versions + ")" private def fail[A]: A = throw new State.Fail(this) def define_version(version: Version, assignment: State.Assignment): State = { val id = version.id copy(versions = versions + (id -> version), assignments = assignments + (id -> assignment.unfinished)) } 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 = { val id = command.id if (commands.isDefinedAt(id)) fail else copy(commands = commands + (id -> command.init_state)) } def defined_command(id: Document_ID.Command): Boolean = commands.isDefinedAt(id) def the_version(id: Document_ID.Version): Version = versions.getOrElse(id, fail) def the_static_state(id: Document_ID.Command): Command.State = commands.getOrElse(id, fail) def the_dynamic_state(id: Document_ID.Exec): Command.State = execs.getOrElse(id, fail) def the_assignment(version: Version): State.Assignment = assignments.getOrElse(version.id, fail) def lookup_id(id: Document_ID.Generic): Option[Command.State] = theories.get(id) orElse commands.get(id) orElse execs.get(id) private def self_id(st: Command.State)(id: Document_ID.Generic): Boolean = 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)] = { for { st <- lookup_id(id) if st.command.node_name == node_name } yield (Symbol.Text_Chunk.Id(st.command.id), st.command.chunk) } private def redirection(st: Command.State): Graph[Document_ID.Command, Unit] = - (commands_redirection /: st.markups.redirection_iterator)({ case (graph, id) => - graph.default_node(id, ()).default_node(st.command.id, ()).add_edge(id, st.command.id) }) + st.markups.redirection_iterator.foldLeft(commands_redirection) { + case (graph, id) => + 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) = { val st1 = st.accumulate(self_id(st), other_id, message, cache) (st1, copy(commands_redirection = redirection(st1))) } execs.get(id).map(update) match { case Some((st1, state1)) => (st1, state1.copy(execs = execs + (id -> st1))) case None => commands.get(id).map(update) match { case Some((st1, state1)) => (st1, state1.copy(commands = commands + (id -> st1))) case None => theories.get(id).map(update) match { case Some((st1, state1)) => (st1, state1.copy(theories = theories + (id -> st1))) case None => fail } } } } 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 { case Some(new_st) => (new_st, copy(execs = execs + (id -> new_st))) case None => fail } case None => commands.get(id) match { case Some(st) => st.add_export(entry) match { case Some(new_st) => (new_st, copy(commands = commands + (id -> new_st))) case None => fail } case None => fail } } } def node_exports(version: Version, node_name: Node.Name): Command.Exports = Command.Exports.merge( for { command <- version.nodes(node_name).commands.iterator st <- command_states(version, command).iterator } yield st.exports) def begin_theory( node_name: Node.Name, id: Document_ID.Exec, source: String, blobs_info: Command.Blobs_Info): State = { if (theories.isDefinedAt(id)) fail else { val command = Command.unparsed(source, theory = true, id = id, node_name = node_name, blobs_info = blobs_info) copy(theories = theories + (id -> command.empty_state)) } } def end_theory(id: Document_ID.Exec): (Snapshot, State) = theories.get(id) match { case None => fail case Some(st) => val command = st.command val node_name = command.node_name val command1 = Command.unparsed(command.source, theory = true, id = id, node_name = node_name, blobs_info = command.blobs_info, results = st.results, markups = st.markups) val state1 = copy(theories = theories - id) (state1.snippet(command1), state1) } 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 val edited_nodes = (for { (name, _) <- version.nodes.iterator if edited_set(name.node) } yield name).toList def upd(exec_id: Document_ID.Exec, st: Command.State) : Option[(Document_ID.Exec, Command.State)] = if (execs.isDefinedAt(exec_id)) None else Some(exec_id -> st) val (changed_commands, new_execs) = - ((Nil: List[Command], execs) /: update) { + update.foldLeft((Nil: List[Command], execs)) { case ((commands1, execs1), (command_id, exec)) => val st = the_static_state(command_id) val command = st.command val commands2 = command :: commands1 val execs2 = exec match { case Nil => execs1 case eval_id :: print_ids => execs1 ++ upd(eval_id, st) ++ (for (id <- print_ids; up <- upd(id, command.empty_state)) yield up) } (commands2, execs2) } val new_assignment = the_assignment(version).assign(update) val new_state = copy(assignments = assignments + (id -> new_assignment), execs = new_execs) ((edited_nodes, changed_commands), new_state) } def is_assigned(version: Version): Boolean = assignments.get(version.id) match { case Some(assgn) => assgn.is_finished case None => false } def is_stable(change: Change): Boolean = change.is_finished && is_assigned(change.version.get_finished) def recent_finished: Change = history.undo_list.find(_.is_finished) getOrElse fail def recent_stable: Change = history.undo_list.find(is_stable) getOrElse fail def stable_tip_version: Option[Version] = if (is_stable(history.tip)) Some(history.tip.version.get_finished) else None def continue_history( previous: Future[Version], edits: List[Edit_Text], version: Future[Version]): State = { val change = Change.make(previous, edits, version) copy(history = history + change) } 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) val removing = old_versions.nonEmpty val state1 = copy(history = history1, removing_versions = removing) (old_versions, state1) case None => fail } } def removed_versions(removed: List[Document_ID.Version]): State = { val versions1 = Version.purge_suppressed(versions -- removed) val assignments1 = assignments -- removed var blobs1_names = Set.empty[Node.Name] var blobs1 = Set.empty[SHA1.Digest] var commands1 = Map.empty[Document_ID.Command, Command.State] var execs1 = Map.empty[Document_ID.Exec, Command.State] for { (version_id, version) <- versions1.iterator command_execs = assignments1(version_id).command_execs (_, node) <- version.nodes.iterator command <- node.commands.iterator } { for ((name, digest) <- command.blobs_defined) { blobs1_names += name blobs1 += digest } if (!commands1.isDefinedAt(command.id)) commands.get(command.id).foreach(st => commands1 += (command.id -> st)) for (exec_id <- command_execs.getOrElse(command.id, Nil)) { if (!execs1.isDefinedAt(exec_id)) execs.get(exec_id).foreach(st => execs1 += (exec_id -> st)) } } copy( versions = versions1, blobs = blobs1, commands = commands1, execs = execs1, commands_redirection = commands_redirection.restrict(commands1.keySet), assignments = assignments1, history = history.purge(versions1), removing_versions = false) } 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 { command <- commands.iterator id <- (command.id :: assignment.command_execs.getOrElse(command.id, Nil)).iterator } yield (id -> command)).toMap } 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 { case eval_id :: print_ids => the_dynamic_state(eval_id).maybe_consolidated && !print_ids.exists(print_id => the_dynamic_state(print_id).consolidating) case Nil => false } } catch { case _: State.Fail => false } } 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) .map(id => id -> the_dynamic_state(id)) match { case Nil => fail case res => res } } catch { case _: State.Fail => try { List(command.id -> the_static_state(command.id)) } catch { case _: State.Fail => List(command.id -> command.init_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)) { (for { command_id <- commands_redirection.imm_succs(command.id).iterator (id, st) <- command_states_self(version, the_static_state(command_id).command) if !self.exists(_._1 == id) } yield (id, st)).toMap.valuesIterator.toList } else Nil self.map(_._2) ::: others.flatMap(_.redirect(command)) } def command_results(version: Version, command: Command): Command.Results = Command.State.merge_results(command_states(version, command)) def command_markup(version: Version, command: Command, index: Command.Markup_Index, range: Text.Range, elements: Markup.Elements): Markup_Tree = Command.State.merge_markup(command_states(version, command), index, range, elements) def xml_markup( version: Version, node_name: Node.Name, range: Text.Range = Text.Range.full, 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 (for { command <- node.commands.iterator command_range <- command.range.try_restrict(range).iterator markup = command_markup(version, command, markup_index, command_range, elements) tree <- markup.to_XML(command_range, command.source, elements).iterator } yield tree).toList } else { val node_source = node.source Text.Range(0, node_source.length).try_restrict(range) match { case None => Nil case Some(node_range) => val markup = version.nodes.commands_loading(node_name).headOption match { case None => Markup_Tree.empty case Some(command) => val chunk_name = Symbol.Text_Chunk.File(node_name.node) val markup_index = Command.Markup_Index(false, chunk_name) command_markup(version, command, markup_index, node_range, elements) } markup.to_XML(node_range, node_source, elements) } } } def node_initialized(version: Version, name: Node.Name): Boolean = name.is_theory && (version.nodes(name).commands.iterator.find(_.potentially_initialized) match { case None => false case Some(command) => command_states(version, command).headOption.exists(_.initialized) }) def node_maybe_consolidated(version: Version, name: Node.Name): Boolean = name.is_theory && version.nodes(name).commands.reverse.iterator.forall(command_maybe_consolidated(version, _)) def node_consolidated(version: Version, name: Node.Name): Boolean = !name.is_theory || { val it = version.nodes(name).commands.reverse.iterator it.hasNext && command_states(version, it.next()).exists(_.consolidated) } def snapshot( node_name: Node.Name = Node.Name.empty, pending_edits: List[Text.Edit] = Nil, snippet_command: Option[Command] = None): Snapshot = { val stable = recent_stable val version = stable.version.get_finished val rev_pending_changes = for { change <- history.undo_list.takeWhile(_ != stable) (name, edits) <- change.rev_edits if name == node_name } yield edits val edits = - (pending_edits /: rev_pending_changes)({ + rev_pending_changes.foldLeft(pending_edits) { case (edits, Node.Edits(es)) => es ::: edits case (edits, _) => edits - }) + } new Snapshot(this, version, node_name, edits, snippet_command) } def snippet(command: Command): Snapshot = snapshot().snippet(command) } } diff --git a/src/Pure/PIDE/document_status.scala b/src/Pure/PIDE/document_status.scala --- a/src/Pure/PIDE/document_status.scala +++ b/src/Pure/PIDE/document_status.scala @@ -1,306 +1,306 @@ /* Title: Pure/PIDE/document_status.scala Author: Makarius Document status based on markup information. */ package isabelle object Document_Status { /* 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) val liberal_elements: Markup.Elements = proper_elements + Markup.WARNING + Markup.LEGACY + Markup.ERROR def make(markup_iterator: Iterator[Markup]): Command_Status = { var touched = false var accepted = false var warned = false var failed = false var canceled = false var finalized = false var forks = 0 var runs = 0 for (markup <- markup_iterator) { markup.name match { case Markup.ACCEPTED => accepted = true case Markup.FORKED => touched = true; forks += 1 case Markup.JOINED => forks -= 1 case Markup.RUNNING => touched = true; runs += 1 case Markup.FINISHED => runs -= 1 case Markup.WARNING | Markup.LEGACY => warned = true case Markup.FAILED | Markup.ERROR => failed = true case Markup.CANCELED => canceled = true case Markup.FINALIZED => finalized = true case _ => } } Command_Status( touched = touched, accepted = accepted, warned = warned, failed = failed, canceled = canceled, finalized = finalized, forks = forks, runs = runs) } val empty: Command_Status = make(Iterator.empty) def merge(status_iterator: Iterator[Command_Status]): Command_Status = if (status_iterator.hasNext) { val status0 = status_iterator.next() - (status0 /: status_iterator)(_ + _) + status_iterator.foldLeft(status0)(_ + _) } else empty } sealed case class Command_Status( private val touched: Boolean, private val accepted: Boolean, private val warned: Boolean, private val failed: Boolean, private val canceled: Boolean, private val finalized: Boolean, forks: Int, runs: Int) { def + (that: Command_Status): Command_Status = Command_Status( touched = touched || that.touched, accepted = accepted || that.accepted, warned = warned || that.warned, failed = failed || that.failed, canceled = canceled || that.canceled, finalized = finalized || that.finalized, forks = forks + that.forks, runs = runs + that.runs) def is_unprocessed: Boolean = accepted && !failed && (!touched || (forks != 0 && runs == 0)) def is_running: Boolean = runs != 0 def is_warned: Boolean = warned def is_failed: Boolean = failed def is_finished: Boolean = !failed && touched && forks == 0 && runs == 0 def is_canceled: Boolean = canceled def is_finalized: Boolean = finalized def is_terminated: Boolean = canceled || touched && forks == 0 && runs == 0 } /* node status */ object Node_Status { def make( state: Document.State, version: Document.Version, name: Document.Node.Name): Node_Status = { val node = version.nodes(name) var unprocessed = 0 var running = 0 var warned = 0 var failed = 0 var finished = 0 var canceled = false var terminated = true var finalized = false for (command <- node.commands.iterator) { val states = state.command_states(version, command) val status = Command_Status.merge(states.iterator.map(_.document_status)) if (status.is_running) running += 1 else if (status.is_failed) failed += 1 else if (status.is_warned) warned += 1 else if (status.is_finished) finished += 1 else unprocessed += 1 if (status.is_canceled) canceled = true if (!status.is_terminated) terminated = false if (status.is_finalized) finalized = true } val initialized = state.node_initialized(version, name) val consolidated = state.node_consolidated(version, name) Node_Status( is_suppressed = version.nodes.is_suppressed(name), unprocessed = unprocessed, running = running, warned = warned, failed = failed, finished = finished, canceled = canceled, terminated = terminated, initialized = initialized, finalized = finalized, consolidated = consolidated) } } sealed case class Node_Status( is_suppressed: Boolean, unprocessed: Int, running: Int, warned: Int, failed: Int, finished: Int, canceled: Boolean, terminated: Boolean, initialized: Boolean, finalized: Boolean, consolidated: Boolean) { def ok: Boolean = failed == 0 def total: Int = unprocessed + running + warned + failed + finished def quasi_consolidated: Boolean = !is_suppressed && !finalized && terminated def percentage: Int = if (consolidated) 100 else if (total == 0) 0 else (((total - unprocessed).toDouble / total) * 100).toInt min 99 def json: JSON.Object.T = JSON.Object("ok" -> ok, "total" -> total, "unprocessed" -> unprocessed, "running" -> running, "warned" -> warned, "failed" -> failed, "finished" -> finished, "canceled" -> canceled, "consolidated" -> consolidated, "percentage" -> percentage) } /* 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 = { var total = 0.0 var command_timings = Map.empty[Command, Double] for { command <- commands.iterator st <- state.command_states(version, command) } { val command_timing = - (0.0 /: st.status)({ + st.status.foldLeft(0.0) { case (timing, Markup.Timing(t)) => timing + t.elapsed.seconds case (timing, _) => timing - }) + } total += command_timing if (command_timing > 0.0 && command_timing >= threshold) { command_timings += (command -> command_timing) } } Overall_Timing(total, command_timings) } } sealed case class Overall_Timing(total: Double, command_timings: Map[Command, Double]) { def command_timing(command: Command): Double = command_timings.getOrElse(command, 0.0) } /* nodes status */ object Overall_Node_Status extends Enumeration { val ok, failed, pending = Value } 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) { 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) def present: List[(Document.Node.Name, Node_Status)] = (for { name <- nodes.topological_order.iterator node_status <- get(name) } yield (name, node_status)).toList def quasi_consolidated(name: Document.Node.Name): Boolean = rep.get(name) match { case Some(st) => st.quasi_consolidated case None => false } def overall_node_status(name: Document.Node.Name): Overall_Node_Status.Value = rep.get(name) match { case Some(st) if st.consolidated => if (st.ok) Overall_Node_Status.ok else Overall_Node_Status.failed case _ => Overall_Node_Status.pending } def update( resources: Resources, state: Document.State, version: Document.Version, domain: Option[Set[Document.Node.Name]] = None, trim: Boolean = false): (Boolean, Nodes_Status) = { val nodes1 = version.nodes val update_iterator = for { name <- domain.getOrElse(nodes1.domain).iterator if !resources.is_hidden(name) && !resources.session_base.loaded_theory(name) st = Document_Status.Node_Status.make(state, version, name) if !rep.isDefinedAt(name) || rep(name) != st } yield (name -> st) val rep1 = rep ++ update_iterator val rep2 = if (trim) rep1 -- rep1.keysIterator.filterNot(nodes1.domain) else rep1 (rep != rep2, new Nodes_Status(rep2, nodes1)) } override def hashCode: Int = rep.hashCode override def equals(that: Any): Boolean = that match { case other: Nodes_Status => rep == other.rep case _ => false } override def toString: String = { var ok = 0 var failed = 0 var pending = 0 var canceled = 0 for (name <- rep.keysIterator) { overall_node_status(name) match { case Overall_Node_Status.ok => ok += 1 case Overall_Node_Status.failed => failed += 1 case Overall_Node_Status.pending => pending += 1 } if (apply(name).canceled) canceled += 1 } "Nodes_Status(ok = " + ok + ", failed = " + failed + ", pending = " + pending + ", canceled = " + canceled + ")" } } } diff --git a/src/Pure/PIDE/headless.scala b/src/Pure/PIDE/headless.scala --- a/src/Pure/PIDE/headless.scala +++ b/src/Pure/PIDE/headless.scala @@ -1,661 +1,663 @@ /* Title: Pure/PIDE/headless.scala Author: Makarius Headless PIDE session and resources from file-system. */ package isabelle import java.io.{File => JFile} import scala.annotation.tailrec import scala.collection.mutable object Headless { /** session **/ private def stable_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 } class Use_Theories_Result private[Headless]( 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 committed = nodes_committed.iterator.map(_._1).toSet nodes.filter(p => !committed(p._1)) } def snapshot(name: Document.Node.Name): Document.Snapshot = stable_snapshot(state, version, name) def ok: Boolean = (nodes.iterator ++ nodes_committed.iterator).forall({ case (_, st) => st.ok }) } class Session private[Headless]( session_name: String, _session_options: => Options, override val resources: Resources) extends isabelle.Session(_session_options, resources) { session => private def loaded_theory(name: Document.Node.Name): Boolean = resources.session_base.loaded_theory(name.theory) /* options */ override def consolidate_delay: Time = session_options.seconds("headless_consolidate_delay") override def prune_delay: Time = session_options.seconds("headless_prune_delay") def default_check_delay: Time = session_options.seconds("headless_check_delay") def default_check_limit: Int = session_options.int("headless_check_limit") def default_nodes_status_delay: Time = session_options.seconds("headless_nodes_status_delay") def default_watchdog_timeout: Time = session_options.seconds("headless_watchdog_timeout") def default_commit_cleanup_delay: Time = session_options.seconds("headless_commit_cleanup_delay") /* temporary directory */ val tmp_dir: JFile = Isabelle_System.tmp_dir("server_session") val tmp_dir_name: String = File.path(tmp_dir).implode def master_directory(master_dir: String): String = proper_string(master_dir) getOrElse tmp_dir_name override def toString: String = session_name override def stop(): Process_Result = { try { super.stop() } finally { Isabelle_System.rm_tree(tmp_dir) } } /* theories */ private object Load_State { def finished: Load_State = Load_State(Nil, Nil, 0) def count_file(name: Document.Node.Name): Long = if (loaded_theory(name)) 0 else name.path.file.length } private case class Load_State( 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) = { val load_theories = dep_graph.all_preds_rev(pending1).filterNot(finished) (load_theories, Load_State(pending1, rest1, load_limit)) } if (!pending.forall(finished)) (Nil, this) else if (rest.isEmpty) (Nil, Load_State.finished) else if (load_limit == 0) load_requirements(rest, Nil) else { val reachable = dep_graph.reachable_limit(load_limit, Load_State.count_file, dep_graph.imm_preds, rest) val (pending1, rest1) = rest.partition(reachable) load_requirements(pending1, rest1) } } } private sealed case class Use_Theories_State( dep_graph: Document.Node.Name.Graph[Unit], load_state: Load_State, watchdog_timeout: Time, commit: Option[(Document.Snapshot, Document_Status.Node_Status) => Unit], 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) { def update(new_nodes_status: Document_Status.Nodes_Status): Use_Theories_State = copy(last_update = Time.now(), nodes_status = new_nodes_status) def watchdog: Boolean = watchdog_timeout > Time.zero && Time.now() - last_update > watchdog_timeout def finished_result: Boolean = result.isDefined def join_result: Option[(Exn.Result[Use_Theories_Result], Use_Theories_State)] = if (finished_result) Some((result.get, this)) else None 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] = { val add = base.filter(name => dep_graph.imm_succs(name).forall(front)) if (add.isEmpty) front else { val preds = add.map(dep_graph.imm_preds) - val base1 = (preds.head /: preds.tail)(_ ++ _).toList.filter(already_committed.keySet) + val base1 = + preds.tail.foldLeft(preds.head)(_ ++ _).toList.filter(already_committed.keySet) frontier(base1, front ++ add) } } if (already_committed.isEmpty) (Nil, this) else { val base = (for { (name, (_, (_, succs))) <- dep_graph.iterator if succs.isEmpty && already_committed.isDefinedAt(name) } yield name).toList val clean = frontier(base, Set.empty) if (clean.isEmpty) (Nil, this) else { (dep_graph.topological_order.filter(clean), copy(dep_graph = dep_graph.exclude(clean))) } } } 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 case Some(commit_fn) => - (already_committed /: dep_graph.topological_order)( - { case (committed, name) => - def parents_committed: Boolean = - version.nodes(name).header.imports.forall(parent => - loaded_theory(parent) || committed.isDefinedAt(parent)) - if (!committed.isDefinedAt(name) && parents_committed && - 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) - committed + (name -> status) - } - else committed - }) + dep_graph.topological_order.foldLeft(already_committed) { + case (committed, name) => + def parents_committed: Boolean = + version.nodes(name).header.imports.forall(parent => + loaded_theory(parent) || committed.isDefinedAt(parent)) + if (!committed.isDefinedAt(name) && parents_committed && + 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) + committed + (name -> status) + } + else committed + } } def finished_theory(name: Document.Node.Name): Boolean = loaded_theory(name) || (if (commit.isDefined) already_committed1.isDefinedAt(name) else state.node_consolidated(version, name)) val result1 = if (!finished_result && (beyond_limit || watchdog || dep_graph.keys_iterator.forall(name => finished_theory(name) || nodes_status.quasi_consolidated(name)))) { val nodes = (for { name <- dep_graph.keys_iterator if !loaded_theory(name) } yield { (name -> Document_Status.Node_Status.make(state, version, name)) }).toList val nodes_committed = (for { name <- dep_graph.keys_iterator status <- already_committed1.get(name) } yield (name -> status)).toList Some(Exn.Res(new Use_Theories_Result(state, version, nodes, nodes_committed))) } else result val (load_theories, load_state1) = load_state.next(dep_graph, finished_theory) (load_theories, copy(already_committed = already_committed1, result = result1, load_state = load_state1)) } } def use_theories( theories: List[String], qualifier: String = Sessions.DRAFT, master_dir: String = "", unicode_symbols: Boolean = false, check_delay: Time = default_check_delay, check_limit: Int = default_check_limit, watchdog_timeout: Time = default_watchdog_timeout, nodes_status_delay: Time = default_nodes_status_delay, id: UUID.T = UUID.random(), // 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 = { val import_names = theories.map(thy => resources.import_name(qualifier, master_directory(master_dir), thy) -> Position.none) resources.dependencies(import_names, progress = progress).check_errors } val dep_theories = dependencies.theories val dep_theories_set = dep_theories.toSet val dep_files = for (path <- dependencies.loaded_files) yield Document.Node.Name(resources.append("", path)) val use_theories_state = { val dep_graph = dependencies.theory_graph val maximals = dep_graph.maximals val rest = if (maximals.isEmpty || maximals.tail.isEmpty) maximals else { val depth = dep_graph.node_depth(Load_State.count_file) maximals.sortBy(node => - depth(node)) } val load_limit = if (commit.isDefined) (session_options.real("headless_load_limit") * 1024 * 1024).round else 0 val load_state = Load_State(Nil, rest, load_limit) Synchronized(Use_Theories_State(dep_graph, load_state, watchdog_timeout, commit)) } def check_state(beyond_limit: Boolean = false): Unit = { val state = session.get_state() for { version <- state.stable_tip_version load_theories = use_theories_state.change_result(_.check(state, version, beyond_limit)) if load_theories.nonEmpty } resources.load_theories(session, id, load_theories, dep_files, unicode_symbols, 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) } } } val consumer = { val delay_nodes_status = Delay.first(nodes_status_delay max Time.zero) { progress.nodes_status(use_theories_state.value.nodes_status) } val delay_commit_clean = Delay.first(commit_cleanup_delay max Time.zero) { val clean_theories = use_theories_state.change_result(_.clean_theories) if (clean_theories.nonEmpty) { progress.echo("Removing " + clean_theories.length + " theories ...") resources.clean_theories(session, id, clean_theories) } } Session.Consumer[Session.Commands_Changed](getClass.getName) { case changed => if (changed.nodes.exists(dep_theories_set)) { val snapshot = session.snapshot() val state = snapshot.state 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 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 } 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.foreach(progress.theory) check_state() if (commit.isDefined && commit_cleanup_delay > Time.zero) { if (use_theories_state.value.finished_result) delay_commit_clean.revoke else delay_commit_clean.invoke } } } } try { session.commands_changed += consumer check_state() use_theories_state.guarded_access(_.join_result) check_progress.cancel } finally { session.commands_changed -= consumer resources.unload_theories(session, id, dep_theories) } Exn.release(use_theories_state.guarded_access(_.join_result)) } def purge_theories( theories: List[String], qualifier: String = Sessions.DRAFT, master_dir: String = "", 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), _))) resources.purge_theories(session, nodes) } } /** resources **/ object Resources { def apply(options: Options, base_info: Sessions.Base_Info, log: Logger = No_Logger): Resources = new Resources(options, base_info, log = log) def make( options: Options, session_name: String, session_dirs: List[Path] = Nil, include_sessions: List[String] = Nil, progress: Progress = new Progress, log: Logger = No_Logger): Resources = { val base_info = Sessions.base_info(options, session_name, dirs = session_dirs, include_sessions = include_sessions, progress = progress) apply(options, base_info, log = log) } final class Theory private[Headless]( val node_name: Document.Node.Name, val node_header: Document.Node.Header, val text: String, val node_required: Boolean) { override def toString: String = node_name.toString def node_perspective: Document.Node.Perspective_Text = Document.Node.Perspective(node_required, Text.Perspective.empty, Document.Node.Overlays.empty) def make_edits(text_edits: List[Text.Edit]): List[Document.Edit_Text] = List(node_name -> Document.Node.Deps(node_header), node_name -> Document.Node.Edits(text_edits), node_name -> node_perspective) 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) if (text_edits.isEmpty && node_required == old_required) Nil else make_edits(text_edits) } def purge_edits: List[Document.Edit_Text] = make_edits(Text.Edit.removes(0, text)) def required(required: Boolean): Theory = if (required == node_required) this else new Theory(node_name, node_header, text, required) } 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) { /* blobs */ def doc_blobs: Document.Blobs = Document.Blobs(blobs) def update_blobs(names: List[Document.Node.Name]): (Document.Blobs, State) = { val new_blobs = names.flatMap(name => { val bytes = Bytes.read(name.path) def new_blob: Document.Blob = { val text = bytes.text Document.Blob(bytes, text, Symbol.Text_Chunk(text), changed = true) } blobs.get(name) match { case Some(blob) => if (blob.bytes == bytes) None else Some(name -> new_blob) case None => Some(name -> new_blob) } }) - val blobs1 = (blobs /: new_blobs)(_ + _) - val blobs2 = (blobs /: new_blobs)({ case (map, (a, b)) => map + (a -> b.unchanged) }) + val blobs1 = new_blobs.foldLeft(blobs)(_ + _) + val blobs2 = new_blobs.foldLeft(blobs) { case (map, (a, b)) => map + (a -> b.unchanged) } (Document.Blobs(blobs1), copy(blobs = blobs2)) } def blob_edits(name: Document.Node.Name, old_blob: Option[Document.Blob]) : List[Document.Edit_Text] = { val blob = blobs.getOrElse(name, error("Missing blob " + quote(name.toString))) val text_edits = old_blob match { case None => List(Text.Edit.insert(0, blob.source)) case Some(blob0) => Text.Edit.replace(0, blob0.source, blob.source) } if (text_edits.isEmpty) Nil else List(name -> Document.Node.Blob(blob), name -> Document.Node.Edits(text_edits)) } /* theories */ lazy val theory_graph: Document.Node.Name.Graph[Unit] = Document.Node.Name.make_graph( for ((name, theory) <- theories.toList) yield ((name, ()), theory.node_header.imports.filter(theories.isDefinedAt))) def is_required(name: Document.Node.Name): Boolean = required.isDefinedAt(name) def insert_required(id: UUID.T, names: List[Document.Node.Name]): State = - copy(required = (required /: names)(_.insert(_, id))) + copy(required = names.foldLeft(required)(_.insert(_, id))) def remove_required(id: UUID.T, names: List[Document.Node.Name]): State = - copy(required = (required /: names)(_.remove(_, id))) + copy(required = names.foldLeft(required)(_.remove(_, id))) def update_theories(update: List[(Document.Node.Name, Theory)]): State = copy(theories = - (theories /: update)({ case (thys, (name, thy)) => - thys.get(name) match { - case Some(thy1) if thy1 == thy => thys - case _ => thys + (name -> thy) - } - })) + update.foldLeft(theories) { + case (thys, (name, thy)) => + thys.get(name) match { + case Some(thy1) if thy1 == thy => thys + case _ => thys + (name -> thy) + } + }) 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) = { val st1 = remove_required(id, theories) val theory_edits = for { node_name <- theories theory <- st1.theories.get(node_name) } yield { val theory1 = theory.required(st1.is_required(node_name)) val edits = theory1.node_edits(Some(theory)) (edits, (node_name, theory1)) } (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) = { val all_nodes = theory_graph.topological_order val purge = nodes.getOrElse(all_nodes).filterNot(is_required).toSet val retain = theory_graph.all_preds(all_nodes.filterNot(purge)).toSet val (retained, purged) = all_nodes.partition(retain) val purge_edits = purged.flatMap(name => theories(name).purge_edits) ((purged, retained, purge_edits), remove_theories(purged)) } } } class Resources private[Headless]( 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) { resources => val store: Sessions.Store = Sessions.store(options) /* 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 + " ...") Isabelle_Process(session, options, session_base_info.sessions_structure, store, logic = session_base_info.session, modes = print_mode).await_startup session } /* theories */ private val state = Synchronized(Resources.State()) def load_theories( session: Session, id: UUID.T, theories: List[Document.Node.Name], files: List[Document.Node.Name], unicode_symbols: Boolean, progress: Progress): Unit = { val loaded_theories = for (node_name <- theories) yield { val path = node_name.path if (!node_name.is_theory) error("Not a theory file: " + path) progress.expose_interrupt() val text0 = File.read(path) val text = if (unicode_symbols) Symbol.decode(text0) else text0 val node_header = resources.check_thy(node_name, Scan.char_reader(text)) new Resources.Theory(node_name, node_header, text, true) } 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)) 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 => { 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 => { val (edits1, st1) = st.unload_theories(session, id, theories) val ((_, _, edits2), st2) = st1.purge_theories(session, None) session.update(st.doc_blobs, edits1 ::: edits2) st2 }) } def purge_theories(session: Session, nodes: Option[List[Document.Node.Name]]) : (List[Document.Node.Name], List[Document.Node.Name]) = { state.change_result(st => { val ((purged, retained, _), st1) = st.purge_theories(session, nodes) ((purged, retained), st1) }) } } } diff --git a/src/Pure/PIDE/line.scala b/src/Pure/PIDE/line.scala --- a/src/Pure/PIDE/line.scala +++ b/src/Pure/PIDE/line.scala @@ -1,253 +1,253 @@ /* Title: Pure/PIDE/line.scala Author: Makarius Line-oriented text documents, with some bias towards VSCode. */ package isabelle import scala.annotation.tailrec object Line { /* logical lines */ def normalize(text: String): String = if (text.contains('\r')) text.replace("\r\n", "\n") else text def logical_lines(text: String): List[String] = split_lines(normalize(text)) /* position */ object Position { val zero: Position = Position() val offside: Position = Position(line = -1) 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) { def line1: Int = line + 1 def column1: Int = column + 1 def print: String = line1.toString + ":" + column1.toString def compare(that: Position): Int = line compare that.line match { case 0 => column compare that.column case i => i } def advance(text: String): Position = if (text.isEmpty) this else { val lines = logical_lines(text) val l = line + lines.length - 1 val c = (if (l == line) column else 0) + Library.trim_line(lines.last).length Position(l, c) } } /* range (right-open interval) */ object Range { def apply(start: Position): Range = Range(start, start) val zero: Range = Range(Position.zero) } sealed case class Range(start: Position, stop: Position) { if (start.compare(stop) > 0) error("Bad line range: " + start.print + ".." + stop.print) def print: String = if (start == stop) start.print else start.print + ".." + stop.print } /* positions within document node */ 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) { 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) { def start: Position = range.start def stop: Position = range.stop } /* document with newline as separator (not terminator) */ object Document { def apply(text: String): Document = Document(logical_lines(text).map(s => Line(Library.isolate_substring(s)))) val empty: Document = apply("") private def split(line_text: String): List[Line] = if (line_text == "") List(Line.empty) else apply(line_text).lines private def chop(lines: List[Line]): (String, List[Line]) = lines match { case Nil => ("", Nil) case line :: rest => (line.text, rest) } private def length(lines: List[Line]): Int = if (lines.isEmpty) 0 - else ((0 /: lines) { case (n, line) => n + line.text.length + 1 }) - 1 + else (lines.foldLeft(0) { case (n, line) => n + line.text.length + 1 }) - 1 def text(lines: List[Line]): String = lines.mkString("", "\n", "") } 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) lazy val text: String = Document.text(lines) def get_text(range: Text.Range): Option[String] = if (text_range.contains(range)) Some(range.substring(text)) else None override def toString: String = text override def equals(that: Any): Boolean = that match { case other: Document => lines == other.lines case _ => false } 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 = { lines_rest match { case Nil => require(i == 0, "bad Line.position.move"); Position(lines_count) case line :: ls => val n = line.text.length if (ls.isEmpty || i <= n) Position(lines_count).advance(line.text.substring(n - i)) else move(i - (n + 1), lines_count + 1, ls) } } move(text_offset, 0, lines) } def range(text_range: Text.Range): Range = Range(position(text_range.start), position(text_range.stop)) def offset(pos: Position): Option[Text.Offset] = { val l = pos.line val c = pos.column val n = lines.length if (0 <= l && l < n) { if (0 <= c && c <= lines(l).text.length) { val line_offset = - (0 /: lines.iterator.take(l)) { case (n, line) => n + line.text.length + 1 } + lines.iterator.take(l).foldLeft(0) { case (n, line) => n + line.text.length + 1 } Some(line_offset + c) } else None } else if (l == n && c == 0) Some(text_length) else None } 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) <- { val c1 = remove.start.column val c2 = remove.stop.column val (prefix, lines1) = lines.splitAt(l1) val (s1, rest1) = Document.chop(lines1) if (l1 == l2) { if (0 <= c1 && c1 <= c2 && c2 <= s1.length) { Some( if (lines1.isEmpty) ("", prefix ::: Document.split(insert)) else { val removed_text = s1.substring(c1, c2) val changed_text = s1.substring(0, c1) + insert + s1.substring(c2) (removed_text, prefix ::: Document.split(changed_text) ::: rest1) }) } else None } else { val (middle, lines2) = rest1.splitAt(l2 - l1 - 1) val (s2, rest2) = Document.chop(lines2) if (0 <= c1 && c1 <= s1.length && 0 <= c2 && c2 <= s2.length) { Some( if (lines1.isEmpty) ("", prefix ::: Document.split(insert)) else { val r1 = s1.substring(c1) val r2 = s2.substring(0, c2) val removed_text = if (lines2.isEmpty) Document.text(Line(r1) :: middle) else Document.text(Line(r1) :: middle ::: List(Line(r2))) val changed_text = s1.substring(0, c1) + insert + s2.substring(c2) (removed_text, prefix ::: Document.split(changed_text) ::: rest2) }) } else None } } } yield (Text.Edit.removes(edit_start, removed_text) ::: Text.Edit.inserts(edit_start, insert), Document(new_lines)) } } /* line text */ val empty: Line = new Line("") def apply(text: String): Line = if (text == "") empty else new Line(text) } final class Line private(val text: String) { require(text.forall(c => c != '\r' && c != '\n'), "bad line text") override def equals(that: Any): Boolean = that match { case other: Line => text == other.text case _ => false } override def hashCode(): Int = text.hashCode override def toString: String = text } diff --git a/src/Pure/PIDE/markup_tree.scala b/src/Pure/PIDE/markup_tree.scala --- a/src/Pure/PIDE/markup_tree.scala +++ b/src/Pure/PIDE/markup_tree.scala @@ -1,268 +1,271 @@ /* Title: Pure/PIDE/markup_tree.scala Author: Fabian Immler, TU Munich Author: Makarius Markup trees over nested / non-overlapping text ranges. */ package isabelle import scala.collection.immutable.SortedMap import scala.collection.mutable import scala.annotation.tailrec object Markup_Tree { /* construct trees */ val empty: Markup_Tree = new Markup_Tree(Branches.empty) def merge(trees: List[Markup_Tree], range: Text.Range, elements: Markup.Elements): Markup_Tree = - (empty /: trees)(_.merge(_, range, elements)) + trees.foldLeft(empty)(_.merge(_, range, elements)) def merge_disjoint(trees: List[Markup_Tree]): Markup_Tree = trees match { case Nil => empty case head :: tail => new Markup_Tree( - (head.branches /: tail) { + tail.foldLeft(head.branches) { case (branches, tree) => - (branches /: tree.branches) { + tree.branches.foldLeft(branches) { case (bs, (r, entry)) => require(!bs.isDefinedAt(r), "cannot merge markup trees") bs + (r -> entry) } }) } /* tree building blocks */ object Entry { def apply(markup: Text.Markup, subtree: Markup_Tree): Entry = Entry(markup.range, List(markup.info), subtree) } sealed case class Entry( range: Text.Range, rev_markup: List[XML.Elem], subtree: Markup_Tree) { def markup: List[XML.Elem] = rev_markup.reverse 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 result } def + (markup: Text.Markup): Entry = copy(rev_markup = markup.info :: rev_markup) def \ (markup: Text.Markup): Entry = copy(subtree = subtree + markup) } object Branches { type T = SortedMap[Text.Range, Entry] val empty: T = SortedMap.empty(Text.Range.Ordering) } /* XML representation */ @tailrec private def strip_elems( elems: List[XML.Elem], body: XML.Body): (List[XML.Elem], XML.Body) = body match { case List(XML.Wrapped_Elem(markup1, body1, body2)) => strip_elems(XML.Elem(markup1, body1) :: elems, body2) case List(XML.Elem(markup1, body1)) => strip_elems(XML.Elem(markup1, Nil) :: elems, body1) case _ => (elems, body) } 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) case (elems, body) => - val (end_offset, subtrees) = ((offset, Nil: List[Markup_Tree]) /: body)(make_trees) + val (end_offset, subtrees) = + body.foldLeft((offset, Nil: List[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(((0, Nil: List[Markup_Tree]) /: body)(make_trees)._2) + merge_disjoint(body.foldLeft((0, Nil: List[Markup_Tree]))(make_trees)._2) } 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) = this(branches + (entry.range -> entry)) private def overlapping(range: Text.Range): Branches.T = if (branches.isEmpty || (range.contains(branches.firstKey.start) && branches.lastKey.stop <= range.stop)) branches else { val start = Text.Range(range.start) val stop = Text.Range(range.stop) val bs = branches.range(start, stop) branches.get(stop) match { case Some(end) if range overlaps end.range => bs + (end.range -> end) case _ => bs } } def restrict(range: Text.Range): Markup_Tree = new Markup_Tree(overlapping(range)) def is_empty: Boolean = branches.isEmpty def + (new_markup: Text.Markup): Markup_Tree = { val new_range = new_markup.range branches.get(new_range) match { case None => new Markup_Tree(branches, Entry(new_markup, empty)) case Some(entry) => if (entry.range == new_range) new Markup_Tree(branches, entry + new_markup) else if (entry.range.contains(new_range)) new Markup_Tree(branches, entry \ new_markup) else if (new_range.contains(branches.head._1) && new_range.contains(branches.last._1)) new Markup_Tree(Branches.empty, Entry(new_markup, this)) else { val body = overlapping(new_range) if (body.forall(e => new_range.contains(e._1))) new Markup_Tree(branches -- body.keys, Entry(new_markup, new Markup_Tree(body))) else { Output.warning("Ignored overlapping markup information: " + new_markup + "\n" + body.filter(e => !new_range.contains(e._1)).values.mkString("\n")) this } } } } 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 = - (tree1 /: tree2.branches)( - { case (tree, (range, entry)) => - if (!range.overlaps(root_range)) tree - else - (merge_trees(tree, entry.subtree) /: entry.filter_markup(elements))( - { case (t, elem) => t + Text.Info(range, elem) }) - }) + tree2.branches.foldLeft(tree1) { + case (tree, (range, entry)) => + if (!range.overlaps(root_range)) tree + else { + entry.filter_markup(elements).foldLeft(merge_trees(tree, entry.subtree)) { + case (t, elem) => t + Text.Info(range, elem) + } + } + } if (this eq other) this else { val tree1 = this.restrict(root_range) val tree2 = other.restrict(root_range) if (tree1.is_empty) tree2 else merge_trees(tree1, tree2) } } 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 { elem <- entry.filter_markup(elements) y1 <- result(y, Text.Info(entry.range, elem)) } { y = y1; changed = true } if (changed) Some(y) else None } def traverse( last: Text.Offset, 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) val subtree = entry.subtree.overlapping(subrange).toList val start = subrange.start results(parent.info, entry) match { case Some(res) => val next = Text.Info(subrange, res) val nexts = traverse(start, (next, subtree) :: (parent, more) :: rest) if (last < start) parent.restrict(Text.Range(last, start)) :: nexts else nexts case None => traverse(last, (parent, subtree ::: more) :: rest) } case (parent, Nil) :: rest => val stop = parent.range.stop val nexts = traverse(stop, rest) if (last < stop) parent.restrict(Text.Range(last, stop)) :: nexts else nexts case Nil => val stop = root_range.stop if (last < stop) List(Text.Info(Text.Range(last, stop), root_info)) else Nil } } traverse(root_range.start, 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 make_text(start: Text.Offset, stop: Text.Offset): XML.Body = if (start == stop) Nil else List(XML.Text(text.subSequence(start, stop).toString)) def make_elems(rev_markups: List[XML.Elem], body: XML.Body): XML.Body = - (body /: rev_markups) { + rev_markups.foldLeft(body) { case (b, elem) => if (!elements(elem.name)) b else if (elem.body.isEmpty) List(XML.Elem(elem.markup, b)) 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 = { val body = new mutable.ListBuffer[XML.Tree] var last = elem_range.start for ((range, entry) <- entries) { val subrange = range.restrict(elem_range) body ++= make_text(last, subrange.start) body ++= make_body(subrange, entry.rev_markup, entry.subtree.overlapping(subrange)) last = subrange.stop } body ++= make_text(last, elem_range.stop) make_elems(elem_markup, body.toList) } make_body(root_range, Nil, overlapping(root_range)) } override def toString: String = branches.toList.map(_._2) match { case Nil => "Empty" case list => list.mkString("Tree(", ",", ")") } } diff --git a/src/Pure/PIDE/protocol_handlers.scala b/src/Pure/PIDE/protocol_handlers.scala --- a/src/Pure/PIDE/protocol_handlers.scala +++ b/src/Pure/PIDE/protocol_handlers.scala @@ -1,85 +1,86 @@ /* Title: Pure/PIDE/protocol_handlers.scala Author: Makarius Management of add-on protocol handlers for PIDE session. */ package isabelle 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) { def get(name: String): Option[Session.Protocol_Handler] = handlers.get(name) def init(handler: Session.Protocol_Handler): State = { val name = handler.getClass.getName try { if (handlers.isDefinedAt(name)) error("Duplicate protocol handler: " + name) else { handler.init(session) val dups = for ((a, _) <- handler.functions if functions.isDefinedAt(a)) yield a if (dups.nonEmpty) error("Duplicate protocol functions: " + commas_quote(dups)) copy(handlers = handlers + (name -> handler), functions = functions ++ handler.functions) } } catch { case exn: Throwable => err_handler(exn, name) } } def init(name: String): State = { val handler = try { Class.forName(name).getDeclaredConstructor().newInstance() .asInstanceOf[Session.Protocol_Handler] } catch { case exn: Throwable => err_handler(exn, name) } init(handler) } def invoke(msg: Prover.Protocol_Output): Boolean = msg.properties match { case (Markup.FUNCTION, a) :: _ if functions.isDefinedAt(a) => try { functions(a)(msg) } catch { case exn: Throwable => Output.error_message( "Failed invocation of protocol function: " + quote(a) + "\n" + Exn.message(exn)) false } case _ => false } def exit(): State = { for ((_, handler) <- handlers) handler.exit() copy(handlers = Map.empty, functions = Map.empty) } } def init(session: Session): Protocol_Handlers = new Protocol_Handlers(session) } class Protocol_Handlers private(session: Session) { private val state = Synchronized(Protocol_Handlers.State(session)) def prover_options(options: Options): Options = - (options /: state.value.handlers)( - { case (opts, (_, handler)) => handler.prover_options(opts) }) + state.value.handlers.foldLeft(options) { + case (opts, (_, handler)) => handler.prover_options(opts) + } def get(name: String): Option[Session.Protocol_Handler] = state.value.get(name) def init(handler: Session.Protocol_Handler): Unit = state.change(_.init(handler)) def init(name: String): Unit = state.change(_.init(name)) def invoke(msg: Prover.Protocol_Output): Boolean = state.value.invoke(msg) def exit(): Unit = state.change(_.exit()) } diff --git a/src/Pure/PIDE/prover.scala b/src/Pure/PIDE/prover.scala --- a/src/Pure/PIDE/prover.scala +++ b/src/Pure/PIDE/prover.scala @@ -1,368 +1,368 @@ /* Title: Pure/PIDE/prover.scala Author: Makarius Options: :folding=explicit: Prover process wrapping. */ package isabelle import java.io.{InputStream, OutputStream, BufferedOutputStream, IOException} object Prover { /* messages */ sealed abstract class Message type Receiver = Message => Unit 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 { def kind: String = message.markup.name def properties: Properties.T = message.markup.properties def body: XML.Body = message.body def is_init: Boolean = kind == Markup.INIT def is_exit: Boolean = kind == Markup.EXIT def is_stdout: Boolean = kind == Markup.STDOUT def is_stderr: Boolean = kind == Markup.STDERR def is_system: Boolean = kind == Markup.SYSTEM def is_status: Boolean = kind == Markup.STATUS def is_report: Boolean = kind == Markup.REPORT def is_syslog: Boolean = is_init || is_exit || is_system || is_stderr 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) if (properties.isEmpty) kind.toString + " [[" + res + "]]" else kind.toString + " " + (for ((x, y) <- properties) yield x + "=" + y).mkString("{", ",", "}") + " [[" + res + "]]" } } class Protocol_Output(props: Properties.T, val bytes: Bytes) extends Output(XML.Elem(Markup(Markup.PROTOCOL, props), Nil)) { lazy val text: String = bytes.text } } class Prover( receiver: Prover.Receiver, cache: XML.Cache, channel: System_Channel, process: Bash.Process) extends Protocol { /** receiver output **/ 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, bytes: Bytes): Unit = { receiver(new Prover.Protocol_Output(cache.props(props), bytes)) } 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 = { output(Markup.EXIT, Markup.Process_Result(result), List(XML.Text(result.print_return_code))) } /** process manager **/ private val process_result: Future[Process_Result] = Future.thread("process_result") { val rc = process.join val timing = process.get_timing Process_Result(rc, timing = timing) } 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") { val stdout = physical_output(false) 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)) { while (finished.isEmpty && process.stderr.ready) { try { val c = process.stderr.read if (c == 2) finished = Some(true) else result += c.toChar } catch { case _: IOException => finished = Some(false) } } Time.seconds(0.05).sleep } (finished.isEmpty || !finished.get, result.toString.trim) } if (startup_errors != "") system_output(startup_errors) if (startup_failed) { terminate_process() process_result.join stdout.join exit_message(Process_Result(127)) } else { val (command_stream, message_stream) = channel.rendezvous() command_input_init(command_stream) val stderr = physical_output(true) val message = message_output(message_stream) val result = process_result.join system_output("process terminated") command_input_close() for (thread <- List(stdout, stderr, message)) thread.join system_output("process_manager terminated") exit_message(result) } channel.shutdown() } /* management methods */ def join(): Unit = process_manager.join() def terminate(): Unit = { system_output("Terminating prover process") command_input_close() var count = 10 while (!process_result.is_finished && count > 0) { Time.seconds(0.1).sleep count -= 1 } if (!process_result.is_finished) terminate_process() } /** process streams **/ /* command input */ private var command_input: Option[Consumer_Thread[List[Bytes]]] = None private def command_input_close(): Unit = command_input.foreach(_.shutdown) private def command_input_init(raw_stream: OutputStream): Unit = { val name = "command_input" val stream = new BufferedOutputStream(raw_stream) command_input = Some( Consumer_Thread.fork(name)( consume = { case chunks => try { Bytes(chunks.map(_.length).mkString("", ",", "\n")).write_stream(stream) chunks.foreach(_.write_stream(stream)) stream.flush true } catch { case e: IOException => system_output(name + ": " + e.getMessage); false } }, finish = { case () => stream.close; system_output(name + " terminated") } ) ) } /* physical output */ 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) Isabelle_Thread.fork(name = name) { try { var result = new StringBuilder(100) var finished = false while (!finished) { //{{{ var c = -1 var done = false while (!done && (result.isEmpty || reader.ready)) { c = reader.read if (c >= 0) result.append(c.asInstanceOf[Char]) else done = true } if (result.nonEmpty) { output(markup, Nil, List(XML.Text(Symbol.decode(result.toString)))) result.clear } else { reader.close finished = true } //}}} } } catch { case e: IOException => system_output(name + ": " + e.getMessage) } system_output(name + " terminated") } } /* message output */ private def message_output(stream: InputStream): Thread = { class EOF extends Exception class Protocol_Error(msg: String) extends Exception(msg) val name = "message_output" Isabelle_Thread.fork(name = name) { val default_buffer = new Array[Byte](65536) var c = -1 def read_int(): Int = //{{{ { var n = 0 c = stream.read if (c == -1) throw new EOF while (48 <= c && c <= 57) { n = 10 * n + (c - 48) c = stream.read } if (c != 10) throw new Protocol_Error("malformed header: expected integer followed by newline") else n } //}}} def read_chunk_bytes(): (Array[Byte], Int) = //{{{ { val n = read_int() val buf = if (n <= default_buffer.length) default_buffer else new Array[Byte](n) var i = 0 var m = 0 do { m = stream.read(buf, i, n - i) if (m != -1) i += m } while (m != -1 && n > i) if (i != n) throw new Protocol_Error("bad chunk (unexpected EOF after " + i + " of " + n + " bytes)") (buf, n) } //}}} def read_chunk(): XML.Body = { val (buf, n) = read_chunk_bytes() YXML.parse_body_failsafe(UTF8.decode_chars(Symbol.decode, buf, 0, n)) } try { do { try { val header = read_chunk() header match { case List(XML.Elem(Markup(name, props), Nil)) => val kind = name.intern if (kind == Markup.PROTOCOL) { val (buf, n) = read_chunk_bytes() protocol_output(props, Bytes(buf, 0, n)) } else { val body = read_chunk() output(kind, props, body) } case _ => read_chunk() throw new Protocol_Error("bad header: " + header.toString) } } catch { case _: EOF => } } while (c != -1) } catch { case e: IOException => system_output("Cannot read message:\n" + e.getMessage) case e: Protocol_Error => system_output("Malformed message:\n" + e.getMessage) } stream.close system_output(name + " terminated") } } /** protocol commands **/ var trace: Boolean = false def protocol_command_raw(name: String, args: List[Bytes]): Unit = command_input match { case Some(thread) if thread.is_active => if (trace) { - val payload = (0 /: args)({ case (n, b) => n + b.length }) + val payload = args.foldLeft(0) { case (n, b) => n + b.length } Output.writeln( "protocol_command " + name + ", args = " + args.length + ", payload = " + payload) } thread.send(Bytes(name) :: args) case _ => error("Inactive prover input thread for command " + quote(name)) } def protocol_command_args(name: String, args: List[String]): Unit = { receiver(new Prover.Input(name, args)) protocol_command_raw(name, args.map(Bytes(_))) } def protocol_command(name: String, args: String*): Unit = protocol_command_args(name, args.toList) } diff --git a/src/Pure/PIDE/rendering.scala b/src/Pure/PIDE/rendering.scala --- a/src/Pure/PIDE/rendering.scala +++ b/src/Pure/PIDE/rendering.scala @@ -1,792 +1,792 @@ /* Title: Pure/PIDE/rendering.scala Author: Makarius Isabelle-specific implementation of quasi-abstract rendering and markup interpretation. */ package isabelle import java.io.{File => JFile} import java.nio.file.FileSystems import scala.collection.immutable.SortedMap object Rendering { /* color */ object Color extends Enumeration { // background val unprocessed1, running1, canceled, bad, intensify, entity, active, active_result, markdown_bullet1, markdown_bullet2, markdown_bullet3, markdown_bullet4 = Value val background_colors: ValueSet = values // foreground val quoted, antiquoted = Value val foreground_colors: ValueSet = values -- background_colors // message underline val writeln, information, warning, legacy, error = Value val message_underline_colors: ValueSet = values -- background_colors -- foreground_colors // message background val writeln_message, information_message, tracing_message, warning_message, legacy_message, error_message = Value val message_background_colors: ValueSet = values -- background_colors -- foreground_colors -- message_underline_colors // text val main, keyword1, keyword2, keyword3, quasi_keyword, improper, operator, tfree, tvar, free, skolem, bound, `var`, inner_numeral, inner_quoted, inner_cartouche, comment1, comment2, comment3, dynamic, class_parameter, antiquote, raw_text, plain_text = Value val text_colors: ValueSet = values -- background_colors -- foreground_colors -- message_underline_colors -- message_background_colors // text overview val unprocessed, running = Value val text_overview_colors = Set(unprocessed, running, error, warning) } /* output messages */ val state_pri = 1 val writeln_pri = 2 val information_pri = 3 val tracing_pri = 4 val warning_pri = 5 val legacy_pri = 6 val error_pri = 7 val message_pri = Map( Markup.STATE -> state_pri, Markup.STATE_MESSAGE -> state_pri, Markup.WRITELN -> writeln_pri, Markup.WRITELN_MESSAGE -> writeln_pri, Markup.INFORMATION -> information_pri, Markup.INFORMATION_MESSAGE -> information_pri, Markup.TRACING -> tracing_pri, Markup.TRACING_MESSAGE -> tracing_pri, Markup.WARNING -> warning_pri, Markup.WARNING_MESSAGE -> warning_pri, Markup.LEGACY -> legacy_pri, Markup.LEGACY_MESSAGE -> legacy_pri, Markup.ERROR -> error_pri, Markup.ERROR_MESSAGE -> error_pri ).withDefaultValue(0) val message_underline_color = Map( writeln_pri -> Color.writeln, information_pri -> Color.information, warning_pri -> Color.warning, legacy_pri -> Color.legacy, error_pri -> Color.error) val message_background_color = Map( writeln_pri -> Color.writeln_message, information_pri -> Color.information_message, tracing_pri -> Color.tracing_message, warning_pri -> Color.warning_message, legacy_pri -> Color.legacy_message, error_pri -> Color.error_message) 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) states ::: other } /* text color */ val text_color = Map( Markup.KEYWORD1 -> Color.keyword1, Markup.KEYWORD2 -> Color.keyword2, Markup.KEYWORD3 -> Color.keyword3, Markup.QUASI_KEYWORD -> Color.quasi_keyword, Markup.IMPROPER -> Color.improper, Markup.OPERATOR -> Color.operator, Markup.STRING -> Color.main, Markup.ALT_STRING -> Color.main, Markup.VERBATIM -> Color.main, Markup.CARTOUCHE -> Color.main, Markup.LITERAL -> Color.keyword1, Markup.DELIMITER -> Color.main, Markup.TFREE -> Color.tfree, Markup.TVAR -> Color.tvar, Markup.FREE -> Color.free, Markup.SKOLEM -> Color.skolem, Markup.BOUND -> Color.bound, Markup.VAR -> Color.`var`, Markup.INNER_STRING -> Color.inner_quoted, Markup.INNER_CARTOUCHE -> Color.inner_cartouche, Markup.DYNAMIC_FACT -> Color.dynamic, Markup.CLASS_PARAMETER -> Color.class_parameter, Markup.ANTIQUOTE -> Color.antiquote, Markup.RAW_TEXT -> Color.raw_text, Markup.PLAIN_TEXT -> Color.plain_text, Markup.ML_KEYWORD1 -> Color.keyword1, Markup.ML_KEYWORD2 -> Color.keyword2, Markup.ML_KEYWORD3 -> Color.keyword3, Markup.ML_DELIMITER -> Color.main, Markup.ML_NUMERAL -> Color.inner_numeral, Markup.ML_CHAR -> Color.inner_quoted, Markup.ML_STRING -> Color.inner_quoted, Markup.ML_COMMENT -> Color.comment1, Markup.COMMENT -> Color.comment1, Markup.COMMENT1 -> Color.comment1, Markup.COMMENT2 -> Color.comment2, Markup.COMMENT3 -> Color.comment3) val foreground = Map( Markup.STRING -> Color.quoted, Markup.ALT_STRING -> Color.quoted, Markup.VERBATIM -> Color.quoted, Markup.CARTOUCHE -> Color.quoted, Markup.ANTIQUOTED -> Color.antiquoted) /* tooltips */ val tooltip_descriptions = Map( Markup.CITATION -> "citation", Markup.TOKEN_RANGE -> "inner syntax token", Markup.FREE -> "free variable", Markup.SKOLEM -> "skolem variable", Markup.BOUND -> "bound variable", Markup.VAR -> "schematic variable", Markup.TFREE -> "free type variable", Markup.TVAR -> "schematic type variable") /* entity 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 = - (empty /: args)({ case (focus1, Text.Info(_, focus2)) => focus1 ++ focus2 }) + args.foldLeft(empty) { case (focus1, Text.Info(_, focus2)) => focus1 ++ focus2 } val full: Focus = 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]) { 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) def ++ (other: Focus): Focus = if (this eq other) this else if (rep.isEmpty) other - else (this /: other.rep.iterator)(_ + _) + else other.rep.iterator.foldLeft(this)(_ + _) override def toString: String = rep.mkString("Focus(", ",", ")") } /* markup elements */ val position_elements = Markup.Elements(Markup.BINDING, Markup.ENTITY, Markup.REPORT, Markup.POSITION) val semantic_completion_elements = Markup.Elements(Markup.COMPLETION, Markup.NO_COMPLETION) val language_context_elements = Markup.Elements(Markup.STRING, Markup.ALT_STRING, Markup.VERBATIM, Markup.CARTOUCHE, Markup.COMMENT, Markup.LANGUAGE, Markup.ML_STRING, Markup.ML_COMMENT) val language_elements = Markup.Elements(Markup.LANGUAGE) val citation_elements = Markup.Elements(Markup.CITATION) val active_elements = Markup.Elements(Markup.DIALOG, Markup.BROWSER, Markup.GRAPHVIEW, Markup.THEORY_EXPORTS, Markup.SENDBACK, Markup.JEDIT_ACTION, Markup.SIMP_TRACE_PANEL) val background_elements = Document_Status.Command_Status.proper_elements + Markup.WRITELN_MESSAGE + Markup.STATE_MESSAGE + Markup.INFORMATION_MESSAGE + Markup.TRACING_MESSAGE + Markup.WARNING_MESSAGE + Markup.LEGACY_MESSAGE + Markup.ERROR_MESSAGE + Markup.BAD + Markup.INTENSIFY + Markup.ENTITY + Markup.Markdown_Bullet.name ++ active_elements val foreground_elements = Markup.Elements(foreground.keySet) val text_color_elements = Markup.Elements(text_color.keySet) val tooltip_elements = Markup.Elements(Markup.LANGUAGE, Markup.EXPRESSION, Markup.TIMING, Markup.ENTITY, Markup.SORTING, Markup.TYPING, Markup.CLASS_PARAMETER, Markup.ML_TYPING, Markup.ML_BREAKPOINT, Markup.PATH, Markup.DOC, Markup.URL, Markup.MARKDOWN_PARAGRAPH, Markup.MARKDOWN_ITEM, Markup.Markdown_List.name) ++ Markup.Elements(tooltip_descriptions.keySet) val tooltip_message_elements = Markup.Elements(Markup.WRITELN, Markup.INFORMATION, Markup.WARNING, Markup.LEGACY, Markup.ERROR, Markup.BAD) val message_elements = Markup.Elements(message_pri.keySet) val warning_elements = Markup.Elements(Markup.WARNING, Markup.LEGACY) val error_elements = Markup.Elements(Markup.ERROR) val entity_elements = Markup.Elements(Markup.ENTITY) val antiquoted_elements = Markup.Elements(Markup.ANTIQUOTED) val meta_data_elements = Markup.Elements(Markup.META_TITLE, Markup.META_CREATOR, Markup.META_CONTRIBUTOR, Markup.META_DATE, Markup.META_DESCRIPTION, Markup.META_LICENSE) val document_tag_elements = Markup.Elements(Markup.Document_Tag.ELEMENT) val markdown_elements = Markup.Elements(Markup.MARKDOWN_PARAGRAPH, Markup.MARKDOWN_ITEM, Markup.Markdown_List.name, Markup.Markdown_Bullet.name) } class Rendering( val snapshot: Document.Snapshot, val options: Options, val session: Session) { override def toString: String = "Rendering(" + snapshot.toString + ")" def get_text(range: Text.Range): Option[String] = None /* caret */ def before_caret_range(caret: Text.Offset): Text.Range = { val former_caret = snapshot.revert(caret) snapshot.convert(Text.Range(former_caret - 1, former_caret)) } /* completion */ def semantic_completion(completed_range: Option[Text.Range], caret_range: Text.Range) : Option[Text.Info[Completion.Semantic]] = if (snapshot.is_outdated) None else { snapshot.select(caret_range, Rendering.semantic_completion_elements, _ => { case Completion.Semantic.Info(info) => completed_range match { case Some(range0) if range0.contains(info.range) && range0 != info.range => None case _ => Some(info) } case _ => None }).headOption.map(_.info) } def semantic_completion_result( history: Completion.History, unicode: Boolean, completed_range: Option[Text.Range], 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)) => get_text(range) match { case Some(original) => (false, names.complete(range, history, unicode, original)) case None => (false, None) } case None => (false, None) } } def language_context(range: Text.Range): Option[Completion.Language_Context] = snapshot.select(range, Rendering.language_context_elements, _ => { case Text.Info(_, XML.Elem(Markup.Language(language, symbols, antiquotes, delimited), _)) => if (delimited) Some(Completion.Language_Context(language, symbols, antiquotes)) else None case Text.Info(_, elem) if elem.name == Markup.ML_STRING || elem.name == Markup.ML_COMMENT => Some(Completion.Language_Context.ML_inner) case Text.Info(_, _) => Some(Completion.Language_Context.inner) }).headOption.map(_.info) def citations(range: Text.Range): List[Text.Info[String]] = snapshot.select(range, Rendering.citation_elements, _ => { case Text.Info(info_range, XML.Elem(Markup.Citation(name), _)) => Some(Text.Info(snapshot.convert(info_range), name)) case _ => None }).map(_.info) /* file-system path completion */ def language_path(range: Text.Range): Option[Text.Info[Boolean]] = snapshot.select(range, Rendering.language_elements, _ => { case Text.Info(info_range, XML.Elem(Markup.Language.Path(delimited), _)) => Some((delimited, snapshot.convert(info_range))) 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])] = { try { val path = Path.explode(text) val (dir, base_name) = if (text == "" || text.endsWith("/")) (path, "") else (path.dir, path.file_name) val directory = new JFile(session.resources.append(snapshot.node_name, dir)) val files = directory.listFiles if (files == null) Nil else { val ignore = space_explode(':', options.string("completion_path_ignore")). map(s => FileSystems.getDefault.getPathMatcher("glob:" + s)) (for { file <- files.iterator name = file.getName if name.startsWith(base_name) path_name = new JFile(name).toPath if !ignore.exists(matcher => matcher.matches(path_name)) text1 = (dir + Path.basic(name)).implode_short if text != text1 is_dir = new JFile(directory, name).isDirectory replacement = text1 + (if (is_dir) "/" else "") descr = List(text1, if (is_dir) "(directory)" else "(file)") } yield (replacement, descr)).take(options.int("completion_limit")).toList } } catch { case ERROR(_) => Nil } } def is_wrapped(s: String): Boolean = s.startsWith("\"") && s.endsWith("\"") || s.startsWith(Symbol.open_decoded) && s.endsWith(Symbol.close_decoded) for { Text.Info(r1, delimited) <- language_path(before_caret_range(caret)) s1 <- get_text(r1) (r2, s2) <- if (is_wrapped(s1)) { Some((Text.Range(r1.start + 1, r1.stop - 1), s1.substring(1, s1.length - 1))) } else if (delimited) Some((r1, s1)) else None if Path.is_valid(s2) paths = complete(s2) if paths.nonEmpty items = paths.map(p => Completion.Item(r2, s2, "", p._2, p._1, 0, false)) } yield Completion.Result(r2, s2, false, items) } /* spell checker */ private lazy val spell_checker_include = Markup.Elements(space_explode(',', options.string("spell_checker_include")): _*) private lazy val spell_checker_elements = spell_checker_include ++ Markup.Elements(space_explode(',', options.string("spell_checker_exclude")): _*) def spell_checker(range: Text.Range): List[Text.Info[Text.Range]] = { val result = snapshot.select(range, spell_checker_elements, _ => { case info => Some( if (spell_checker_include(info.info.name)) Some(snapshot.convert(info.range)) else None) }) for (Text.Info(range, Some(range1)) <- result) yield Text.Info(range, range1) } def spell_checker_point(range: Text.Range): Option[Text.Range] = spell_checker(range).headOption.map(_.info) /* text background */ 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])]( range, (List(Markup.Empty), None), elements, command_states => { case ((markups, color), Text.Info(_, XML.Elem(markup, _))) if markups.nonEmpty && Document_Status.Command_Status.proper_elements(markup.name) => Some((markup :: markups, color)) case (_, Text.Info(_, XML.Elem(Markup(Markup.BAD, _), _))) => Some((Nil, Some(Rendering.Color.bad))) case (_, Text.Info(_, XML.Elem(Markup(Markup.INTENSIFY, _), _))) => Some((Nil, Some(Rendering.Color.intensify))) case (_, Text.Info(_, XML.Elem(Markup.Entity.Occ(i), _))) if focus(i) => Some((Nil, Some(Rendering.Color.entity))) case (_, Text.Info(_, XML.Elem(Markup.Markdown_Bullet(depth), _))) => val color = depth % 4 match { case 1 => Rendering.Color.markdown_bullet1 case 2 => Rendering.Color.markdown_bullet2 case 3 => Rendering.Color.markdown_bullet3 case _ => Rendering.Color.markdown_bullet4 } Some((Nil, Some(color))) case (_, Text.Info(_, Protocol.Dialog(_, serial, result))) => command_states.collectFirst( { case st if st.results.defined(serial) => st.results.get(serial).get }) match { case Some(Protocol.Dialog_Result(res)) if res == result => Some((Nil, Some(Rendering.Color.active_result))) case _ => Some((Nil, Some(Rendering.Color.active))) } case (_, Text.Info(_, elem)) => if (Rendering.active_elements(elem.name)) Some((Nil, Some(Rendering.Color.active))) else None }) color <- result match { case (markups, opt_color) if markups.nonEmpty => val status = Document_Status.Command_Status.make(markups.iterator) if (status.is_unprocessed) Some(Rendering.Color.unprocessed1) else if (status.is_running) Some(Rendering.Color.running1) else if (status.is_canceled) Some(Rendering.Color.canceled) else opt_color case (_, opt_color) => opt_color } } yield Text.Info(r, color) } /* text foreground */ def foreground(range: Text.Range): List[Text.Info[Rendering.Color.Value]] = snapshot.select(range, Rendering.foreground_elements, _ => { case info => Some(Rendering.foreground(info.info.name)) }) /* entity focus */ def entity_focus_defs(range: Text.Range): Rendering.Focus = Rendering.Focus.make( snapshot.cumulate(range, Rendering.Focus.empty, Rendering.entity_elements, _ => { case (focus, Text.Info(_, XML.Elem(Markup.Entity.Def(i), _))) => Some(focus + i) case _ => None })) def entity_focus(range: Text.Range, defs_focus: Rendering.Focus): Rendering.Focus = Rendering.Focus.make( snapshot.cumulate(range, Rendering.Focus.empty, Rendering.entity_elements, _ => { case (focus, Text.Info(_, XML.Elem(Markup.Entity.Ref(i), _))) if defs_focus(i) => Some(focus + i) case _ => None })) /* caret 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 else { val defs_focus = if (defs_range == Text.Range.full) Rendering.Focus.full else entity_focus_defs(defs_range) entity_focus(caret_range, defs_focus) } } 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, _ => { case (_, Text.Info(_, XML.Elem(Markup.Entity.Occ(i), _))) if focus(i) => Some(true) case _ => None }).flatMap(info => if (info.info) Some(info.range) else None) } else Nil } /* messages */ def message_underline_color(elements: Markup.Elements, range: Text.Range) : List[Text.Info[Rendering.Color.Value]] = { val results = snapshot.cumulate[Int](range, 0, elements, _ => { case (pri, Text.Info(_, elem)) => Some(pri max Rendering.message_pri(elem.name)) }) for { Text.Info(r, pri) <- results color <- Rendering.message_underline_color.get(pri) } yield Text.Info(r, color) } 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 => { case (res, Text.Info(_, elem)) => Command.State.get_result_proper(command_states, elem.markup.properties) .map(res :+ _) case _ => None }) var seen_serials = Set.empty[Long] def seen(i: Long): Boolean = { val b = seen_serials(i) seen_serials += i b } for { Text.Info(range, entries) <- results (i, elem) <- entries if !seen(i) } yield Text.Info(range, elem) } /* tooltips */ def timing_threshold: Double = 0.0 private sealed case class Tooltip_Info( range: Text.Range, timing: Timing = Timing.zero, messages: List[(Long, 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 = { 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 = { val r = snapshot.convert(r0) if (range == r) copy(rev_infos = (important -> tree) :: rev_infos) else copy (range = r, rev_infos = List(important -> tree)) } def timing_info(tree: XML.Tree): Option[XML.Tree] = tree match { case XML.Elem(Markup(Markup.TIMING, _), _) => if (timing.elapsed.seconds >= timing_threshold) Some(XML.Text(timing.message)) else None case _ => Some(tree) } def infos(important: Boolean): List[XML.Tree] = for { (is_important, tree) <- rev_infos.reverse if is_important == important tree1 <- timing_info(tree) } yield tree1 } 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]]] = { val results = snapshot.cumulate[Tooltip_Info](range, Tooltip_Info(range), elements, command_states => { case (info, Text.Info(_, XML.Elem(Markup.Timing(t), _))) => Some(info + t) case (info, Text.Info(r0, msg @ XML.Elem(Markup(Markup.BAD, Markup.Serial(i)), body))) if body.nonEmpty => Some(info + (r0, i, msg)) case (info, Text.Info(r0, XML.Elem(Markup(name, props), _))) if Rendering.tooltip_message_elements(name) => for ((i, tree) <- Command.State.get_result_proper(command_states, props)) yield (info + (r0, i, tree)) case (info, Text.Info(r0, XML.Elem(Markup.Entity(kind, name), _))) if kind != "" && kind != Markup.ML_DEF => val kind1 = Word.implode(Word.explode('_', kind)) val txt1 = if (name == "") kind1 else if (kind1 == "") quote(name) else kind1 + " " + quote(name) val info1 = info + (r0, true, XML.Text(txt1)) Some(if (kind == Markup.COMMAND) info1 + (r0, true, XML.elem(Markup.TIMING)) else info1) case (info, Text.Info(r0, XML.Elem(Markup.Path(name), _))) => val file = perhaps_append_file(snapshot.node_name, name) val text = if (name == file) "file " + quote(file) else "path " + quote(name) + "\nfile " + quote(file) Some(info + (r0, true, XML.Text(text))) case (info, Text.Info(r0, XML.Elem(Markup.Doc(name), _))) => val text = "doc " + quote(name) Some(info + (r0, true, XML.Text(text))) case (info, Text.Info(r0, XML.Elem(Markup.Url(name), _))) => Some(info + (r0, true, XML.Text("URL " + quote(name)))) case (info, Text.Info(r0, XML.Elem(Markup(name, _), body))) if name == Markup.SORTING || name == Markup.TYPING => Some(info + (r0, true, Pretty.block(XML.Text("::") :: Pretty.brk(1) :: body))) case (info, Text.Info(r0, XML.Elem(Markup(Markup.CLASS_PARAMETER, _), body))) => Some(info + (r0, true, Pretty.block(0, body))) case (info, Text.Info(r0, XML.Elem(Markup(Markup.ML_TYPING, _), body))) => Some(info + (r0, false, Pretty.block(XML.Text("ML:") :: Pretty.brk(1) :: body))) case (info, Text.Info(r0, Protocol.ML_Breakpoint(breakpoint))) => val text = if (session.debugger.breakpoint_state(breakpoint)) "breakpoint (enabled)" else "breakpoint (disabled)" Some(info + (r0, true, XML.Text(text))) case (info, Text.Info(r0, XML.Elem(Markup.Language(language, _, _, _), _))) => val lang = Word.implode(Word.explode('_', language)) Some(info + (r0, true, XML.Text("language: " + lang))) case (info, Text.Info(r0, XML.Elem(Markup.Expression(kind), _))) => val descr = if (kind == "") "expression" else "expression: " + kind Some(info + (r0, true, XML.Text(descr))) case (info, Text.Info(r0, XML.Elem(Markup(Markup.MARKDOWN_PARAGRAPH, _), _))) => Some(info + (r0, true, XML.Text("Markdown: paragraph"))) case (info, Text.Info(r0, XML.Elem(Markup(Markup.MARKDOWN_ITEM, _), _))) => Some(info + (r0, true, XML.Text("Markdown: item"))) case (info, Text.Info(r0, XML.Elem(Markup.Markdown_List(kind), _))) => Some(info + (r0, true, XML.Text("Markdown: " + kind))) case (info, Text.Info(r0, XML.Elem(Markup(name, _), _))) => Rendering.tooltip_descriptions.get(name).map(desc => info + (r0, true, XML.Text(desc))) }).map(_.info) if (results.isEmpty) None else { val r = Text.Range(results.head.range.start, results.last.range.stop) val all_tips = - (SortedMap.empty[Long, XML.Tree] /: results.flatMap(_.messages))(_ + _) + results.flatMap(_.messages).foldLeft(SortedMap.empty[Long, XML.Tree])(_ + _) .iterator.map(_._2).toList ::: results.flatMap(res => res.infos(true)) ::: results.flatMap(res => res.infos(false)).lastOption.toList if (all_tips.isEmpty) None else Some(Text.Info(r, all_tips)) } } /* messages */ def warnings(range: Text.Range): List[Text.Markup] = snapshot.select(range, Rendering.warning_elements, _ => Some(_)).map(_.info) def errors(range: Text.Range): List[Text.Markup] = snapshot.select(range, Rendering.error_elements, _ => Some(_)).map(_.info) /* command status overview */ def overview_color(range: Text.Range): Option[Rendering.Color.Value] = { if (snapshot.is_outdated) None else { val results = snapshot.cumulate[List[Markup]](range, Nil, Document_Status.Command_Status.liberal_elements, _ => { case (status, Text.Info(_, elem)) => Some(elem.markup :: status) }, status = true) if (results.isEmpty) None else { val status = Document_Status.Command_Status.make(results.iterator.flatMap(_.info)) if (status.is_running) Some(Rendering.Color.running) else if (status.is_failed) Some(Rendering.Color.error) else if (status.is_warned) Some(Rendering.Color.warning) else if (status.is_unprocessed) Some(Rendering.Color.unprocessed) else None } } } /* antiquoted text */ def antiquoted(range: Text.Range): Option[Text.Range] = snapshot.cumulate[Option[Text.Range]](range, None, Rendering.antiquoted_elements, _ => { case (res, info) => if (res.isEmpty) Some(Some(info.range)) else None }).headOption.flatMap(_.info) /* meta data */ def meta_data(range: Text.Range): Properties.T = { val results = snapshot.cumulate[Properties.T](range, Nil, Rendering.meta_data_elements, _ => { case (res, Text.Info(_, elem)) => val plain_text = XML.content(elem) Some((elem.name -> plain_text) :: res) }) Library.distinct(results.flatMap(_.info.reverse)) } /* document tags */ def document_tags(range: Text.Range): List[String] = { val results = snapshot.cumulate[List[String]](range, Nil, Rendering.document_tag_elements, _ => { case (res, Text.Info(_, XML.Elem(Markup.Document_Tag(name), _))) => Some(name :: res) case _ => None }) Library.distinct(results.flatMap(_.info.reverse)) } } diff --git a/src/Pure/PIDE/resources.scala b/src/Pure/PIDE/resources.scala --- a/src/Pure/PIDE/resources.scala +++ b/src/Pure/PIDE/resources.scala @@ -1,449 +1,451 @@ /* Title: Pure/PIDE/resources.scala Author: Makarius Resources for theories and auxiliary files. */ package isabelle import scala.util.parsing.input.Reader import java.io.{File => JFile} object Resources { def empty: Resources = new Resources(Sessions.Structure.empty, Sessions.Structure.empty.bootstrap) } class Resources( val sessions_structure: Sessions.Structure, val session_base: Sessions.Base, val log: Logger = No_Logger, command_timings: List[Properties.T] = Nil) { resources => /* init session */ def init_session_yxml: String = { import XML.Encode._ YXML.string_of_body( pair(list(pair(string, properties)), pair(list(pair(string, string)), pair(list(pair(string, string)), pair(list(pair(string, list(string))), pair(list(properties), pair(list(pair(string, properties)), pair(list(pair(string, string)), list(string))))))))( (sessions_structure.session_positions, (sessions_structure.dest_session_directories, (sessions_structure.session_chapters, (sessions_structure.bibtex_entries, (command_timings, (Scala.functions.map(fun => (fun.name, fun.position)), (session_base.global_theories.toList, session_base.loaded_theories.keys))))))))) } /* file formats */ def make_theory_name(name: Document.Node.Name): Option[Document.Node.Name] = File_Format.registry.get(name).flatMap(_.make_theory_name(resources, name)) def make_theory_content(thy_name: Document.Node.Name): Option[String] = File_Format.registry.get_theory(thy_name).flatMap(_.make_theory_content(resources, thy_name)) def is_hidden(name: Document.Node.Name): Boolean = !name.is_theory || name.theory == Sessions.root_name || File_Format.registry.is_theory(name) def html_document(snapshot: Document.Snapshot): Option[Presentation.HTML_Document] = File_Format.registry.get(snapshot.node_name).flatMap(_.html_document(snapshot)) /* file-system operations */ def append(dir: String, source_path: Path): String = (Path.explode(dir) + source_path).expand.implode 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 = { val node = append(dir, file) val master_dir = append(dir, file.dir) Document.Node.Name(node, master_dir, theory) } def loaded_theory_node(theory: String): Document.Node.Name = Document.Node.Name(theory, "", theory) /* source files of Isabelle/ML bootstrap */ 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 val path = Path.explode(raw_name) val path1 = if (path.is_absolute || path.is_current) check(path) else { check(Path.explode("~~/src/Pure") + path) orElse (if (Isabelle_System.getenv("ML_SOURCES") == "") None else check(Path.explode("$ML_SOURCES") + path)) } Some(File.platform_path(path1 getOrElse path)) } else None } else Some(raw_name) } /* theory files */ 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)) () => { if (syntax.has_load_commands(raw_text)) { val text = Symbol.decode(Scan.reader_decode_utf8(is_utf8, raw_text)) syntax.parse_spans(text).filter(_.is_load_command(syntax)) } else Nil } } 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] = { val pure_dir = Path.explode("~~/src/Pure") for { (name, theory) <- Thy_Header.ml_roots path = (pure_dir + Path.explode(name)).expand node_name = Document.Node.Name(path.implode, path.dir.implode, theory) file <- loaded_files(syntax, node_name, load_commands(syntax, node_name)()) } yield file } def theory_name(qualifier: String, theory: String): String = if (Long_Name.is_qualified(theory) || session_base.global_theories.isDefinedAt(theory)) theory else Long_Name.qualify(qualifier, theory) 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 = sessions_structure.get(session) match { case Some(info) => info.dirs case None => Nil } dirs.collectFirst({ 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 = { val theory = theory_name(qualifier, Thy_Header.import_name(s)) def theory_node = file_node(Path.explode(s).thy, dir = dir, theory = theory) if (!Thy_Header.is_base_name(s)) theory_node else if (session_base.loaded_theory(theory)) loaded_theory_node(theory) else { find_theory_node(theory) match { case Some(node_name) => node_name case None => if (Long_Name.is_qualified(s)) loaded_theory_node(theory) else theory_node } } } def import_name(name: Document.Node.Name, s: String): Document.Node.Name = import_name(session_base.theory_qualifier(name), name.master_dir, s) 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] = { for { qualifier <- session_base.session_directories.get(File.canonical(file).getParentFile) theory_base <- proper_string(Thy_Header.theory_name(file.getName)) theory = theory_name(qualifier, theory_base) theory_node <- find_theory_node(theory) if File.eq(theory_node.path.file, file) } yield theory_node } 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) } catch { case ERROR(_) => None } (for { (session, (info, _)) <- sessions_structure.imports_graph.iterator dir <- (if (session == context_session) context_dir.toList else info.dirs).iterator theory <- Thy_Header.try_read_dir(dir).iterator if Completion.completed(s)(theory) } yield { if (session == context_session || session_base.global_theories.isDefinedAt(theory)) theory else Long_Name.qualify(session, theory) }).toList.sorted } 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) error("Cannot load theory " + quote(name.theory)) 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 = { if (node_name.is_theory && reader.source.length > 0) { try { val header = Thy_Header.read(node_name, reader, command = command, strict = strict) val imports = header.imports.map({ case (s, pos) => val name = import_name(node_name, s) if (Sessions.exclude_theory(name.theory_base_name)) error("Bad theory name " + quote(name.theory_base_name) + Position.here(pos)) (name, pos) }) Document.Node.Header(imports, header.keywords, header.abbrevs) } catch { case exn: Throwable => Document.Node.bad_header(Exn.message(exn)) } } else Document.Node.no_header } /* special 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)) else if (Thy_Header.is_bootstrap(name.theory)) List(import_name(name, Thy_Header.PURE)) else Nil if (imports.isEmpty) None else Some(Document.Node.Header(imports.map((_, Position.none)))) } /* blobs */ def undefined_blobs(nodes: Document.Nodes): List[Document.Node.Name] = (for { (node_name, node) <- nodes.iterator if !session_base.loaded_theory(node_name) cmd <- node.load_commands.iterator name <- cmd.blobs_undefined.iterator } yield name).toList /* document changes */ def parse_change( reparse_limit: Int, previous: Document.Version, doc_blobs: Document.Blobs, edits: List[Document.Edit_Text], consolidate: List[Document.Node.Name]): Session.Change = Thy_Syntax.parse_change(resources, reparse_limit, previous, doc_blobs, edits, consolidate) def commit(change: Session.Change): Unit = {} /* theory and file dependencies */ def dependencies( thys: List[(Document.Node.Name, Position.T)], 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] = { - (Dependencies.empty[Options] /: info.theories)({ case (dependencies, (options, thys)) => - dependencies.require_thys(options, - for { (thy, pos) <- thys } yield (import_name(info, thy), pos), - progress = progress) - }) + info.theories.foldLeft(Dependencies.empty[Options]) { + case (dependencies, (options, thys)) => + dependencies.require_thys(options, + for { (thy, pos) <- thys } yield (import_name(info, thy), pos), + progress = progress) + } } object Dependencies { def empty[A]: Dependencies[A] = new Dependencies[A](Nil, Map.empty) private def show_path(names: List[Document.Node.Name]): String = names.map(name => quote(name.theory)).mkString(" via ") private def cycle_msg(names: List[Document.Node.Name]): String = "Cyclic dependency of " + show_path(names) private def required_by(initiators: List[Document.Node.Name]): String = if (initiators.isEmpty) "" else "\n(required by " + show_path(initiators.reverse) + ")" } final class Dependencies[A] private( rev_entries: List[Document.Node.Entry], 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] = { val (name, pos) = thy def message: String = "The error(s) above occurred for theory " + quote(name.theory) + Dependencies.required_by(initiators) + Position.here(pos) if (seen.isDefinedAt(name)) this else { val dependencies1 = new Dependencies[A](rev_entries, seen + (name -> adjunct)) if (session_base.loaded_theory(name)) dependencies1 else { try { if (initiators.contains(name)) error(Dependencies.cycle_msg(initiators)) progress.expose_interrupt() val header = try { with_thy_reader(name, check_thy(name, _, command = false)).cat_errors(message) } catch { case ERROR(msg) => cat_error(msg, message) } val entry = Document.Node.Entry(name, header) dependencies1.require_thys(adjunct, header.imports_pos, initiators = name :: initiators, progress = progress).cons(entry) } catch { case e: Throwable => dependencies1.cons(Document.Node.Entry(name, Document.Node.bad_header(Exn.message(e)))) } } } } def require_thys(adjunct: A, thys: List[(Document.Node.Name, Position.T)], progress: Progress = new Progress, initiators: List[Document.Node.Name] = Nil): Dependencies[A] = - (this /: thys)(_.require_thy(adjunct, _, progress = progress, initiators = initiators)) + thys.foldLeft(this)(_.require_thy(adjunct, _, progress = progress, initiators = initiators)) def entries: List[Document.Node.Entry] = rev_entries.reverse def theories: List[Document.Node.Name] = entries.map(_.name) def theories_adjunct: List[(Document.Node.Name, A)] = theories.map(name => (name, seen(name))) def errors: List[String] = entries.flatMap(_.header.errors) def check_errors: Dependencies[A] = errors match { case Nil => this case errs => error(cat_lines(errs)) } lazy val theory_graph: Document.Node.Name.Graph[Unit] = { val regular = theories.toSet val irregular = (for { entry <- entries.iterator imp <- entry.header.imports if !regular(imp) } yield imp).toSet Document.Node.Name.make_graph( irregular.toList.map(name => ((name, ()), Nil)) ::: entries.map(entry => ((entry.name, ()), entry.header.imports))) } lazy val loaded_theories: Graph[String, Outer_Syntax] = - (session_base.loaded_theories /: entries)({ case (graph, entry) => - val name = entry.name.theory - val imports = entry.header.imports.map(_.theory) + entries.foldLeft(session_base.loaded_theories) { + case (graph, entry) => + val name = entry.name.theory + val imports = entry.header.imports.map(_.theory) - val graph1 = (graph /: (name :: imports))(_.default_node(_, Outer_Syntax.empty)) - val graph2 = (graph1 /: imports)(_.add_edge(_, name)) + val graph1 = (name :: imports).foldLeft(graph)(_.default_node(_, Outer_Syntax.empty)) + val graph2 = imports.foldLeft(graph1)(_.add_edge(_, name)) - val syntax0 = if (name == Thy_Header.PURE) List(Thy_Header.bootstrap_syntax) else Nil - val syntax1 = (name :: graph2.imm_preds(name).toList).map(graph2.get_node) - val syntax = Outer_Syntax.merge(syntax0 ::: syntax1) + entry.header + val syntax0 = if (name == Thy_Header.PURE) List(Thy_Header.bootstrap_syntax) else Nil + val syntax1 = (name :: graph2.imm_preds(name).toList).map(graph2.get_node) + val syntax = Outer_Syntax.merge(syntax0 ::: syntax1) + entry.header - graph2.map_node(name, _ => syntax) - }) + graph2.map_node(name, _ => syntax) + } def get_syntax(name: Document.Node.Name): Outer_Syntax = loaded_theories.get_node(name.theory) def load_commands: List[(Document.Node.Name, List[Command_Span.Span])] = theories.zip( Par_List.map((e: () => List[Command_Span.Span]) => e(), 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]) = { val theory = name.theory val syntax = get_syntax(name) val files1 = resources.loaded_files(syntax, name, spans) val files2 = if (theory == Thy_Header.PURE) pure_files(syntax) else Nil (theory, files1 ::: files2) } def loaded_files: List[Path] = for { (name, spans) <- load_commands file <- loaded_files(name, spans)._2 } yield file def imported_files: List[Path] = { val base_theories = loaded_theories.all_preds(theories.map(_.theory)). filter(session_base.loaded_theories.defined) base_theories.map(theory => session_base.known_theories(theory).name.path) ::: base_theories.flatMap(session_base.known_loaded_files.withDefaultValue(Nil)) } lazy val overall_syntax: Outer_Syntax = Outer_Syntax.merge(loaded_theories.maximals.map(loaded_theories.get_node)) override def toString: String = entries.toString } } diff --git a/src/Pure/PIDE/xml.scala b/src/Pure/PIDE/xml.scala --- a/src/Pure/PIDE/xml.scala +++ b/src/Pure/PIDE/xml.scala @@ -1,457 +1,457 @@ /* Title: Pure/PIDE/xml.scala Author: Makarius Untyped XML trees and basic data representation. */ package isabelle object XML { /** XML trees **/ /* datatype representation */ type Attribute = Properties.Entry type Attributes = Properties.T 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 { private lazy val hash: Int = (markup, body).hashCode() override def hashCode(): Int = hash def name: String = markup.name def update_attributes(more_attributes: Attributes): Elem = if (more_attributes.isEmpty) this else Elem(markup.update_properties(more_attributes), body) def + (att: Attribute): Elem = Elem(markup + att, body) } case class Text(content: String) extends Tree { private lazy val hash: Int = content.hashCode() override def hashCode(): Int = hash } def elem(markup: Markup): XML.Elem = XML.Elem(markup, Nil) def elem(name: String, body: Body): XML.Elem = XML.Elem(Markup(name, Nil), body) def elem(name: String): XML.Elem = XML.Elem(Markup(name, Nil), Nil) val no_text: Text = Text("") val newline: Text = Text("\n") /* name space */ object Namespace { def apply(prefix: String, target: String): Namespace = new Namespace(prefix, target) } final class Namespace private(prefix: String, target: String) { def apply(name: String): String = prefix + ":" + name val attribute: XML.Attribute = ("xmlns:" + prefix, target) override def toString: String = attribute.toString } /* wrapped elements */ val XML_ELEM = "xml_elem" val XML_NAME = "xml_name" val XML_BODY = "xml_body" 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) def unapply(tree: Tree): Option[(Markup, Body, Body)] = tree match { case XML.Elem(Markup(XML_ELEM, (XML_NAME, name) :: props), XML.Elem(Markup(XML_BODY, Nil), body1) :: body2) => Some(Markup(name, props), body1, body2) case _ => None } } object Root_Elem { def apply(body: Body): XML.Elem = XML.Elem(Markup(XML_ELEM, Nil), body) def unapply(tree: Tree): Option[Body] = tree match { case XML.Elem(Markup(XML_ELEM, Nil), body) => Some(body) case _ => None } } /* traverse text */ 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) => (x /: ts)(traverse) - case XML.Elem(_, ts) => (x /: ts)(traverse) + case XML.Wrapped_Elem(_, _, ts) => ts.foldLeft(x)(traverse) + case XML.Elem(_, ts) => ts.foldLeft(x)(traverse) case XML.Text(s) => op(x, s) } - (a /: body)(traverse) + body.foldLeft(a)(traverse) } def text_length(body: Body): Int = traverse_text(body)(0) { case (n, s) => n + s.length } /* text content */ def content(body: Body): String = { val text = new StringBuilder(text_length(body)) traverse_text(body)(()) { case (_, s) => text.append(s) } text.toString } def content(tree: Tree): String = content(List(tree)) /** string representation **/ val header: String = "\n" def output_char(s: StringBuilder, c: Char, permissive: Boolean = false): Unit = { c match { case '<' => s ++= "<" case '>' => s ++= ">" case '&' => s ++= "&" case '"' if !permissive => s ++= """ case '\'' if !permissive => s ++= "'" case _ => s += c } } 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 = { s += '<' s ++= markup.name for ((a, b) <- markup.properties) { s += ' ' s ++= a s += '=' s += '"' output_string(s, b) s += '"' } if (end) s += '/' s += '>' } def output_elem_end(s: StringBuilder, name: String): Unit = { s += '<' s += '/' s ++= name s += '>' } def string_of_body(body: Body): String = { val s = new StringBuilder def tree(t: Tree): Unit = t match { case XML.Elem(markup, Nil) => output_elem(s, markup, end = true) case XML.Elem(markup, ts) => output_elem(s, markup) ts.foreach(tree) output_elem_end(s, markup.name) case XML.Text(txt) => output_string(s, txt) } body.foreach(tree) s.toString } def string_of_tree(tree: XML.Tree): String = string_of_body(List(tree)) /** cache **/ object Cache { def make( xz: XZ.Cache = XZ.Cache.make(), max_string: Int = isabelle.Cache.default_max_string, initial_size: Int = isabelle.Cache.default_initial_size): Cache = new Cache(xz, max_string, initial_size) val none: Cache = make(XZ.Cache.none, max_string = 0) } class Cache private[XML](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 = { if (x.isEmpty) x else lookup(x) match { case Some(y) => y case None => store(x.map(p => (Library.isolate_substring(p._1).intern, cache_string(p._2)))) } } protected def cache_markup(x: Markup): Markup = { lookup(x) match { case Some(y) => y case None => x match { case Markup(name, props) => store(Markup(cache_string(name), cache_props(props))) } } } protected def cache_tree(x: XML.Tree): XML.Tree = { lookup(x) match { case Some(y) => y case None => x match { case XML.Elem(markup, body) => store(XML.Elem(cache_markup(markup), cache_body(body))) case XML.Text(text) => store(XML.Text(cache_string(text))) } } } protected def cache_body(x: XML.Body): XML.Body = { if (x.isEmpty) x else lookup(x) match { case Some(y) => y case None => x.map(cache_tree) } } // support hash-consing def tree0(x: XML.Tree): XML.Tree = if (no_cache) x else synchronized { lookup(x) getOrElse store(x) } // main methods def props(x: Properties.T): Properties.T = if (no_cache) x else synchronized { cache_props(x) } def markup(x: Markup): Markup = if (no_cache) x else synchronized { cache_markup(x) } def tree(x: XML.Tree): XML.Tree = if (no_cache) x else synchronized { cache_tree(x) } def body(x: XML.Body): XML.Body = if (no_cache) x else synchronized { cache_body(x) } def elem(x: XML.Elem): XML.Elem = if (no_cache) x else synchronized { cache_tree(x).asInstanceOf[XML.Elem] } } /** XML as data representation language **/ abstract class Error(s: String) extends Exception(s) class XML_Atom(s: String) extends Error(s) class XML_Body(body: XML.Body) extends Error("") object Encode { type T[A] = A => XML.Body type V[A] = PartialFunction[A, (List[String], XML.Body)] type P[A] = PartialFunction[A, List[String]] /* atomic values */ def long_atom(i: Long): String = Library.signed_string_of_long(i) def int_atom(i: Int): String = Library.signed_string_of_int(i) def bool_atom(b: Boolean): String = if (b) "1" else "0" def unit_atom(u: Unit) = "" /* structural nodes */ private def node(ts: XML.Body): XML.Tree = XML.Elem(Markup(":", Nil), ts) private def vector(xs: List[String]): XML.Attributes = xs.zipWithIndex.map({ case (x, i) => (int_atom(i), x) }) private def tagged(tag: Int, data: (List[String], XML.Body)): XML.Tree = XML.Elem(Markup(int_atom(tag), vector(data._1)), data._2) /* representation of standard types */ val tree: T[XML.Tree] = (t => List(t)) val properties: T[Properties.T] = (props => List(XML.Elem(Markup(":", props), Nil))) val string: T[String] = (s => if (s.isEmpty) Nil else List(XML.Text(s))) val long: T[Long] = (x => string(long_atom(x))) val int: T[Int] = (x => string(int_atom(x))) val bool: T[Boolean] = (x => string(bool_atom(x))) val unit: T[Unit] = (x => string(unit_atom(x))) def pair[A, B](f: T[A], g: T[B]): T[(A, B)] = (x => List(node(f(x._1)), node(g(x._2)))) def triple[A, B, C](f: T[A], g: T[B], h: T[C]): T[(A, B, C)] = (x => List(node(f(x._1)), node(g(x._2)), node(h(x._3)))) 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]] = { case None => Nil case Some(x) => List(node(f(x))) } 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 { type T[A] = XML.Body => A type V[A] = (List[String], XML.Body) => A type P[A] = PartialFunction[List[String], A] /* atomic values */ def long_atom(s: String): Long = try { java.lang.Long.parseLong(s) } catch { case e: NumberFormatException => throw new XML_Atom(s) } def int_atom(s: String): Int = try { Integer.parseInt(s) } catch { case e: NumberFormatException => throw new XML_Atom(s) } def bool_atom(s: String): Boolean = if (s == "1") true else if (s == "0") false else throw new XML_Atom(s) def unit_atom(s: String): Unit = if (s == "") () else throw new XML_Atom(s) /* structural nodes */ private def node(t: XML.Tree): XML.Body = t match { case XML.Elem(Markup(":", Nil), ts) => ts case _ => throw new XML_Body(List(t)) } private def vector(atts: XML.Attributes): List[String] = atts.iterator.zipWithIndex.map( { case ((a, x), i) => if (int_atom(a) == i) x else throw new XML_Atom(a) }).toList private def tagged(t: XML.Tree): (Int, (List[String], XML.Body)) = t match { case XML.Elem(Markup(name, atts), ts) => (int_atom(name), (vector(atts), ts)) case _ => throw new XML_Body(List(t)) } /* representation of standard types */ val tree: T[XML.Tree] = { case List(t) => t case ts => throw new XML_Body(ts) } val properties: T[Properties.T] = { case List(XML.Elem(Markup(":", props), Nil)) => props case ts => throw new XML_Body(ts) } val string: T[String] = { case Nil => "" case List(XML.Text(s)) => s case ts => throw new XML_Body(ts) } val long: T[Long] = (x => long_atom(string(x))) val int: T[Int] = (x => int_atom(string(x))) val bool: T[Boolean] = (x => bool_atom(string(x))) val unit: T[Unit] = (x => unit_atom(string(x))) 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)] = { case List(t1, t2, t3) => (f(node(t1)), g(node(t2)), h(node(t3))) case ts => throw new XML_Body(ts) } 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]] = { 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] = { case List(t) => val (tag, (xs, ts)) = tagged(t) val f = try { fs(tag) } catch { case _: IndexOutOfBoundsException => throw new XML_Body(List(t)) } f(xs, ts) case ts => throw new XML_Body(ts) } } } diff --git a/src/Pure/System/getopts.scala b/src/Pure/System/getopts.scala --- a/src/Pure/System/getopts.scala +++ b/src/Pure/System/getopts.scala @@ -1,72 +1,72 @@ /* Title: Pure/System/getopts.scala Author: Makarius Support for command-line options as in GNU bash. */ package isabelle import scala.annotation.tailrec object Getopts { def apply(usage_text: String, option_specs: (String, String => Unit)*): Getopts = { val options = - (Map.empty[Char, (Boolean, String => Unit)] /: option_specs) { + option_specs.foldLeft(Map.empty[Char, (Boolean, String => Unit)]) { case (m, (s, f)) => val (a, entry) = if (s.size == 1) (s(0), (false, f)) else if (s.size == 2 && s.endsWith(":")) (s(0), (true, f)) else error("Bad option specification: " + quote(s)) if (m.isDefinedAt(a)) error("Duplicate option specification: " + quote(a.toString)) else m + (a -> entry) } new Getopts(usage_text, options) } } class Getopts private(usage_text: String, options: Map[Char, (Boolean, String => Unit)]) { def usage(): Nothing = { Output.writeln(usage_text, stdout = true) sys.exit(1) } private def is_option(opt: Char): Boolean = options.isDefinedAt(opt) private def is_option0(opt: Char): Boolean = is_option(opt) && !options(opt)._1 private def is_option1(opt: Char): Boolean = is_option(opt) && options(opt)._1 private def print_option(opt: Char): String = quote("-" + opt.toString) private def option(opt: Char, opt_arg: List[Char]): Unit = try { options(opt)._2.apply(opt_arg.mkString) } catch { case ERROR(msg) => cat_error(msg, "The error(s) above occurred in command-line option " + print_option(opt)) } @tailrec private def getopts(args: List[List[Char]]): List[List[Char]] = args match { case List('-', '-') :: rest_args => rest_args case ('-' :: opt :: rest_opts) :: rest_args if is_option0(opt) && rest_opts.nonEmpty => option(opt, Nil); getopts(('-' :: rest_opts) :: rest_args) case List('-', opt) :: rest_args if is_option0(opt) => option(opt, Nil); getopts(rest_args) case ('-' :: opt :: opt_arg) :: rest_args if is_option1(opt) && opt_arg.nonEmpty => option(opt, opt_arg); getopts(rest_args) case List('-', opt) :: opt_arg :: rest_args if is_option1(opt) => option(opt, opt_arg); getopts(rest_args) case List(List('-', opt)) if is_option1(opt) => Output.error_message("Command-line option " + print_option(opt) + " requires an argument") usage() case ('-' :: opt :: _) :: _ if !is_option(opt) => if (opt != '?') Output.error_message("Illegal command-line option " + print_option(opt)) usage() case _ => args } def apply(args: List[String]): List[String] = getopts(args.map(_.toList)).map(_.mkString) def apply(args: Array[String]): List[String] = apply(args.toList) } diff --git a/src/Pure/System/options.scala b/src/Pure/System/options.scala --- a/src/Pure/System/options.scala +++ b/src/Pure/System/options.scala @@ -1,454 +1,454 @@ /* Title: Pure/System/options.scala Author: Makarius System options with external string representation. */ package isabelle object Options { type Spec = (String, Option[String]) val empty: Options = new Options() /* representation */ sealed abstract class Type { def print: String = Word.lowercase(toString) } case object Bool extends Type case object Int extends Type case object Real extends Type case object String extends Type case object Unknown extends Type case class Opt( public: Boolean, pos: Position.T, name: String, typ: Type, value: String, default_value: String, description: String, section: String) { private def print(default: Boolean): String = { val x = if (default) default_value else value "option " + name + " : " + typ.print + " = " + (if (typ == Options.String) quote(x) else x) + (if (description == "") "" else "\n -- " + quote(description)) } def print: String = print(false) def print_default: String = print(true) def title(strip: String = ""): String = { val words = Word.explode('_', name) val words1 = words match { case word :: rest if word == strip => rest case _ => words } Word.implode(words1.map(Word.perhaps_capitalize)) } def unknown: Boolean = typ == Unknown } /* parsing */ private val SECTION = "section" private val PUBLIC = "public" private val OPTION = "option" private val OPTIONS = Path.explode("etc/options") private val PREFS = Path.explode("$ISABELLE_HOME_USER/etc/preferences") val options_syntax: Outer_Syntax = Outer_Syntax.empty + ":" + "=" + "--" + Symbol.comment + Symbol.comment_decoded + (SECTION, Keyword.DOCUMENT_HEADING) + (PUBLIC, Keyword.BEFORE_COMMAND) + (OPTION, Keyword.THY_DECL) val prefs_syntax: Outer_Syntax = Outer_Syntax.empty + "=" 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] = opt(token("-", tok => tok.is_sym_ident && tok.content == "-")) ~ atom("nat", _.is_nat) ^^ { case s ~ n => if (s.isDefined) "-" + n else n } | atom("option value", tok => tok.is_name || tok.is_float) } private object Parser extends Parser { def comment_marker: Parser[String] = $$$("--") | $$$(Symbol.comment) | $$$(Symbol.comment_decoded) 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 ~ $$$("=") ~ option_value ~ (comment_marker ~! text ^^ { case _ ~ x => x } | success(""))) ^^ { case a ~ _ ~ ((b, pos) ~ _ ~ c ~ _ ~ d ~ e) => (options: Options) => options.declare(a.isDefined, pos, b, c, d, e) } } 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, syntax: Outer_Syntax = options_syntax, 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 { case Success(result, _) => result case bad => error(bad.toString) } - try { (options.set_section("") /: ops) { case (opts, op) => op(opts) } } + try { ops.foldLeft(options.set_section("")) { case (opts, op) => op(opts) } } catch { case ERROR(msg) => error(msg + Position.here(Position.File(file_name))) } } def parse_prefs(options: Options, content: String): Options = parse_file(options, PREFS.file_name, content, syntax = prefs_syntax, parser = prefs_entry) } 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 = { var options = empty for { dir <- Isabelle_System.components() file = dir + OPTIONS if file.is_file } { options = Parser.parse_file(options, file.implode, File.read(file)) } - (Options.Parser.parse_prefs(options, prefs) /: opts)(_ + _) + opts.foldLeft(Options.Parser.parse_prefs(options, prefs))(_ + _) } /* encode */ val encode: XML.Encode.T[Options] = (options => options.encode) /* Isabelle tool wrapper */ val isabelle_tool = Isabelle_Tool("options", "print Isabelle system options", Scala_Project.here, args => { var build_options = false var get_option = "" var list_options = false var export_file = "" val getopts = Getopts(""" Usage: isabelle options [OPTIONS] [MORE_OPTIONS ...] Options are: -b include $ISABELLE_BUILD_OPTIONS -g OPTION get value of OPTION -l list options -x FILE export options to FILE in YXML format Report Isabelle system options, augmented by MORE_OPTIONS given as arguments NAME=VAL or NAME. """, "b" -> (_ => build_options = true), "g:" -> (arg => get_option = arg), "l" -> (_ => list_options = true), "x:" -> (arg => export_file = arg)) val more_options = getopts(args) if (get_option == "" && !list_options && export_file == "") getopts.usage() val options = { val options0 = Options.init() val options1 = if (build_options) - (options0 /: Word.explode(Isabelle_System.getenv("ISABELLE_BUILD_OPTIONS")))(_ + _) + Word.explode(Isabelle_System.getenv("ISABELLE_BUILD_OPTIONS")).foldLeft(options0)(_ + _) else options0 - (options1 /: more_options)(_ + _) + more_options.foldLeft(options1)(_ + _) } if (get_option != "") Output.writeln(options.check_name(get_option).value, stdout = true) if (export_file != "") File.write(Path.explode(export_file), YXML.string_of_body(options.encode)) if (get_option == "" && export_file == "") Output.writeln(options.print, stdout = true) }) } final class Options private( val options: Map[String, Options.Opt] = Map.empty, val section: String = "") { override def toString: String = options.iterator.mkString("Options(", ",", ")") private def print_opt(opt: Options.Opt): String = if (opt.public) "public " + opt.print else opt.print def print: String = cat_lines(options.toList.sortBy(_._1).map(p => print_opt(p._2))) def description(name: String): String = check_name(name).description /* check */ def check_name(name: String): Options.Opt = options.get(name) match { case Some(opt) if !opt.unknown => opt case _ => error("Unknown option " + quote(name)) } private def check_type(name: String, typ: Options.Type): Options.Opt = { val opt = check_name(name) if (opt.typ == typ) opt else error("Ill-typed option " + quote(name) + " : " + opt.typ.print + " vs. " + typ.print) } /* basic operations */ private def put[A](name: String, typ: Options.Type, value: String): Options = { val opt = check_type(name, typ) new Options(options + (name -> opt.copy(value = value)), section) } private def get[A](name: String, typ: Options.Type, parse: String => Option[A]): A = { val opt = check_type(name, typ) parse(opt.value) match { case Some(x) => x case None => error("Malformed value for option " + quote(name) + " : " + typ.print + " =\n" + quote(opt.value)) } } /* internal lookup and update */ class Bool_Access { def apply(name: String): Boolean = get(name, Options.Bool, 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 { 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 { 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 { def apply(name: String): String = get(name, Options.String, s => Some(s)) def update(name: String, x: String): Options = put(name, Options.String, x) } val string = new String_Access def proper_string(name: String): Option[String] = Library.proper_string(string(name)) def seconds(name: String): Time = Time.seconds(real(name)) /* external updates */ private def check_value(name: String): Options = { val opt = check_name(name) opt.typ match { case Options.Bool => bool(name); this case Options.Int => int(name); this case Options.Real => real(name); this case Options.String => string(name); this case Options.Unknown => this } } def declare( public: Boolean, pos: Position.T, name: String, typ_name: String, value: String, description: String): Options = { options.get(name) match { case Some(other) => error("Duplicate declaration of option " + quote(name) + Position.here(pos) + Position.here(other.pos)) case None => val typ = typ_name match { case "bool" => Options.Bool case "int" => Options.Int case "real" => Options.Real case "string" => Options.String case _ => error("Unknown type for option " + quote(name) + " : " + quote(typ_name) + Position.here(pos)) } val opt = Options.Opt(public, pos, name, typ, value, value, description, section) (new Options(options + (name -> opt), section)).check_value(name) } } def add_permissive(name: String, value: String): Options = { if (options.isDefinedAt(name)) this + (name, value) else { val opt = Options.Opt(false, Position.none, name, Options.Unknown, value, value, "", "") new Options(options + (name -> opt), section) } } def + (name: String, value: String): Options = { val opt = check_name(name) (new Options(options + (name -> opt.copy(value = value)), section)).check_value(name) } def + (name: String, opt_value: Option[String]): Options = { val opt = check_name(name) opt_value match { case Some(value) => this + (name, value) case None if opt.typ == Options.Bool => this + (name, "true") case None => error("Missing value for option " + quote(name) + " : " + opt.typ.print) } } def + (str: String): Options = { str.indexOf('=') match { case -1 => this + (str, None) case i => this + (str.substring(0, i), str.substring(i + 1)) } } def ++ (specs: List[Options.Spec]): Options = - (this /: specs)({ case (x, (y, z)) => x + (y, z) }) + specs.foldLeft(this) { case (x, (y, z)) => x + (y, z) } /* sections */ def set_section(new_section: String): Options = new Options(options, new_section) def sections: List[(String, List[Options.Opt])] = options.groupBy(_._2.section).toList.map({ case (a, opts) => (a, opts.toList.map(_._2)) }) /* encode */ def encode: XML.Body = { val opts = for ((_, opt) <- options.toList; if !opt.unknown) yield (opt.pos, (opt.name, (opt.typ.print, opt.value))) import XML.Encode.{string => string_, _} list(pair(properties, pair(string_, pair(string_, string_))))(opts) } /* save preferences */ def save_prefs(file: Path = Options.PREFS): Unit = { val defaults: Options = Options.init(prefs = "") val changed = (for { (name, opt2) <- options.iterator opt1 = defaults.options.get(name) if opt1.isEmpty || opt1.get.value != opt2.value } yield (name, opt2.value, if (opt1.isEmpty) " (* unknown *)" else "")).toList val prefs = changed.sortBy(_._1) .map({ case (x, y, z) => x + " = " + Outer_Syntax.quote_string(y) + z + "\n" }).mkString Isabelle_System.make_directory(file.dir) File.write_backup(file, "(* generated by Isabelle " + Date.now() + " *)\n\n" + prefs) } } class Options_Variable(init_options: Options) { private var options = init_options def value: Options = synchronized { options } 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 { 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 { 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 { 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 { def apply(name: String): String = value.string(name) def update(name: String, x: String): Unit = upd(opts => opts.string.update(name, x)) } val string = new String_Access def proper_string(name: String): Option[String] = Library.proper_string(string(name)) def seconds(name: String): Time = value.seconds(name) } diff --git a/src/Pure/Thy/bibtex.scala b/src/Pure/Thy/bibtex.scala --- a/src/Pure/Thy/bibtex.scala +++ b/src/Pure/Thy/bibtex.scala @@ -1,702 +1,704 @@ /* Title: Pure/Thy/bibtex.scala Author: Makarius BibTeX support. */ package isabelle import java.io.{File => JFile} import scala.collection.mutable import scala.util.parsing.combinator.RegexParsers import scala.util.parsing.input.Reader object Bibtex { /** file format **/ def is_bibtex(name: String): Boolean = name.endsWith(".bib") class File_Format extends isabelle.File_Format { val format_name: String = "bibtex" val file_ext: String = "bib" override def theory_suffix: String = "bibtex_file" override def theory_content(name: String): String = """theory "bib" imports Pure begin bibtex_file """ + Outer_Syntax.quote_string(name) + """ end""" 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) val content = Isabelle_System.with_tmp_file("bib", "bib") { bib => File.write(bib, snapshot.node.source) Bibtex.html_output(List(bib), style = "unsort", title = title) } Some(Presentation.HTML_Document(title, content)) } else None } } /** bibtex errors **/ 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 val Error2 = """^(.+)---line (\d+) of file (.+)""".r Line.logical_lines(File.read(log_path)).flatMap(line => line match { case Error1(msg) => Some("Bibtex error: " + msg) case Error2(msg, Value.Int(l), file) => val path = File.standard_path(file) if (Path.is_wellformed(path)) { val pos = Position.Line_File(l, (dir + Path.explode(path)).canonical.implode) Some("Bibtex error" + Position.here(pos) + ":\n " + msg) } else None case _ => None }) } else Nil } /** check database **/ 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)] var line = 1 var offset = 1 def make_pos(length: Int): Position.T = Position.Offset(offset) ::: Position.End_Offset(offset + length) ::: Position.Line(line) def advance_pos(tok: Token): Unit = { for (s <- Symbol.iterator(tok.source)) { if (Symbol.is_newline(s)) line += 1 offset += 1 } } def get_line_pos(l: Int): Position.T = if (0 < l && l <= tokens.length) tokens(l - 1)._2 else Position.none for (chunk <- chunks) { val name = chunk.name if (name != "" && !chunk_pos.isDefinedAt(name)) { chunk_pos += (name -> make_pos(chunk.heading_length)) } for (tok <- chunk.tokens) { tokens += (tok.copy(source = tok.source.replace("\n", " ")) -> make_pos(tok.source.length)) advance_pos(tok) } } 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"), "\\bibstyle{plain}\n\\bibdata{root}\n\\citation{*}") Isabelle_System.bash("\"$ISABELLE_BIBTEX\" root", cwd = tmp_dir.file) val Error = """^(.*)---line (\d+) of file root.bib$""".r val Warning = """^Warning--(.+)$""".r val Warning_Line = """--line (\d+) of file root.bib$""".r val Warning_in_Chunk = """^Warning--(.+) in (.+)$""".r val log_file = tmp_dir + Path.explode("root.blg") val lines = if (log_file.is_file) Line.logical_lines(File.read(log_file)) else Nil val (errors, warnings) = if (lines.isEmpty) (Nil, Nil) else { lines.zip(lines.tail ::: List("")).flatMap( { case (Error(msg, Value.Int(l)), _) => Some((true, (msg, get_line_pos(l)))) case (Warning_in_Chunk(msg, name), _) if chunk_pos.isDefinedAt(name) => Some((false, (Word.capitalize(msg + " in entry " + quote(name)), chunk_pos(name)))) case (Warning(msg), Warning_Line(Value.Int(l))) => Some((false, (Word.capitalize(msg), get_line_pos(l)))) case (Warning(msg), _) => Some((false, (Word.capitalize(msg), Position.none))) case _ => None } ).partition(_._1) } (errors.map(_._2), warnings.map(_._2)) }) } object Check_Database extends Scala.Fun("bibtex_check_database") { val here = Scala_Project.here def apply(database: String): String = { import XML.Encode._ YXML.string_of_body(pair(list(pair(string, properties)), list(pair(string, properties)))( check_database(database))) } } /** document model **/ /* entries */ def entries(text: String): List[Text.Info[String]] = { val result = new mutable.ListBuffer[Text.Info[String]] var offset = 0 for (chunk <- Bibtex.parse(text)) { val end_offset = offset + chunk.source.length if (chunk.name != "" && !chunk.is_command) result += Text.Info(Text.Range(offset, end_offset), chunk.name) offset = end_offset } result.toList } 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 } yield info.map((_, model)) } /* completion */ def completion[A, B <: Document.Model]( 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) original <- rendering.get_text(r) original1 <- Completion.clean_name(Library.perhaps_unquote(original)) entries = (for { Text.Info(_, (entry, _)) <- entries_iterator(models) if entry.toLowerCase.containsSlice(name1.toLowerCase) && entry != original1 } yield entry).toList if entries.nonEmpty items = entries.sorted.map({ case entry => val full_name = Long_Name.qualify(Markup.CITATION, entry) val description = List(entry, "(BibTeX entry)") val replacement = quote(entry) Completion.Item(r, original, full_name, description, replacement, 0, false) }).sorted(history.ordering).take(rendering.options.int("completion_limit")) } yield Completion.Result(r, original, false, items) } /** content **/ private val months = List( "jan", "feb", "mar", "apr", "may", "jun", "jul", "aug", "sep", "oct", "nov", "dec") def is_month(s: String): Boolean = months.contains(s.toLowerCase) private val commands = List("preamble", "string") def is_command(s: String): Boolean = commands.contains(s.toLowerCase) sealed case class Entry( kind: String, required: List[String], optional_crossref: 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) def is_optional(s: String): Boolean = optional_crossref.contains(s.toLowerCase) || optional_other.contains(s.toLowerCase) || optional_standard.contains(s.toLowerCase) def fields: List[String] = required ::: optional_crossref ::: optional_other ::: optional_standard def template: String = "@" + kind + "{,\n" + fields.map(x => " " + x + " = {},\n").mkString + "}\n" } val known_entries: List[Entry] = List( Entry("Article", List("author", "title"), List("journal", "year"), List("volume", "number", "pages", "month", "note")), Entry("InProceedings", List("author", "title"), List("booktitle", "year"), List("editor", "volume", "number", "series", "pages", "month", "address", "organization", "publisher", "note")), Entry("InCollection", List("author", "title", "booktitle"), List("publisher", "year"), List("editor", "volume", "number", "series", "type", "chapter", "pages", "edition", "month", "address", "note")), Entry("InBook", List("author", "editor", "title", "chapter"), List("publisher", "year"), List("volume", "number", "series", "type", "address", "edition", "month", "pages", "note")), Entry("Proceedings", List("title", "year"), List(), List("booktitle", "editor", "volume", "number", "series", "address", "month", "organization", "publisher", "note")), Entry("Book", List("author", "editor", "title"), List("publisher", "year"), List("volume", "number", "series", "address", "edition", "month", "note")), Entry("Booklet", List("title"), List(), List("author", "howpublished", "address", "month", "year", "note")), Entry("PhdThesis", List("author", "title", "school", "year"), List(), List("type", "address", "month", "note")), Entry("MastersThesis", List("author", "title", "school", "year"), List(), List("type", "address", "month", "note")), Entry("TechReport", List("author", "title", "institution", "year"), List(), List("type", "number", "address", "month", "note")), Entry("Manual", List("title"), List(), List("author", "organization", "address", "edition", "month", "year", "note")), Entry("Unpublished", List("author", "title", "note"), List(), List("month", "year")), Entry("Misc", List(), List(), List("author", "title", "howpublished", "month", "year", "note"))) def get_entry(kind: String): Option[Entry] = known_entries.find(entry => entry.kind.toLowerCase == kind.toLowerCase) def is_entry(kind: String): Boolean = get_entry(kind).isDefined /** tokens and chunks **/ object Token { object Kind extends Enumeration { val COMMAND = Value("command") val ENTRY = Value("entry") val KEYWORD = Value("keyword") val NAT = Value("natural number") val STRING = Value("string") val NAME = Value("name") val IDENT = Value("identifier") val SPACE = Value("white space") val COMMENT = Value("ignored text") val ERROR = Value("bad input") } } sealed case class Token(kind: Token.Kind.Value, source: String) { def is_kind: Boolean = kind == Token.Kind.COMMAND || kind == Token.Kind.ENTRY || kind == Token.Kind.IDENT def is_name: Boolean = kind == Token.Kind.NAME || kind == Token.Kind.IDENT def is_ignored: Boolean = kind == Token.Kind.SPACE || kind == Token.Kind.COMMENT def is_malformed: Boolean = kind == Token.Kind.ERROR def is_open: Boolean = kind == Token.Kind.KEYWORD && (source == "{" || source == "(") } case class Chunk(kind: String, tokens: List[Token]) { val source = tokens.map(_.source).mkString private val content: Option[List[Token]] = tokens match { case Token(Token.Kind.KEYWORD, "@") :: body if body.nonEmpty => (body.init.filterNot(_.is_ignored), body.last) match { case (tok :: Token(Token.Kind.KEYWORD, "{") :: toks, Token(Token.Kind.KEYWORD, "}")) if tok.is_kind => Some(toks) case (tok :: Token(Token.Kind.KEYWORD, "(") :: toks, Token(Token.Kind.KEYWORD, ")")) if tok.is_kind => Some(toks) case _ => None } case _ => None } def name: String = content match { case Some(tok :: _) if tok.is_name => tok.source case _ => "" } def heading_length: Int = if (name == "") 1 - else (0 /: tokens.takeWhile(tok => !tok.is_open)){ case (n, tok) => n + tok.source.length } + else { + tokens.takeWhile(tok => !tok.is_open).foldLeft(0) { case (n, tok) => n + tok.source.length } + } def is_ignored: Boolean = kind == "" && tokens.forall(_.is_ignored) def is_malformed: Boolean = kind == "" || tokens.exists(_.is_malformed) def is_command: Boolean = Bibtex.is_command(kind) && name != "" && content.isDefined def is_entry: Boolean = Bibtex.is_entry(kind) && name != "" && content.isDefined } /** parsing **/ // context of partial line-oriented scans abstract class Line_Context case object Ignored extends Line_Context case object At extends Line_Context case class Item_Start(kind: String) extends Line_Context case class Item_Open(kind: String, end: String) extends Line_Context case class Item(kind: String, end: String, delim: Delimited) extends Line_Context case class Delimited(quoted: Boolean, depth: Int) val Closed = Delimited(false, 0) private def token(kind: Token.Kind.Value)(source: String): Token = Token(kind, source) private def keyword(source: String): Token = Token(Token.Kind.KEYWORD, source) // See also https://ctan.org/tex-archive/biblio/bibtex/base/bibtex.web // module @. object Parsers extends RegexParsers { /* white space and comments */ override val whiteSpace = "".r private val space = """[ \t\n\r]+""".r ^^ token(Token.Kind.SPACE) private val spaces = rep(space) /* ignored text */ private val ignored: Parser[Chunk] = rep1("""(?i)([^@]+|@[ \t]*comment)""".r) ^^ { case ss => Chunk("", List(Token(Token.Kind.COMMENT, ss.mkString))) } private def ignored_line: Parser[(Chunk, Line_Context)] = ignored ^^ { case a => (a, Ignored) } /* delimited string: outermost "..." or {...} and body with balanced {...} */ // 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)] { require(if (delim.quoted) delim.depth > 0 else delim.depth >= 0, "bad delimiter depth") def apply(in: Input) = { val start = in.offset val end = in.source.length var i = start var q = delim.quoted var d = delim.depth var finished = false while (!finished && i < end) { val c = in.source.charAt(i) if (c == '"' && d == 0) { i += 1; d = 1; q = true } else if (c == '"' && d == 1 && q) { i += 1; d = 0; q = false; finished = true } else if (c == '{') { i += 1; d += 1 } else if (c == '}') { if (d == 1 && !q || d > 1) { i += 1; d -= 1; if (d == 0) finished = true } else {i = start; finished = true } } else if (d > 0) i += 1 else finished = true } if (i == start) Failure("bad input", in) else { val s = in.source.subSequence(start, i).toString Success((s, Delimited(q, d)), in.drop(i - start)) } } }.named("delimited_depth") private def delimited: Parser[Token] = delimited_depth(Closed) ^? { case (s, delim) if delim == Closed => Token(Token.Kind.STRING, s) } private def recover_delimited: Parser[Token] = """["{][^@]*""".r ^^ token(Token.Kind.ERROR) def delimited_line(ctxt: Item): Parser[(Chunk, Line_Context)] = delimited_depth(ctxt.delim) ^^ { case (s, delim1) => (Chunk(ctxt.kind, List(Token(Token.Kind.STRING, s))), ctxt.copy(delim = delim1)) } | recover_delimited ^^ { case a => (Chunk(ctxt.kind, List(a)), Ignored) } /* other tokens */ private val at = "@" ^^ keyword private val nat = "[0-9]+".r ^^ token(Token.Kind.NAT) private val name = """[\x21-\x7f&&[^"#%'(),={}]]+""".r ^^ token(Token.Kind.NAME) private val identifier = """[\x21-\x7f&&[^"#%'(),={}0-9]][\x21-\x7f&&[^"#%'(),={}]]*""".r private val ident = identifier ^^ token(Token.Kind.IDENT) val other_token = "[=#,]".r ^^ keyword | (nat | (ident | space)) /* body */ private val body = delimited | (recover_delimited | other_token) private def body_line(ctxt: Item) = if (ctxt.delim.depth > 0) delimited_line(ctxt) else delimited_line(ctxt) | other_token ^^ { case a => (Chunk(ctxt.kind, List(a)), ctxt) } | ctxt.end ^^ { case a => (Chunk(ctxt.kind, List(keyword(a))), Ignored) } /* items: command or entry */ private val item_kind = identifier ^^ { case a => val kind = if (is_command(a)) Token.Kind.COMMAND else if (is_entry(a)) Token.Kind.ENTRY else Token.Kind.IDENT Token(kind, a) } private val item_begin = "{" ^^ { case a => ("}", keyword(a)) } | "(" ^^ { case a => (")", keyword(a)) } private def item_name(kind: String) = kind.toLowerCase match { case "preamble" => failure("") case "string" => identifier ^^ token(Token.Kind.NAME) case _ => name } private val item_start = at ~ spaces ~ item_kind ~ spaces ^^ { case a ~ b ~ c ~ d => (c.source, List(a) ::: b ::: List(c) ::: d) } private val item: Parser[Chunk] = (item_start ~ item_begin ~ spaces) into { case (kind, a) ~ ((end, b)) ~ c => opt(item_name(kind)) ~ rep(body) ~ opt(end ^^ keyword) ^^ { case d ~ e ~ f => Chunk(kind, a ::: List(b) ::: c ::: d.toList ::: e ::: f.toList) } } private val recover_item: Parser[Chunk] = at ~ "[^@]*".r ^^ { case a ~ b => Chunk("", List(a, Token(Token.Kind.ERROR, b))) } /* chunks */ val chunk: Parser[Chunk] = ignored | (item | recover_item) def chunk_line(ctxt: Line_Context): Parser[(Chunk, Line_Context)] = { ctxt match { case Ignored => ignored_line | at ^^ { case a => (Chunk("", List(a)), At) } case At => space ^^ { case a => (Chunk("", List(a)), ctxt) } | item_kind ^^ { case a => (Chunk(a.source, List(a)), Item_Start(a.source)) } | recover_item ^^ { case a => (a, Ignored) } | ignored_line case Item_Start(kind) => space ^^ { case a => (Chunk(kind, List(a)), ctxt) } | item_begin ^^ { case (end, a) => (Chunk(kind, List(a)), Item_Open(kind, end)) } | recover_item ^^ { case a => (a, Ignored) } | ignored_line case Item_Open(kind, end) => space ^^ { case a => (Chunk(kind, List(a)), ctxt) } | item_name(kind) ^^ { case a => (Chunk(kind, List(a)), Item(kind, end, Closed)) } | body_line(Item(kind, end, Closed)) | ignored_line case item_ctxt: Item => body_line(item_ctxt) | ignored_line case _ => failure("") } } } /* parse */ def parse(input: CharSequence): List[Chunk] = Parsers.parseAll(Parsers.rep(Parsers.chunk), Scan.char_reader(input)) match { case Parsers.Success(result, _) => result case _ => error("Unexpected failure to parse input:\n" + input.toString) } 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 while (!in.atEnd) { Parsers.parse(Parsers.chunk_line(ctxt), in) match { case Parsers.Success((x, c), rest) => chunks += x; ctxt = c; in = rest case Parsers.NoSuccess(_, rest) => error("Unepected failure to parse input:\n" + rest.source.toString) } } (chunks.toList, ctxt) } /** HTML output **/ private val output_styles = List( "" -> "html-n", "plain" -> "html-n", "alpha" -> "html-a", "named" -> "html-n", "paragraph" -> "html-n", "unsort" -> "html-u", "unsortlist" -> "html-u") def html_output(bib: List[Path], title: String = "Bibliography", body: Boolean = false, citations: List[String] = List("*"), style: String = "", 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 names0 = bib_files.map(_.file_name) if (Library.duplicates(names0).isEmpty) names0 else names0.zipWithIndex.map({ case (name, i) => (i + 1).toString + "-" + name }) } for ((a, b) <- bib_files zip bib_names) { Isabelle_System.copy_file(a.ext("bib"), tmp_dir + Path.basic(b).ext("bib")) } /* style file */ val bst = output_styles.toMap.get(style) match { case Some(base) => base + (if (chronological) "c" else "") + ".bst" case None => error("Bad style for bibtex HTML output: " + quote(style) + "\n(expected: " + commas_quote(output_styles.map(_._1)) + ")") } Isabelle_System.copy_file(Path.explode("$BIB2XHTML_HOME/bst") + Path.explode(bst), tmp_dir) /* result */ val in_file = Path.explode("bib.aux") val out_file = Path.explode("bib.html") File.write(tmp_dir + in_file, bib_names.mkString("\\bibdata{", ",", "}\n") + citations.map(cite => "\\citation{" + cite + "}\n").mkString) Isabelle_System.bash( "\"$BIB2XHTML_HOME/main/bib2xhtml.pl\" -B \"$ISABELLE_BIBTEX\"" + " -u -s " + Bash.string(proper_string(style) getOrElse "empty") + (if (chronological) " -c" else "") + (if (title != "") " -h " + Bash.string(title) + " " else "") + " " + File.bash_path(in_file) + " " + File.bash_path(out_file), cwd = tmp_dir.file).check val html = File.read(tmp_dir + out_file) if (body) { cat_lines( split_lines(html). dropWhile(line => !line.startsWith("

", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\", "\\