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,222 @@ /* 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() { 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) }) 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) }) }) 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 + 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) } } (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)))) } } } } } 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 + 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_fonts.scala b/src/Pure/Admin/build_fonts.scala --- a/src/Pure/Admin/build_fonts.scala +++ b/src/Pure/Admin/build_fonts.scala @@ -1,364 +1,364 @@ /* Title: Pure/Admin/build_fonts.scala Author: Makarius Build standard Isabelle fonts: DejaVu base + Isabelle symbols. */ package isabelle object Build_Fonts { /* relevant codepoint ranges */ object Range { def base_font: Seq[Int] = (0x0020 to 0x007e) ++ // ASCII (0x00a0 to 0x024f) ++ // Latin Extended-A/B (0x0400 to 0x04ff) ++ // Cyrillic (0x0600 to 0x06ff) ++ // Arabic Seq( 0x02dd, // hungarumlaut 0x2018, // single quote 0x2019, // single quote 0x201a, // single quote 0x201c, // double quote 0x201d, // double quote 0x201e, // double quote 0x2039, // single guillemet 0x203a, // single guillemet 0x204b, // FOOTNOTE 0x20ac, // Euro 0x2710, // pencil 0xfb01, // ligature fi 0xfb02, // ligature fl 0xfffd, // replacement character ) def isabelle_font: Seq[Int] = Seq( 0x05, // X 0x06, // Y 0x07, // EOF 0x7f, // DEL 0xaf, // INVERSE 0xac, // logicalnot 0xb0, // degree 0xb1, // plusminus 0xb7, // periodcentered 0xd7, // multiply 0xf7, // divide ) ++ (0x0370 to 0x03ff) ++ // Greek (pseudo math) (0x0590 to 0x05ff) ++ // Hebrew (non-mono) Seq( 0x2010, // hyphen 0x2013, // dash 0x2014, // dash 0x2015, // dash 0x2020, // dagger 0x2021, // double dagger 0x2022, // bullet 0x2026, // ellipsis 0x2030, // perthousand 0x2032, // minute 0x2033, // second 0x2038, // caret 0x20cd, // currency symbol ) ++ (0x2100 to 0x214f) ++ // Letterlike Symbols (0x2190 to 0x21ff) ++ // Arrows (0x2200 to 0x22ff) ++ // Mathematical Operators (0x2300 to 0x23ff) ++ // Miscellaneous Technical Seq( 0x2423, // graphic for space 0x2500, // box drawing 0x2501, // box drawing 0x2508, // box drawing 0x2509, // box drawing 0x2550, // box drawing ) ++ (0x25a0 to 0x25ff) ++ // Geometric Shapes Seq( 0x261b, // ACTION 0x2660, // spade suit 0x2661, // heart suit 0x2662, // diamond suit 0x2663, // club suit 0x266d, // musical flat 0x266e, // musical natural 0x266f, // musical sharp 0x2756, // UNDEFINED 0x2759, // BOLD 0x27a7, // DESCR 0x27e6, // left white square bracket 0x27e7, // right white square bracket 0x27e8, // left angle bracket 0x27e9, // right angle bracket 0x27ea, // left double angle bracket 0x27eb, // right double angle bracket ) ++ (0x27f0 to 0x27ff) ++ // Supplemental Arrows-A (0x2900 to 0x297f) ++ // Supplemental Arrows-B (0x2980 to 0x29ff) ++ // Miscellaneous Mathematical Symbols-B (0x2a00 to 0x2aff) ++ // Supplemental Mathematical Operators Seq(0x2b1a) ++ // VERBATIM (0x1d400 to 0x1d7ff) ++ // Mathematical Alphanumeric Symbols Seq( 0x1f310, // globe with meridians (Symbola font) 0x1f4d3, // notebook (Symbola font) 0x1f5c0, // folder (Symbola font) 0x1f5cf, // page (Symbola font) ) def isabelle_math_font: Seq[Int] = (0x21 to 0x2f) ++ // bang .. slash (0x3a to 0x40) ++ // colon .. atsign (0x5b to 0x5f) ++ // leftbracket .. underscore (0x7b to 0x7e) ++ // leftbrace .. tilde Seq( 0xa9, // copyright 0xae, // registered ) val vacuous_font: Seq[Int] = Seq(0x3c) // "<" as template } /* font families */ sealed case class Family( plain: String = "", bold: String = "", italic: String = "", bold_italic: String = "") { val fonts: List[String] = proper_string(plain).toList ::: proper_string(bold).toList ::: proper_string(italic).toList ::: proper_string(bold_italic).toList def get(index: Int): String = fonts(index % fonts.length) } object Family { def isabelle_symbols: Family = Family( plain = "IsabelleSymbols.sfd", bold = "IsabelleSymbolsBold.sfd") def deja_vu_sans_mono: Family = Family( plain = "DejaVuSansMono.ttf", bold = "DejaVuSansMono-Bold.ttf", italic = "DejaVuSansMono-Oblique.ttf", bold_italic = "DejaVuSansMono-BoldOblique.ttf") def deja_vu_sans: Family = Family( plain = "DejaVuSans.ttf", bold = "DejaVuSans-Bold.ttf", italic = "DejaVuSans-Oblique.ttf", bold_italic = "DejaVuSans-BoldOblique.ttf") def deja_vu_serif: Family = Family( plain = "DejaVuSerif.ttf", bold = "DejaVuSerif-Bold.ttf", italic = "DejaVuSerif-Italic.ttf", bold_italic = "DejaVuSerif-BoldItalic.ttf") def vacuous: Family = Family(plain = "Vacuous.sfd") } /* hinting */ // see https://www.freetype.org/ttfautohint/doc/ttfautohint.html private def auto_hint(source: Path, target: Path) { Isabelle_System.bash("ttfautohint -i " + File.bash_path(source) + " " + File.bash_path(target)).check } private def hinted_path(hinted: Boolean): Path = if (hinted) Path.basic("ttf-hinted") else Path.basic("ttf") private val hinting = List(false, true) /* build fonts */ private def find_file(dirs: List[Path], name: String): Path = { val path = Path.explode(name) dirs.collectFirst({ case dir if (dir + path).is_file => dir + path }) match { case Some(file) => file case None => error(cat_lines( ("Failed to find font file " + path + " in directories:") :: dirs.map(dir => " " + dir.toString))) } } val default_sources: List[Family] = List(Family.deja_vu_sans_mono, Family.deja_vu_sans, Family.deja_vu_serif) val default_target_dir: Path = Path.explode("isabelle_fonts") def build_fonts( sources: List[Family] = default_sources, source_dirs: List[Path] = Nil, target_prefix: String = "Isabelle", target_version: String = "", target_dir: Path = default_target_dir, progress: Progress = No_Progress) { progress.echo("Directory " + target_dir) hinting.foreach(hinted => Isabelle_System.mkdirs(target_dir + hinted_path(hinted))) val font_dirs = source_dirs ::: List(Path.explode("~~/Admin/isabelle_fonts")) for (dir <- font_dirs if !dir.is_dir) error("Bad source directory: " + dir) // Isabelle fonts val targets = for { source <- sources; (source_font, index) <- source.fonts.zipWithIndex } yield { val isabelle_file = find_file(font_dirs, Family.isabelle_symbols.get(index)) val source_file = find_file(font_dirs, source_font) val source_names = Fontforge.font_names(source_file) val target_names = source_names.update(prefix = target_prefix, version = target_version) Isabelle_System.with_tmp_file("font", "ttf")(tmp_file => { for (hinted <- hinting) { val target_file = target_dir + hinted_path(hinted) + target_names.ttf progress.echo("Font " + target_file.toString + " ...") if (hinted) auto_hint(source_file, tmp_file) else File.copy(source_file, tmp_file) Fontforge.execute( Fontforge.commands( Fontforge.open(isabelle_file), Fontforge.select(Range.isabelle_font), Fontforge.copy, Fontforge.close, Fontforge.open(tmp_file), Fontforge.select(Range.base_font), Fontforge.select_invert, Fontforge.clear, Fontforge.select(Range.isabelle_font), Fontforge.paste, target_names.set, Fontforge.generate(target_file), Fontforge.close) ).check } }) (target_names.ttf, index) } // Vacuous font { val vacuous_file = find_file(font_dirs, Family.vacuous.get(0)) val target_names = Fontforge.font_names(vacuous_file) val target_file = target_dir + hinted_path(false) + target_names.ttf progress.echo("Font " + target_file.toString + " ...") val domain = (for ((name, index) <- targets if index == 0) yield Fontforge.font_domain(target_dir + hinted_path(false) + name)) - .flatten.toSet.toList.sorted + .flatten.distinct.sorted Fontforge.execute( Fontforge.commands( Fontforge.open(vacuous_file), Fontforge.select(Range.vacuous_font), Fontforge.copy) + domain.map(code => Fontforge.commands( Fontforge.select(Seq(code)), Fontforge.paste)) .mkString("\n", "\n", "\n") + Fontforge.commands( Fontforge.generate(target_file), Fontforge.close) ).check } // etc/settings val settings_path = Components.settings(target_dir) Isabelle_System.mkdirs(settings_path.dir) def fonts_settings(hinted: Boolean): String = "\n isabelle_fonts \\\n" + (for ((target, _) <- targets) yield """ "$COMPONENT/""" + hinted_path(hinted).file_name + "/" + target.file_name + "\"") .mkString(" \\\n") File.write(settings_path, """# -*- shell-script -*- :mode=shellscript: if grep "isabelle_fonts_hinted.*=.*false" "$ISABELLE_HOME_USER/etc/preferences" >/dev/null 2>/dev/null then""" + fonts_settings(false) + """ else""" + fonts_settings(true) + """ fi isabelle_fonts_hidden "$COMPONENT/""" + hinted_path(false).file_name + """/Vacuous.ttf" """) // README File.copy(Path.explode("~~/Admin/isabelle_fonts/README"), target_dir) } /* Isabelle tool wrapper */ val isabelle_tool = Isabelle_Tool("build_fonts", "construct Isabelle fonts", args => { var source_dirs: List[Path] = Nil val getopts = Getopts(""" Usage: isabelle build_fonts [OPTIONS] Options are: -d DIR additional source directory Construct Isabelle fonts from DejaVu font families and Isabelle symbols. """, "d:" -> (arg => source_dirs = source_dirs ::: List(Path.explode(arg)))) val more_args = getopts(args) if (more_args.nonEmpty) getopts.usage() val target_version = Date.Format("uuuuMMdd")(Date.now()) val target_dir = Path.explode("isabelle_fonts-" + target_version) val progress = new Console_Progress build_fonts(source_dirs = source_dirs, target_dir = target_dir, target_version = target_version, progress = progress) }) } 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,595 +1,595 @@ /* Title: Pure/Admin/build_history.scala Author: Makarius Build other history versions. */ package isabelle import java.io.{File => JFile} import java.time.format.DateTimeFormatter import java.util.Locale object Build_History { /* log files */ val engine = "build_history" val log_prefix = engine + "_" val META_INFO_MARKER = "\fmeta_info = " /* augment settings */ def augment_settings( other_isabelle: Other_Isabelle, threads: Int, arch_64: Boolean, 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 = No_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, 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, 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 afp = AFP.init(options, afp_repos) (List("-d", "~~/AFP/thys"), afp.partition(afp_partition)) } /* 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( components_base = components_base, catalogs = List("main", "optional")) other_isabelle.init_settings(component_settings ::: init_settings) other_isabelle.resolve_components(verbose) val ml_platform = augment_settings(other_isabelle, threads, arch_64, heap, max_heap, more_settings) 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(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.mkdirs(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.mkdirs(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_result = Other_Isabelle(root, isabelle_identifier = isabelle_identifier, user_home = user_home, progress = build_out_progress)( "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 => Build_Log.Log_File.print_props(Build_Log.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.ext("xz") + " ...") File.write_xz(log_path.ext("xz"), terminate_lines( Build_Log.Log_File.print_props(META_INFO_MARKER, meta_info) :: build_result.out_lines ::: session_build_info ::: ml_statistics.map(Build_Log.Log_File.print_props(Build_Log.ML_STATISTICS_MARKER, _)) ::: session_errors.map(Build_Log.Log_File.print_props(Build_Log.ERROR_MESSAGE_MARKER, _)) ::: heap_sizes), XZ.options(6)) /* next build */ if (multicore_base && first_build && isabelle_output_log.is_dir) Isabelle_System.copy_dir(isabelle_output_log, isabelle_base_log) Isabelle_System.rm_tree(isabelle_output) first_build = false (build_result, log_path.ext("xz")) } } /* command line entry point */ private object Multicore { private val Pat1 = """^(\d+)$""".r private val Pat2 = """^(\d+)x(\d+)$""".r def parse(s: String): (Int, Int) = s match { case Pat1(Value.Int(x)) => (x, 1) case Pat2(Value.Int(x), Value.Int(y)) => (x, y) case _ => error("Bad multicore configuration: " + quote(s)) } } def main(args: Array[String]) { Command_Line.tool0 { var 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 more_settings: 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) -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) -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(_))), + "M:" -> (arg => multicore_list = space_explode(',', arg).map(Multicore.parse)), "N:" -> (arg => isabelle_identifier = arg), "P:" -> (arg => afp_partition = Value.Int.parse(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), "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, 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, 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 } 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 = No_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) } val other_hg = 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") val afp_hg = 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") execute("Admin/build_history", "-o " + ssh.bash_path(output_file) + (if (rev == "") "" else " -r " + Bash.string(rev_id)) + " " + options + afp_options + " " + ssh.bash_path(isabelle_repos_other) + " " + args, echo = true, strict = false) 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_jdk.scala b/src/Pure/Admin/build_jdk.scala --- a/src/Pure/Admin/build_jdk.scala +++ b/src/Pure/Admin/build_jdk.scala @@ -1,240 +1,240 @@ /* Title: Pure/Admin/build_jdk.scala Author: Makarius Build Isabelle jdk component from original platform installations. */ package isabelle import java.io.{File => JFile} import java.nio.file.Files import java.nio.file.attribute.PosixFilePermission import scala.util.matching.Regex object Build_JDK { /* version */ def detect_version(s: String): String = { val Version_Dir_Entry = """^jdk-([0-9.]+\+\d+)$""".r s match { case Version_Dir_Entry(version) => version case _ => error("Cannot detect JDK version from " + quote(s)) } } /* platform */ sealed case class JDK_Platform(name: String, home: String, exe: String, regex: Regex) { override def toString: String = name def detect(jdk_dir: Path): Boolean = { val path = jdk_dir + Path.explode(exe) if (path.is_file) { val file_descr = Isabelle_System.bash("file -b " + File.bash_path(path)).check.out regex.pattern.matcher(file_descr).matches } else false } } val jdk_platforms = List( JDK_Platform("x86_64-linux", ".", "bin/java", """.*ELF 64-bit.*x86[-_]64.*""".r), JDK_Platform("x86_64-windows", ".", "bin/java.exe", """.*PE32\+ executable.*x86[-_]64.*""".r), JDK_Platform("x86_64-darwin", "Contents/Home", "Contents/Home/bin/java", """.*Mach-O 64-bit.*x86[-_]64.*""".r)) /* README */ def readme(version: String): String = """This is OpenJDK """ + version + """ as required for Isabelle. See https://adoptopenjdk.net for the original downloads, which are covered the GPL2 (with various liberal exceptions, see legal/*). Linux, Windows, Mac OS X all work uniformly, depending on certain platform-specific subdirectories. """ /* settings */ val settings = """# -*- shell-script -*- :mode=shellscript: case "$ISABELLE_PLATFORM_FAMILY" in linux) ISABELLE_JAVA_PLATFORM="$ISABELLE_PLATFORM64" ISABELLE_JDK_HOME="$COMPONENT/$ISABELLE_JAVA_PLATFORM" ;; windows) ISABELLE_JAVA_PLATFORM="$ISABELLE_WINDOWS_PLATFORM64" ISABELLE_JDK_HOME="$COMPONENT/$ISABELLE_JAVA_PLATFORM" ;; macos) ISABELLE_JAVA_PLATFORM="$ISABELLE_PLATFORM64" ISABELLE_JDK_HOME="$COMPONENT/$ISABELLE_JAVA_PLATFORM/Contents/Home" ;; esac """ /* extract archive */ private def suppress_name(name: String): Boolean = name.startsWith("._") def extract_archive(dir: Path, archive: Path): (String, JDK_Platform) = { try { val tmp_dir = dir + Path.explode("tmp") Isabelle_System.mkdirs(tmp_dir) if (archive.get_ext == "zip") { Isabelle_System.bash( "unzip -x " + File.bash_path(archive.absolute), cwd = tmp_dir.file).check } else { Isabelle_System.gnutar("-xzf " + File.bash_path(archive), dir = tmp_dir).check } val dir_entry = - File.read_dir(tmp_dir).filterNot(suppress_name(_)) match { + File.read_dir(tmp_dir).filterNot(suppress_name) match { case List(s) => s case _ => error("Archive contains multiple directories") } val version = detect_version(dir_entry) val jdk_dir = tmp_dir + Path.explode(dir_entry) val platform = jdk_platforms.find(_.detect(jdk_dir)) getOrElse error("Failed to detect JDK platform") val platform_dir = dir + Path.explode(platform.name) if (platform_dir.is_dir) error("Directory already exists: " + platform_dir) File.link(Path.current, jdk_dir + Path.explode(platform.home) + Path.explode("jre")) File.move(jdk_dir, platform_dir) (version, platform) } catch { case ERROR(msg) => cat_error(msg, "The error(s) above occurred for " + archive) } } /* build jdk */ def build_jdk( archives: List[Path], progress: Progress = No_Progress, target_dir: Path = Path.current) { if (Platform.is_windows) error("Cannot build jdk on Windows") Isabelle_System.with_tmp_dir("jdk")(dir => { progress.echo("Extracting ...") val extracted = archives.map(extract_archive(dir, _)) val version = - extracted.map(_._1).toSet.toList match { + extracted.map(_._1).distinct match { case List(version) => version case Nil => error("No archives") case versions => error("Archives contain multiple JDK versions: " + commas_quote(versions)) } val missing_platforms = jdk_platforms.filterNot(p1 => extracted.exists({ case (_, p2) => p1.name == p2.name })) if (missing_platforms.nonEmpty) error("Missing platforms: " + commas_quote(missing_platforms.map(_.name))) val jdk_name = "jdk-" + version val jdk_path = Path.explode(jdk_name) val component_dir = dir + jdk_path Isabelle_System.mkdirs(component_dir + Path.explode("etc")) File.write(Components.settings(component_dir), settings) File.write(component_dir + Path.explode("README"), readme(version)) for ((_, platform) <- extracted) File.move(dir + Path.explode(platform.name), component_dir) for (file <- File.find_files(component_dir.file, include_dirs = true)) { val path = file.toPath val perms = Files.getPosixFilePermissions(path) perms.add(PosixFilePermission.OWNER_READ) perms.add(PosixFilePermission.GROUP_READ) perms.add(PosixFilePermission.OTHERS_READ) perms.add(PosixFilePermission.OWNER_WRITE) if (file.isDirectory) { perms.add(PosixFilePermission.OWNER_WRITE) perms.add(PosixFilePermission.OWNER_EXECUTE) perms.add(PosixFilePermission.GROUP_EXECUTE) perms.add(PosixFilePermission.OTHERS_EXECUTE) } Files.setPosixFilePermissions(path, perms) } File.find_files((component_dir + Path.explode("x86_64-darwin")).file, file => suppress_name(file.getName)).foreach(_.delete) progress.echo("Sharing ...") val main_dir :: other_dirs = jdk_platforms.map(platform => (component_dir + Path.explode(platform.name)).file.toPath) for { file1 <- File.find_files(main_dir.toFile).iterator path1 = file1.toPath dir2 <- other_dirs.iterator } { val path2 = dir2.resolve(main_dir.relativize(path1)) val file2 = path2.toFile if (!Files.isSymbolicLink(path2) && file2.isFile && File.eq_content(file1, file2)) { file2.delete Files.createLink(path2, path1) } } progress.echo("Archiving ...") Isabelle_System.gnutar( "-czf " + File.bash_path(target_dir + jdk_path.ext("tar.gz")) + " " + jdk_name, dir = dir).check }) } /* Isabelle tool wrapper */ val isabelle_tool = Isabelle_Tool("build_jdk", "build Isabelle jdk component from original archives", args => { var target_dir = Path.current val getopts = Getopts(""" Usage: isabelle build_jdk [OPTIONS] ARCHIVES... Options are: -D DIR target directory (default ".") Build jdk component from tar.gz archives, with original jdk archives for x86_64 Linux, Windows, Mac OS X. """, "D:" -> (arg => target_dir = Path.explode(arg))) val more_args = getopts(args) if (more_args.isEmpty) getopts.usage() - val archives = more_args.map(Path.explode(_)) + val archives = more_args.map(Path.explode) val progress = new Console_Progress() build_jdk(archives = archives, progress = progress, target_dir = target_dir) }) } diff --git a/src/Pure/Admin/build_release.scala b/src/Pure/Admin/build_release.scala --- a/src/Pure/Admin/build_release.scala +++ b/src/Pure/Admin/build_release.scala @@ -1,849 +1,848 @@ /* Title: Pure/Admin/build_release.scala Author: Makarius Build full Isabelle distribution from repository. */ package isabelle object Build_Release { /** release info **/ sealed case class Bundle_Info( platform: Platform.Family.Value, platform_description: String, name: String) { def path: Path = Path.explode(name) } class Release private[Build_Release]( progress: Progress, val date: Date, val dist_name: String, val dist_dir: Path, val dist_version: String, val ident: String) { val isabelle_dir: Path = dist_dir + Path.explode(dist_name) val isabelle_archive: Path = dist_dir + Path.explode(dist_name + ".tar.gz") val isabelle_library_archive: Path = dist_dir + Path.explode(dist_name + "_library.tar.gz") def other_isabelle(dir: Path): Other_Isabelle = Other_Isabelle(dir + Path.explode(dist_name), isabelle_identifier = dist_name + "-build", progress = progress) def bundle_info(platform: Platform.Family.Value): Bundle_Info = platform match { case Platform.Family.linux => Bundle_Info(platform, "Linux", dist_name + "_linux.tar.gz") case Platform.Family.macos => Bundle_Info(platform, "macOS", dist_name + "_macos.tar.gz") case Platform.Family.windows => Bundle_Info(platform, "Windows", dist_name + ".exe") } } /** generated content **/ /* patch release */ private def change_file(dir: Path, name: String, f: String => String) { val file = dir + Path.explode(name) File.write(file, f(File.read(file))) } private val getsettings_file: String = "lib/scripts/getsettings" private val ISABELLE_ID = """ISABELLE_ID="(.+)"""".r def patch_release(release: Release, is_official: Boolean) { val dir = release.isabelle_dir for (name <- List("src/Pure/System/distribution.ML", "src/Pure/System/distribution.scala")) { change_file(dir, name, s => s.replaceAllLiterally("val is_identified = false", "val is_identified = true") .replaceAllLiterally("val is_official = false", "val is_official = " + is_official)) } change_file(dir, getsettings_file, s => s.replaceAllLiterally("ISABELLE_ID=\"\"", "ISABELLE_ID=" + quote(release.ident)) .replaceAllLiterally("ISABELLE_IDENTIFIER=\"\"", "ISABELLE_IDENTIFIER=" + quote(release.dist_name))) change_file(dir, "lib/html/library_index_header.template", s => s.replaceAllLiterally("{ISABELLE}", release.dist_name)) for { name <- List( "src/Pure/System/distribution.ML", "src/Pure/System/distribution.scala", "lib/Tools/version") } { change_file(dir, name, s => s.replaceAllLiterally("repository version", release.dist_version)) } change_file(dir, "README", s => s.replaceAllLiterally("some repository version of Isabelle", release.dist_version)) } /* ANNOUNCE */ def make_announce(release: Release) { File.write(release.isabelle_dir + Path.explode("ANNOUNCE"), """ IMPORTANT NOTE ============== This is a snapshot of Isabelle/""" + release.ident + """ from the repository. """) } /* NEWS */ def make_news(other_isabelle: Other_Isabelle, dist_version: String) { val target = other_isabelle.isabelle_home + Path.explode("doc") val target_fonts = target + Path.explode("fonts") Isabelle_System.mkdirs(target_fonts) other_isabelle.copy_fonts(target_fonts) HTML.write_document(target, "NEWS.html", List(HTML.title("NEWS (" + dist_version + ")")), List( HTML.chapter("NEWS"), HTML.source( Symbol.decode(File.read(other_isabelle.isabelle_home + Path.explode("NEWS")))))) } /* bundled components */ class Bundled(platform: Option[Platform.Family.Value] = None) { def detect(s: String): Boolean = s.startsWith("#bundled") && !s.startsWith("#bundled ") def apply(name: String): String = "#bundled" + (platform match { case None => "" case Some(plat) => "-" + plat }) + ":" + name private val Pattern1 = ("""^#bundled:(.*)$""").r private val Pattern2 = ("""^#bundled-(.*):(.*)$""").r def unapply(s: String): Option[String] = s match { case Pattern1(name) => Some(name) case Pattern2(Platform.Family(plat), name) if platform == Some(plat) => Some(name) case _ => None } } def record_bundled_components(dir: Path) { val catalogs = List("main", "bundled").map((_, new Bundled())) ::: default_platform_families.flatMap(platform => List(platform.toString, "bundled-" + platform.toString). map((_, new Bundled(platform = Some(platform))))) File.append(Components.components(dir), terminate_lines("#bundled components" :: (for { (catalog, bundled) <- catalogs.iterator - val path = Components.admin(dir) + Path.basic(catalog) + path = Components.admin(dir) + Path.basic(catalog) if path.is_file line <- split_lines(File.read(path)) if line.nonEmpty && !line.startsWith("#") && !line.startsWith("jedit_build") } yield bundled(line)).toList)) } def get_bundled_components(dir: Path, platform: Platform.Family.Value): (List[String], String) = { val Bundled = new Bundled(platform = Some(platform)) val components = for { Bundled(name) <- Components.read_components(dir) if !name.startsWith("jedit_build") } yield name val jdk_component = components.find(_.startsWith("jdk")) getOrElse error("Missing jdk component") (components, jdk_component) } def activate_components(dir: Path, platform: Platform.Family.Value, more_names: List[String]) { def contrib_name(name: String): String = Components.contrib(name = name).implode val Bundled = new Bundled(platform = Some(platform)) Components.write_components(dir, Components.read_components(dir).flatMap(line => line match { case Bundled(name) => if (Components.check_dir(Components.contrib(dir, name))) Some(contrib_name(name)) else None case _ => if (Bundled.detect(line)) None else Some(line) - }) ::: more_names.map(contrib_name(_))) + }) ::: more_names.map(contrib_name)) } def make_contrib(dir: Path) { Isabelle_System.mkdirs(Components.contrib(dir)) File.write(Components.contrib(dir, "README"), """This directory contains add-on components that contribute to the main Isabelle distribution. Separate licensing conditions apply, see each directory individually. """) } /** build release **/ private def execute(dir: Path, script: String): Unit = Isabelle_System.bash(script, cwd = dir.file).check private def execute_tar(dir: Path, args: String): Unit = Isabelle_System.gnutar(args, dir = dir).check /* build heaps on remote server */ private def remote_build_heaps( options: Options, platform: Platform.Family.Value, build_sessions: List[String], local_dir: Path) { val server_option = "build_release_server_" + platform.toString options.string(server_option) match { case SSH.Target(user, host) => using(SSH.open_session(options, host = host, user = user))(ssh => { Isabelle_System.with_tmp_file("tmp", "tar")(local_tmp_tar => { execute_tar(local_dir, "-cf " + File.bash_path(local_tmp_tar) + " .") ssh.with_tmp_dir(remote_dir => { val remote_tmp_tar = remote_dir + Path.basic("tmp.tar") ssh.write_file(remote_tmp_tar, local_tmp_tar) val remote_commands = List( "cd " + File.bash_path(remote_dir), "tar -xf tmp.tar", "./bin/isabelle build -o system_heaps -b -- " + Bash.strings(build_sessions), "tar -cf tmp.tar heaps") ssh.execute(remote_commands.mkString(" && ")).check ssh.read_file(remote_tmp_tar, local_tmp_tar) }) execute_tar(local_dir, "-xf " + File.bash_path(local_tmp_tar)) }) }) case s => error("Bad " + server_option + ": " + quote(s)) } } /* main */ private val default_platform_families: List[Platform.Family.Value] = List(Platform.Family.linux, Platform.Family.windows, Platform.Family.macos) def build_release(base_dir: Path, options: Options, components_base: Path = Components.default_components_base, progress: Progress = No_Progress, rev: String = "", afp_rev: String = "", official_release: Boolean = false, proper_release_name: Option[String] = None, platform_families: List[Platform.Family.Value] = default_platform_families, more_components: List[Path] = Nil, website: Option[Path] = None, build_sessions: List[String] = Nil, build_library: Boolean = false, parallel_jobs: Int = 1): Release = { val hg = Mercurial.repository(Path.explode("$ISABELLE_HOME")) val release = { val date = Date.now() val dist_name = proper_release_name getOrElse ("Isabelle_" + Date.Format.date(date)) val dist_dir = (base_dir + Path.explode("dist-" + dist_name)).absolute val version = proper_string(rev) orElse proper_release_name getOrElse "tip" val ident = try { hg.id(version) } catch { case ERROR(_) => error("Bad repository version: " + version) } val dist_version = proper_release_name match { case Some(name) => name + ": " + Date.Format("LLLL uuuu")(date) case None => "Isabelle repository snapshot " + ident + " " + Date.Format.date(date) } new Release(progress, date, dist_name, dist_dir, dist_version, ident) } /* make distribution */ if (release.isabelle_archive.is_file) { progress.echo_warning("Release archive already exists: " + release.isabelle_archive) val archive_ident = Isabelle_System.with_tmp_dir("build_release")(tmp_dir => { val getsettings = Path.explode(release.dist_name + "/" + getsettings_file) execute_tar(tmp_dir, "-xzf " + File.bash_path(release.isabelle_archive) + " " + File.bash_path(getsettings)) split_lines(File.read(tmp_dir + getsettings)) .collectFirst({ case ISABELLE_ID(ident) => ident }) .getOrElse(error("Failed to read ISABELLE_ID from " + release.isabelle_archive)) }) if (release.ident != archive_ident) { error("Mismatch of release identification " + release.ident + " vs. archive " + archive_ident) } } else { progress.echo_warning("Producing release archive " + release.isabelle_archive + " ...") Isabelle_System.mkdirs(release.dist_dir) if (release.isabelle_dir.is_dir) error("Directory " + release.isabelle_dir + " already exists") progress.echo_warning("Retrieving Mercurial repository version " + release.ident) hg.archive(release.isabelle_dir.expand.implode, rev = release.ident, options = "--type files") for (name <- List(".hg_archival.txt", ".hgtags", ".hgignore", "README_REPOSITORY")) { (release.isabelle_dir + Path.explode(name)).file.delete } progress.echo_warning("Preparing distribution " + quote(release.dist_name)) patch_release(release, proper_release_name.isDefined && official_release) if (proper_release_name.isEmpty) make_announce(release) make_contrib(release.isabelle_dir) execute(release.isabelle_dir, """find . -print | xargs chmod -f u+rw""") record_bundled_components(release.isabelle_dir) /* build tools and documentation */ val other_isabelle = release.other_isabelle(release.dist_dir) other_isabelle.init_settings( other_isabelle.init_components(components_base = components_base, catalogs = List("main"))) other_isabelle.resolve_components(echo = true) try { val export_classpath = "export CLASSPATH=" + Bash.string(other_isabelle.getenv("ISABELLE_CLASSPATH")) + "\n" other_isabelle.bash(export_classpath + "./Admin/build all", echo = true).check other_isabelle.bash(export_classpath + "./bin/isabelle jedit -b", echo = true).check } catch { case ERROR(_) => error("Failed to build tools") } try { other_isabelle.bash( "./bin/isabelle build_doc -a -o system_heaps -j " + parallel_jobs, echo = true).check } catch { case ERROR(_) => error("Failed to build documentation") } make_news(other_isabelle, release.dist_version) for (name <- List("Admin", "browser_info", "heaps")) { Isabelle_System.rm_tree(other_isabelle.isabelle_home + Path.explode(name)) } other_isabelle.cleanup() progress.echo_warning("Creating distribution archive " + release.isabelle_archive) def execute_dist_name(script: String): Unit = Isabelle_System.bash(script, cwd = release.dist_dir.file, env = Isabelle_System.settings() + ("DIST_NAME" -> release.dist_name)).check execute_dist_name(""" set -e chmod -R a+r "$DIST_NAME" chmod -R u+w "$DIST_NAME" chmod -R g=o "$DIST_NAME" find "$DIST_NAME" -type f "(" -name "*.thy" -o -name "*.ML" -o -name "*.scala" ")" -print | xargs chmod -f u-w """) execute_tar(release.dist_dir, "-czf " + File.bash_path(release.isabelle_archive) + " " + Bash.string(release.dist_name)) execute_dist_name(""" set -e mv "$DIST_NAME" "${DIST_NAME}-old" mkdir "$DIST_NAME" mv "${DIST_NAME}-old/README" "${DIST_NAME}-old/NEWS" "${DIST_NAME}-old/ANNOUNCE" \ "${DIST_NAME}-old/COPYRIGHT" "${DIST_NAME}-old/CONTRIBUTORS" "$DIST_NAME" mkdir "$DIST_NAME/doc" mv "${DIST_NAME}-old/doc/"*.pdf \ "${DIST_NAME}-old/doc/"*.html \ "${DIST_NAME}-old/doc/"*.css \ "${DIST_NAME}-old/doc/fonts" \ "${DIST_NAME}-old/doc/Contents" "$DIST_NAME/doc" rm -f Isabelle && ln -sf "$DIST_NAME" Isabelle rm -rf "${DIST_NAME}-old" """) } /* make application bundles */ val bundle_infos = platform_families.map(release.bundle_info) for (bundle_info <- bundle_infos) { - val bundle_archive = release.dist_dir + bundle_info.path val isabelle_name = release.dist_name val platform = bundle_info.platform progress.echo("\nApplication bundle for " + platform) Isabelle_System.with_tmp_dir("build_release")(tmp_dir => { // release archive execute_tar(tmp_dir, "-xzf " + File.bash_path(release.isabelle_archive)) val other_isabelle = release.other_isabelle(tmp_dir) val isabelle_target = other_isabelle.isabelle_home // bundled components progress.echo("Bundled components:") val contrib_dir = Components.contrib(isabelle_target) val (bundled_components, jdk_component) = get_bundled_components(isabelle_target, platform) Components.resolve(components_base, bundled_components, target_dir = Some(contrib_dir), copy_dir = Some(release.dist_dir + Path.explode("contrib")), progress = progress) val more_components_names = more_components.map(Components.unpack(contrib_dir, _, progress = progress)) Components.purge(contrib_dir, platform) activate_components(isabelle_target, platform, more_components_names) // Java parameters val java_options_title = "# Java runtime options" val java_options: List[String] = (for { variable <- List( "ISABELLE_JAVA_SYSTEM_OPTIONS", "JEDIT_JAVA_SYSTEM_OPTIONS", "JEDIT_JAVA_OPTIONS") opt <- Word.explode(other_isabelle.getenv(variable)) } yield opt) ::: List("-Disabelle.jedit_server=" + isabelle_name) val classpath: List[Path] = { val base = isabelle_target.absolute Path.split(other_isabelle.getenv("ISABELLE_CLASSPATH")).map(path => { val abs_path = path.absolute File.relative_path(base, abs_path) match { case Some(rel_path) => rel_path case None => error("Bad ISABELLE_CLASSPATH element: " + abs_path) } }) ::: List(Path.explode("src/Tools/jEdit/dist/jedit.jar")) } val jedit_options = Path.explode("src/Tools/jEdit/etc/options") val jedit_props = Path.explode("src/Tools/jEdit/dist/properties/jEdit.props") // build heaps if (build_sessions.nonEmpty) { progress.echo("Building heaps ...") remote_build_heaps(options, platform, build_sessions, isabelle_target) } // application bundling platform match { case Platform.Family.linux => File.write(isabelle_target + jedit_options, File.read(isabelle_target + jedit_options) .replaceAll("jedit_reset_font_size : int =.*", "jedit_reset_font_size : int = 24")) File.write(isabelle_target + jedit_props, File.read(isabelle_target + jedit_props) .replaceAll("console.fontsize=.*", "console.fontsize=18") .replaceAll("helpviewer.fontsize=.*", "helpviewer.fontsize=18") .replaceAll("metal.primary.fontsize=.*", "metal.primary.fontsize=18") .replaceAll("metal.secondary.fontsize=.*", "metal.secondary.fontsize=18") .replaceAll("view.fontsize=.*", "view.fontsize=24") .replaceAll("view.gutter.fontsize=.*", "view.gutter.fontsize=16")) File.write(isabelle_target + Path.explode("Isabelle.options"), terminate_lines(java_options_title :: java_options)) val isabelle_app = isabelle_target + Path.explode("lib/scripts/Isabelle_app") File.write(isabelle_app, File.read(Path.explode("~~/Admin/Linux/Isabelle_app")) .replaceAllLiterally("{CLASSPATH}", classpath.map("$ISABELLE_HOME/" + _).mkString(":")) .replaceAllLiterally("/jdk/", "/" + jdk_component + "/")) File.set_executable(isabelle_app, true) val linux_app = isabelle_target + Path.explode("contrib/linux_app") File.move(linux_app + Path.explode("Isabelle"), isabelle_target + Path.explode(isabelle_name)) Isabelle_System.rm_tree(linux_app) val archive_name = isabelle_name + "_linux.tar.gz" progress.echo("Packaging " + archive_name + " ...") execute_tar(tmp_dir, "-czf " + File.bash_path(release.dist_dir + Path.explode(archive_name)) + " " + Bash.string(isabelle_name)) case Platform.Family.macos => File.write(isabelle_target + jedit_props, File.read(isabelle_target + jedit_props) .replaceAll("lookAndFeel=.*", "lookAndFeel=com.apple.laf.AquaLookAndFeel") .replaceAll("delete-line.shortcut=.*", "delete-line.shortcut=C+d") .replaceAll("delete.shortcut2=.*", "delete.shortcut2=A+d") .replaceAll("plugin-blacklist.MacOSX.jar=true", "plugin-blacklist.MacOSX.jar=")) // MacOS application bundle File.move(isabelle_target + Path.explode("contrib/macos_app"), tmp_dir) val isabelle_app = Path.explode(isabelle_name + ".app") val app_dir = tmp_dir + isabelle_app File.move(tmp_dir + Path.explode("macos_app/Isabelle.app"), app_dir) val app_contents = app_dir + Path.explode("Contents") val app_resources = app_contents + Path.explode("Resources") File.move(tmp_dir + Path.explode(isabelle_name), app_resources) File.write(app_contents + Path.explode("Info.plist"), File.read(Path.explode("~~/Admin/MacOS/Info.plist")) .replaceAllLiterally("{ISABELLE_NAME}", isabelle_name) .replaceAllLiterally("{JAVA_OPTIONS}", terminate_lines(java_options.map(opt => "" + opt + "")))) for (cp <- classpath) { File.link( Path.explode("../Resources/" + isabelle_name + "/") + cp, app_contents + Path.explode("Java"), force = true) } File.link( Path.explode("../Resources/" + isabelle_name + "/contrib/" + jdk_component + "/x86_64-darwin"), app_contents + Path.explode("PlugIns/bundled.jdk"), force = true) File.link( Path.explode("../../Info.plist"), app_resources + Path.explode(isabelle_name + "/" + isabelle_name + ".plist"), force = true) File.link( Path.explode("Contents/Resources/" + isabelle_name), app_dir + Path.explode("Isabelle"), force = true) // application archive val archive_name = isabelle_name + "_macos.tar.gz" progress.echo("Packaging " + archive_name + " ...") execute_tar(tmp_dir, "-czf " + File.bash_path(release.dist_dir + Path.explode(archive_name)) + " " + File.bash_path(isabelle_app)) case Platform.Family.windows => File.write(isabelle_target + jedit_props, File.read(isabelle_target + jedit_props) .replaceAll("lookAndFeel=.*", "lookAndFeel=com.sun.java.swing.plaf.windows.WindowsLookAndFeel") .replaceAll("foldPainter=.*", "foldPainter=Square")) // application launcher File.move(isabelle_target + Path.explode("contrib/windows_app"), tmp_dir) val app_template = Path.explode("~~/Admin/Windows/launch4j") File.write(isabelle_target + Path.explode(isabelle_name + ".l4j.ini"), (java_options_title :: java_options).map(_ + "\r\n").mkString) val isabelle_xml = Path.explode("isabelle.xml") val isabelle_exe = Path.explode(isabelle_name + ".exe") File.write(tmp_dir + isabelle_xml, File.read(app_template + isabelle_xml) .replaceAllLiterally("{ISABELLE_NAME}", isabelle_name) .replaceAllLiterally("{OUTFILE}", File.platform_path(isabelle_target + isabelle_exe)) .replaceAllLiterally("{ICON}", File.platform_path(app_template + Path.explode("isabelle_transparent.ico"))) .replaceAllLiterally("{SPLASH}", File.platform_path(app_template + Path.explode("isabelle.bmp"))) .replaceAllLiterally("{CLASSPATH}", cat_lines(classpath.map(cp => " %EXEDIR%\\" + File.platform_path(cp).replace('/', '\\') + ""))) .replaceAllLiterally("\\jdk\\", "\\" + jdk_component + "\\")) execute(tmp_dir, "\"windows_app/launch4j-${ISABELLE_PLATFORM_FAMILY}/launch4j\" isabelle.xml") File.copy(app_template + Path.explode("manifest.xml"), isabelle_target + isabelle_exe.ext("manifest")) // Cygwin setup val cygwin_template = Path.explode("~~/Admin/Windows/Cygwin") File.copy(cygwin_template + Path.explode("Cygwin-Terminal.bat"), isabelle_target) val cygwin_mirror = File.read(isabelle_target + Path.explode("contrib/cygwin/isabelle/cygwin_mirror")) val cygwin_bat = Path.explode("Cygwin-Setup.bat") File.write(isabelle_target + cygwin_bat, File.read(cygwin_template + cygwin_bat) .replaceAllLiterally("{MIRROR}", cygwin_mirror)) File.set_executable(isabelle_target + cygwin_bat, true) for (name <- List("isabelle/postinstall", "isabelle/rebaseall")) { val path = Path.explode(name) File.copy(cygwin_template + path, isabelle_target + Path.explode("contrib/cygwin") + path) } execute(isabelle_target, """find . -type f -not -name "*.exe" -not -name "*.dll" """ + (if (Platform.is_macos) "-perm +100" else "-executable") + " -print0 > contrib/cygwin/isabelle/executables") execute(isabelle_target, """find . -type l -exec echo "{}" ";" -exec readlink "{}" ";" """ + """> contrib/cygwin/isabelle/symlinks""") execute(isabelle_target, """find . -type l -exec rm "{}" ";" """) File.write(isabelle_target + Path.explode("contrib/cygwin/isabelle/uninitialized"), "") // executable archive (self-extracting 7z) val archive_name = isabelle_name + ".7z" val exe_archive = tmp_dir + Path.explode(archive_name) exe_archive.file.delete progress.echo("Packaging " + archive_name + " ...") execute(tmp_dir, "7z -y -bd a " + File.bash_path(exe_archive) + " " + Bash.string(isabelle_name)) if (!exe_archive.is_file) error("Failed to create archive: " + exe_archive) val sfx_exe = tmp_dir + Path.explode("windows_app/7zsd_All_x64.sfx") val sfx_txt = File.read(Path.explode("~~/Admin/Windows/Installer/sfx.txt")). replaceAllLiterally("{ISABELLE_NAME}", isabelle_name) Bytes.write(release.dist_dir + isabelle_exe, Bytes.read(sfx_exe) + Bytes(sfx_txt) + Bytes.read(exe_archive)) File.set_executable(release.dist_dir + isabelle_exe, true) } }) progress.echo("DONE") } /* minimal website */ for (dir <- website) { val website_platform_bundles = for { bundle_info <- bundle_infos if (release.dist_dir + bundle_info.path).is_file } yield (bundle_info.name, bundle_info) val isabelle_link = HTML.link(Isabelle_Cronjob.isabelle_repos_source + "/rev/" + release.ident, HTML.text("Isabelle/" + release.ident)) val afp_link = HTML.link(AFP.repos_source + "/rev/" + afp_rev, HTML.text("AFP/" + afp_rev)) HTML.write_document(dir, "index.html", List(HTML.title(release.dist_name)), List( HTML.section(release.dist_name + " (" + release.ident + ")"), HTML.subsection("Platforms"), HTML.itemize( website_platform_bundles.map({ case (bundle, bundle_info) => List(HTML.link(bundle, HTML.text(bundle_info.platform_description))) })), HTML.subsection("Repositories"), HTML.itemize( List(List(isabelle_link)) ::: (if (afp_rev == "") Nil else List(List(afp_link)))))) for ((bundle, _) <- website_platform_bundles) File.copy(release.dist_dir + Path.explode(bundle), dir) } /* HTML library */ if (build_library) { if (release.isabelle_library_archive.is_file) { progress.echo_warning("Library archive already exists: " + release.isabelle_library_archive) } else { Isabelle_System.with_tmp_dir("build_release")(tmp_dir => { val bundle = release.dist_dir + Path.explode(release.dist_name + "_" + Platform.family + ".tar.gz") execute_tar(tmp_dir, "-xzf " + File.bash_path(bundle)) val other_isabelle = release.other_isabelle(tmp_dir) Isabelle_System.mkdirs(other_isabelle.etc) File.write(other_isabelle.etc_preferences, "ML_system_64 = true\n") other_isabelle.bash("bin/isabelle build -f -j " + parallel_jobs + " -o browser_info -o document=pdf -o document_variants=document:outline=/proof,/ML" + " -o system_heaps -c -a -d '~~/src/Benchmarks'", echo = true).check other_isabelle.isabelle_home_user.file.delete execute(tmp_dir, "chmod -R a+r " + Bash.string(release.dist_name)) execute(tmp_dir, "chmod -R g=o " + Bash.string(release.dist_name)) execute_tar(tmp_dir, "-czf " + File.bash_path(release.isabelle_library_archive) + " " + Bash.string(release.dist_name + "/browser_info")) }) } } release } /** command line entry point **/ def main(args: Array[String]) { Command_Line.tool0 { var afp_rev = "" var components_base: Path = Components.default_components_base var official_release = false var proper_release_name: Option[String] = None var website: Option[Path] = None var build_sessions: List[String] = Nil var more_components: List[Path] = Nil var parallel_jobs = 1 var build_library = false var options = Options.init() var platform_families = default_platform_families var rev = "" val getopts = Getopts(""" Usage: Admin/build_release [OPTIONS] BASE_DIR Options are: -A REV corresponding AFP changeset id -C DIR base directory for Isabelle components (default: """ + Components.default_components_base + """) -O official release (not release-candidate) -R RELEASE proper release with name -W WEBSITE produce minimal website in given directory -b SESSIONS build platform-specific session images (separated by commas) -c ARCHIVE clean bundling with additional component .tar.gz archive -j INT maximum number of parallel jobs (default 1) -l build library -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) -p NAMES platform families (default: """ + default_platform_families.mkString(",") + """) -r REV Mercurial changeset id (default: RELEASE or tip) Build Isabelle release in base directory, using the local repository clone. """, "A:" -> (arg => afp_rev = arg), "C:" -> (arg => components_base = Path.explode(arg)), "O" -> (_ => official_release = true), "R:" -> (arg => proper_release_name = Some(arg)), "W:" -> (arg => website = Some(Path.explode(arg))), "b:" -> (arg => build_sessions = space_explode(',', arg)), "c:" -> (arg => { val path = Path.explode(arg) Components.Archive.get_name(path.file_name) more_components = more_components ::: List(path) }), "j:" -> (arg => parallel_jobs = Value.Int.parse(arg)), "l" -> (_ => build_library = true), "o:" -> (arg => options = options + arg), "p:" -> (arg => platform_families = space_explode(',', arg).map(Platform.Family.parse)), "r:" -> (arg => rev = arg)) val more_args = getopts(args) val base_dir = more_args match { case List(base_dir) => base_dir case _ => getopts.usage() } val progress = new Console_Progress() if (platform_families.contains(Platform.Family.windows) && !Isabelle_System.bash("7z i").ok) error("Building for windows requires 7z") build_release(Path.explode(base_dir), options, components_base = components_base, progress = progress, rev = rev, afp_rev = afp_rev, official_release = official_release, proper_release_name = proper_release_name, website = website, platform_families = if (platform_families.isEmpty) default_platform_families else platform_families, more_components = more_components, build_sessions = build_sessions, build_library = build_library, parallel_jobs = parallel_jobs) } } } diff --git a/src/Pure/Admin/components.scala b/src/Pure/Admin/components.scala --- a/src/Pure/Admin/components.scala +++ b/src/Pure/Admin/components.scala @@ -1,301 +1,301 @@ /* Title: Pure/Admin/components.scala Author: Makarius Isabelle system components. */ package isabelle import java.io.{File => JFile} object Components { /* archive name */ object Archive { val suffix: String = ".tar.gz" def apply(name: String): String = if (name == "") error("Bad component name: " + quote(name)) else name + suffix def unapply(archive: String): Option[String] = { for { name0 <- Library.try_unsuffix(suffix, archive) name <- proper_string(name0) } yield name } def get_name(archive: String): String = unapply(archive) getOrElse error("Bad component archive name (expecting .tar.gz): " + quote(archive)) } /* component collections */ - val default_components_base = Path.explode("$ISABELLE_COMPONENTS_BASE") + val default_components_base: Path = Path.explode("$ISABELLE_COMPONENTS_BASE") def admin(dir: Path): Path = dir + Path.explode("Admin/components") def contrib(dir: Path = Path.current, name: String = ""): Path = dir + Path.explode("contrib") + Path.explode(name) def unpack(dir: Path, archive: Path, progress: Progress = No_Progress): String = { val name = Archive.get_name(archive.file_name) progress.echo("Unpacking " + name) Isabelle_System.gnutar("-xzf " + File.bash_path(archive), dir = dir).check name } def resolve(base_dir: Path, names: List[String], target_dir: Option[Path] = None, copy_dir: Option[Path] = None, progress: Progress = No_Progress) { Isabelle_System.mkdirs(base_dir) for (name <- names) { val archive_name = Archive(name) val archive = base_dir + Path.explode(archive_name) if (!archive.is_file) { val remote = Isabelle_System.getenv("ISABELLE_COMPONENT_REPOSITORY") + "/" + archive_name progress.echo("Getting " + remote) Bytes.write(archive, Url.read_bytes(Url(remote))) } for (dir <- copy_dir) { Isabelle_System.mkdirs(dir) File.copy(archive, dir) } unpack(target_dir getOrElse base_dir, archive, progress = progress) } } def purge(dir: Path, platform: Platform.Family.Value) { def purge_platforms(platforms: String*): Set[String] = platforms.flatMap(name => List("x86-" + name, "x86_64_32-" + name, "x86_64-" + name)).toSet + "ppc-darwin" val purge_set = platform match { case Platform.Family.linux => purge_platforms("darwin", "cygwin", "windows") case Platform.Family.macos => purge_platforms("linux", "cygwin", "windows") case Platform.Family.windows => purge_platforms("linux", "darwin") } File.find_files(dir.file, (file: JFile) => file.isDirectory && purge_set(file.getName), include_dirs = true).foreach(Isabelle_System.rm_tree) } /* component directory content */ def settings(dir: Path = Path.current): Path = dir + Path.explode("etc/settings") def components(dir: Path = Path.current): Path = dir + Path.explode("etc/components") def check_dir(dir: Path): Boolean = settings(dir).is_file || components(dir).is_file def read_components(dir: Path): List[String] = split_lines(File.read(components(dir))).filter(_.nonEmpty) def write_components(dir: Path, lines: List[String]): Unit = File.write(components(dir), terminate_lines(lines)) /* component repository content */ val components_sha1: Path = Path.explode("~~/Admin/components/components.sha1") sealed case class SHA1_Digest(sha1: String, file_name: String) { override def toString: String = sha1 + " " + file_name } def read_components_sha1(lines: List[String] = Nil): List[SHA1_Digest] = (proper_list(lines) getOrElse split_lines(File.read(components_sha1))).flatMap(line => Word.explode(line) match { case Nil => None case List(sha1, name) => Some(SHA1_Digest(sha1, name)) case _ => error("Bad components.sha1 entry: " + quote(line)) }) - def write_components_sha1(entries: List[SHA1_Digest]) = + def write_components_sha1(entries: List[SHA1_Digest]): Unit = File.write(components_sha1, entries.sortBy(_.file_name).mkString("", "\n", "\n")) /** build and publish components **/ def build_components( options: Options, components: List[Path], progress: Progress = No_Progress, publish: Boolean = false, force: Boolean = false, update_components_sha1: Boolean = false) { val archives: List[Path] = for (path <- components) yield { path.file_name match { case Archive(_) => path case name => if (!path.is_dir) error("Bad component directory: " + path) else if (!check_dir(path)) { error("Malformed component directory: " + path + "\n (requires " + settings() + " or " + Components.components() + ")") } else { val component_path = path.expand val archive_dir = component_path.dir val archive_name = Archive(name) val archive = archive_dir + Path.explode(archive_name) if (archive.is_file && !force) { error("Component archive already exists: " + archive) } progress.echo("Packaging " + archive_name) Isabelle_System.gnutar("-czf " + File.bash_path(archive) + " " + Bash.string(name), dir = archive_dir).check archive } } } if ((publish && archives.nonEmpty) || update_components_sha1) { options.string("isabelle_components_server") match { case SSH.Target(user, host) => using(SSH.open_session(options, host = host, user = user))(ssh => { val components_dir = Path.explode(options.string("isabelle_components_dir")) val contrib_dir = Path.explode(options.string("isabelle_components_contrib_dir")) for (dir <- List(components_dir, contrib_dir) if !ssh.is_dir(dir)) { error("Bad remote directory: " + dir) } if (publish) { for (archive <- archives) { val archive_name = archive.file_name val name = Archive.get_name(archive_name) val remote_component = components_dir + archive.base val remote_contrib = contrib_dir + Path.explode(name) // component archive if (ssh.is_file(remote_component) && !force) { error("Remote component archive already exists: " + remote_component) } progress.echo("Uploading " + archive_name) ssh.write_file(remote_component, archive) // contrib directory val is_standard_component = Isabelle_System.with_tmp_dir("component")(tmp_dir => { Isabelle_System.gnutar("-xzf " + File.bash_path(archive), dir = tmp_dir).check check_dir(tmp_dir + Path.explode(name)) }) if (is_standard_component) { if (ssh.is_dir(remote_contrib)) { if (force) ssh.rm_tree(remote_contrib) else error("Remote component directory already exists: " + remote_contrib) } progress.echo("Unpacking remote " + archive_name) ssh.execute("tar -C " + ssh.bash_path(contrib_dir) + " -xzf " + ssh.bash_path(remote_component)).check } else { progress.echo_warning("No unpacking of non-standard component: " + archive_name) } } } // remote SHA1 digests if (update_components_sha1) { val lines = for { entry <- ssh.read_dir(components_dir) if entry.is_file && entry.name.endsWith(Archive.suffix) } yield { progress.echo("Digesting remote " + entry.name) Library.trim_line( ssh.execute("cd " + ssh.bash_path(components_dir) + "; sha1sum " + Bash.string(entry.name)).check.out) } write_components_sha1(read_components_sha1(lines)) } }) case s => error("Bad isabelle_components_server: " + quote(s)) } } // local SHA1 digests { val new_entries = for (archive <- archives) yield { val file_name = archive.file_name progress.echo("Digesting local " + file_name) val sha1 = SHA1.digest(archive).rep SHA1_Digest(sha1, file_name) } val new_names = new_entries.map(_.file_name).toSet write_components_sha1( new_entries ::: read_components_sha1().filterNot(entry => new_names.contains(entry.file_name))) } } /* Isabelle tool wrapper */ private val relevant_options = List("isabelle_components_server", "isabelle_components_dir", "isabelle_components_contrib_dir") val isabelle_tool = Isabelle_Tool("build_components", "build and publish Isabelle components", args => { var publish = false var update_components_sha1 = false var force = false var options = Options.init() def show_options: String = cat_lines(relevant_options.map(name => options.options(name).print)) val getopts = Getopts(""" Usage: isabelle build_components [OPTIONS] ARCHIVES... DIRS... Options are: -P publish on SSH server (see options below) -f force: overwrite existing component archives and directories -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) -u update all SHA1 keys in Isabelle repository Admin/components Build and publish Isabelle components as .tar.gz archives on SSH server, depending on system options: """ + Library.prefix_lines(" ", show_options) + "\n", "P" -> (_ => publish = true), "f" -> (_ => force = true), "o:" -> (arg => options = options + arg), "u" -> (_ => update_components_sha1 = true)) val more_args = getopts(args) if (more_args.isEmpty && !update_components_sha1) getopts.usage() val progress = new Console_Progress build_components(options, more_args.map(Path.explode), progress = progress, publish = publish, force = force, update_components_sha1 = update_components_sha1) }) } 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,608 +1,608 @@ /* 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 import scala.collection.mutable object Isabelle_Cronjob { /* global resources: owned by main cronjob */ val backup = "lxbroy10:cronjob" - val main_dir = Path.explode("~/cronjob") - val main_state_file = main_dir + Path.explode("run/main.state") - val current_log = main_dir + Path.explode("run/main.log") // owned by log service - val cumulative_log = main_dir + Path.explode("log/main.log") // owned by log service + val 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 = main_dir + Path.explode("isabelle") - val afp_repos = main_dir + Path.explode("AFP") + val isabelle_repos: Path = main_dir + Path.explode("isabelle") + val afp_repos: Path = main_dir + Path.explode("AFP") 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 = + 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 = + val exit: Logger_Task = Logger_Task("exit", logger => { Isabelle_System.bash( "rsync -a " + File.bash_path(main_dir) + "/log/." + " " + Bash.string(backup) + "/log/.") .check }) /* build release */ - val 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()) }) /* integrity test of build_history vs. build_history_base */ - val build_history_base = + val build_history_base: Logger_Task = Logger_Task("build_history_base", logger => { using(logger.ssh_context.open_session("lxbroy10"))(ssh => { val results = Build_History.remote_build_history(ssh, isabelle_repos, isabelle_repos.ext("build_history_base"), isabelle_identifier = "cronjob_build_history", self_update = true, rev = "build_history_base", options = "-f", args = "HOL") for ((log_name, bytes) <- results) { Bytes.write(logger.log_dir + Path.explode(log_name), bytes) } }) }) /* 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", 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 + " = " + 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 }) 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) }) } } val remote_builds_old: List[Remote_Build] = List( Remote_Build("AFP 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 + " = " + SQL.string("AFP")), Remote_Build("AFP", "lxbroy7", args = "-N -X large -X slow", afp = true, detect = Build_Log.Prop.build_tags + " = " + 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 + " = " + 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 + " = " + 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 Mac OS X", "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 + " = " + SQL.string("polyml-5.7")), Remote_Build("Poly/ML 5.7.1 Mac OS X", "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 + " = " + SQL.string("polyml-5.7.1-pre1") + " OR " + Build_Log.Prop.build_tags + " = " + SQL.string("polyml-5.7.1-pre2")), 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 + " = " + SQL.string("polyml-test")), Remote_Build("Mac OS X 10.8 Mountain Lion", "macbroy30", options = "-m32 -M2", args = "-a", detect = Build_Log.Prop.build_start + " < date '2017-03-03'")) val remote_builds1: List[List[Remote_Build]] = { List( List(Remote_Build("Linux A", "lxbroy9", options = "-m32 -B -M1x2,2", args = "-N -g timing")), List(Remote_Build("Linux B", "lxbroy10", historic = true, history = 90, options = "-m32 -B -M1x4,2,4,6", args = "-N -g timing")), List(Remote_Build("Linux Benchmarks", "lxbroy5", historic = true, history = 90, 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 + " = " + SQL.string("Benchmarks"))), List( Remote_Build("Mac OS X", "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", args = "-a", detect = Build_Log.Prop.build_tags.undefined, history_base = "2c0f24e927dd"), Remote_Build("Mac OS X, quick_and_dirty", "macbroy2", options = "-m32 -M8 -t quick_and_dirty", args = "-a -o quick_and_dirty", detect = Build_Log.Prop.build_tags + " = " + SQL.string("quick_and_dirty"), history_base = "2c0f24e927dd"), Remote_Build("Mac OS X, skip_proofs", "macbroy2", options = "-m32 -M8 -t skip_proofs", args = "-a -o skip_proofs", detect = Build_Log.Prop.build_tags + " = " + SQL.string("skip_proofs"), history_base = "2c0f24e927dd")), List( Remote_Build("Mac OS X 10.12 Sierra", "macbroy30", options = "-m32 -M2", args = "-a", detect = Build_Log.Prop.build_start + " > date '2017-03-03'")), List(Remote_Build("Mac OS X 10.10 Yosemite", "macbroy31", options = "-m32 -M2", args = "-a")), List(Remote_Build("Mac OS X 10.13 High Sierra", "lapbroy68", options = "-m32 -M1,2,4 -e ISABELLE_GHC_SETUP=true", self_update = true, args = "-a -d '~~/src/Benchmarks'")), List(Remote_Build("Mac OS X 10.14 Mojave", "lapnipkow3", options = "-m32 -M1,2 -e ISABELLE_GHC_SETUP=true", self_update = true, args = "-a -d '~~/src/Benchmarks'")), List(Remote_Build("Mac OS X 10.15 Catalina", "laramac01", user = "makarius", proxy_host = "laraserver", proxy_user = "makarius", self_update = true, options = "-m32 -M4 -e ISABELLE_GHC_SETUP=true", args = "-a -d '~~/src/Benchmarks'")), List( Remote_Build("Windows", "vmnipkow9", historic = true, history = 90, self_update = true, options = "-m32 -M4" + " -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 + " = " + SQL.string("x86-windows")), Remote_Build("Windows", "vmnipkow9", historic = true, history = 90, self_update = true, options = "-m64 -M4" + " -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 + " = " + SQL.string("x86_64-windows"))) ) ::: { for { (n, hosts) <- List(1 -> List("lxbroy6"), 2 -> List("lxbroy8", "lxbroy5")) } yield { List(Remote_Build("AFP", 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 + " = " + SQL.string("AFP"))) } } } val remote_builds2: List[List[Remote_Build]] = List( List( Remote_Build("AFP2", "lrzcloud2", actual_host = "10.195.2.10", self_update = true, proxy_host = "lxbroy10", proxy_user = "i21isatest", options = "-m32 -M1x8 -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 + " = " + SQL.string("AFP")), Remote_Build("AFP bulky2", "lrzcloud2", actual_host = "10.195.2.10", self_update = true, proxy_host = "lxbroy10", proxy_user = "i21isatest", options = "-m64 -M8 -U30000 -s10 -t AFP", args = "-g large -g slow", afp = true, bulky = true, detect = Build_Log.Prop.build_tags + " = " + 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) + " -f -h " + Bash.string(r.host) + " " + r.options, args = "-o timeout=10800 " + r.args) for ((log_name, bytes) <- results) { logger.log(Date.now(), log_name) Bytes.write(logger.log_dir + Path.explode(log_name), bytes) } }) }) } val build_status_profiles: List[Build_Status.Profile] = (remote_builds_old :: remote_builds1 ::: remote_builds2).flatten.map(_.profile) /** task logging **/ object Log_Service { def apply(options: Options, progress: Progress = No_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() { thread.shutdown() } - val hostname = Isabelle_System.hostname() + 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) { 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]) { val elapsed_time = end_date.time - start_date.time val msg = (if (err.isEmpty) "finished" else "ERROR " + err.get) + (if (elapsed_time.seconds < 3.0) "" else " (" + elapsed_time.message_hms + " elapsed time)") log(end_date, msg) } val log_dir: Path = main_dir + Build_Log.log_subdir(start_date) Isabelle_System.mkdirs(log_dir) log(start_date, "started") } class Task private[Isabelle_Cronjob](name: String, body: => Unit) { private val future: Future[Unit] = Future.thread("cronjob: " + name) { body } def is_finished: Boolean = future.is_finished } /** cronjob **/ def cronjob(progress: Progress, exclude_task: Set[String]) { /* soft lock */ val still_running = try { Some(File.read(main_state_file)) } catch { case ERROR(_) => None } still_running match { case None | Some("") => case Some(running) => error("Isabelle cronjob appears to be still running: " + running) } /* log service */ val log_service = Log_Service(Options.init(), progress = progress) def run(start_date: Date, task: Logger_Task) { log_service.run_task(start_date, task) } def run_now(task: Logger_Task) { run(Date.now(), task) } /* structured tasks */ def SEQ(tasks: List[Logger_Task]): Logger_Task = Logger_Task(body = _ => for (task <- tasks.iterator if !exclude_task(task.name) || task.name == "") run_now(task)) def PAR(tasks: List[Logger_Task]): Logger_Task = Logger_Task(body = _ => { @tailrec def join(running: List[Task]) { running.partition(_.is_finished) match { case (Nil, Nil) => case (Nil, _ :: _) => Thread.sleep(500); join(running) case (_ :: _, remaining) => join(remaining) } } val start_date = Date.now() val running = for (task <- tasks if !exclude_task(task.name)) yield log_service.fork_task(start_date, task) join(running) }) /* 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 + " " + log_service.hostname) run(main_start_date, Logger_Task("isabelle_cronjob", logger => run_now( SEQ(List( init, build_history_base, 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]) { Command_Line.tool0 { var force = false var verbose = false var exclude_task = Set.empty[String] val getopts = Getopts(""" Usage: Admin/cronjob/main [OPTIONS] Options are: -f apply force to do anything -v verbose -x NAME exclude tasks with this name """, "f" -> (_ => force = true), "v" -> (_ => verbose = true), "x:" -> (arg => exclude_task += arg)) val more_args = getopts(args) if (more_args.nonEmpty) getopts.usage() val progress = if (verbose) new Console_Progress() else No_Progress if (force) cronjob(progress, exclude_task) else error("Need to apply force to do anything") } } } diff --git a/src/Pure/Admin/isabelle_devel.scala b/src/Pure/Admin/isabelle_devel.scala --- a/src/Pure/Admin/isabelle_devel.scala +++ b/src/Pure/Admin/isabelle_devel.scala @@ -1,76 +1,76 @@ /* Title: Pure/Admin/isabelle_devel.scala Author: Makarius Website for Isabelle development resources. */ package isabelle object Isabelle_Devel { val RELEASE_SNAPSHOT = "release_snapshot" val BUILD_LOG_DB = "build_log.db" val BUILD_STATUS = "build_status" val CRONJOB_LOG = "cronjob-main.log" - val root = Path.explode("~/html-data/devel") - val cronjob_log = root + Path.basic(CRONJOB_LOG) + val root: Path = Path.explode("~/html-data/devel") + val cronjob_log: Path = root + Path.basic(CRONJOB_LOG) /* index */ def make_index() { val redirect = "https://isabelle-dev.sketis.net/home/menu/view/20" HTML.write_document(root, "index.html", List( XML.Elem(Markup("meta", List("http-equiv" -> "Refresh", "content" -> ("0; url=" + redirect))), Nil)), List(HTML.link(redirect, HTML.text("Isabelle Development Resources")))) } /* release snapshot */ def release_snapshot( options: Options, rev: String = "", afp_rev: String = "", parallel_jobs: Int = 1) { Isabelle_System.with_tmp_dir("isadist")(base_dir => { Isabelle_System.update_directory(root + Path.explode(RELEASE_SNAPSHOT), website_dir => Build_Release.build_release(base_dir, options, rev = rev, afp_rev = afp_rev, parallel_jobs = parallel_jobs, build_sessions = List(Isabelle_System.getenv("ISABELLE_LOGIC")), website = Some(website_dir))) }) } /* maintain build_log database */ def build_log_database(options: Options, log_dirs: List[Path]) { val store = Build_Log.store(options) using(store.open_database())(db => { store.update_database(db, log_dirs) store.update_database(db, log_dirs, ml_statistics = true) store.snapshot_database(db, root + Path.explode(BUILD_LOG_DB)) }) } /* present build status */ def build_status(options: Options) { Isabelle_System.update_directory(root + Path.explode(BUILD_STATUS), dir => Build_Status.build_status(options, target_dir = dir, ml_statistics = true)) } } diff --git a/src/Pure/Admin/jenkins.scala b/src/Pure/Admin/jenkins.scala --- a/src/Pure/Admin/jenkins.scala +++ b/src/Pure/Admin/jenkins.scala @@ -1,147 +1,147 @@ /* Title: Pure/Admin/jenkins.scala Author: Makarius Support for Jenkins continuous integration service. */ package isabelle import java.net.URL import scala.util.matching.Regex object Jenkins { /* server API */ def root(): String = Isabelle_System.getenv_strict("ISABELLE_JENKINS_ROOT") def invoke(url: String, args: String*): Any = { val req = url + "/api/json?" + args.mkString("&") val result = Url.read(req) try { JSON.parse(result) } catch { case ERROR(_) => error("Malformed JSON from " + quote(req)) } } /* build jobs */ def build_job_names(): List[String] = for { job <- JSON.array(invoke(root()), "jobs").getOrElse(Nil) _class <- JSON.string(job, "_class") if _class == "hudson.model.FreeStyleProject" name <- JSON.string(job, "name") } yield name def download_logs( options: Options, job_names: List[String], dir: Path, progress: Progress = No_Progress) { val store = Sessions.store(options) - val infos = job_names.flatMap(build_job_infos(_)) + val infos = job_names.flatMap(build_job_infos) Par_List.map((info: Job_Info) => info.download_log(store, dir, progress), infos) } /* build log status */ val build_log_jobs = List("isabelle-nightly-benchmark", "isabelle-nightly-slow") val build_status_profiles: List[Build_Status.Profile] = build_log_jobs.map(job_name => Build_Status.Profile("jenkins " + job_name, sql = Build_Log.Prop.build_engine + " = " + SQL.string(Build_Log.Jenkins.engine) + " AND " + Build_Log.Data.session_name + " <> " + SQL.string("Pure") + " AND " + Build_Log.Data.status + " = " + SQL.string(Build_Log.Session_Status.finished.toString) + " AND " + Build_Log.Data.log_name + " LIKE " + SQL.string("%" + job_name))) /* job info */ sealed case class Job_Info( job_name: String, timestamp: Long, main_log: URL, session_logs: List[(String, String, URL)]) { val date: Date = Date(Time.ms(timestamp), Date.timezone_berlin) def log_filename: Path = Build_Log.log_filename(Build_Log.Jenkins.engine, date, List(job_name)) def read_ml_statistics(store: Sessions.Store, session_name: String): List[Properties.T] = { def get_log(ext: String): Option[URL] = session_logs.collectFirst({ case (a, b, url) if a == session_name && b == ext => url }) get_log("db") match { case Some(url) => Isabelle_System.with_tmp_file(session_name, "db") { database => Bytes.write(database, Bytes.read(url)) using(SQLite.open_database(database))(db => store.read_ml_statistics(db, session_name)) } case None => get_log("gz") match { case Some(url) => val log_file = Build_Log.Log_File(session_name, Url.read_gzip(url)) log_file.parse_session_info(ml_statistics = true).ml_statistics case None => Nil } } } def download_log(store: Sessions.Store, dir: Path, progress: Progress = No_Progress) { val log_dir = dir + Build_Log.log_subdir(date) val log_path = log_dir + log_filename.ext("xz") if (!log_path.is_file) { progress.echo(log_path.expand.implode) Isabelle_System.mkdirs(log_dir) val ml_statistics = - session_logs.map(_._1).toSet.toList.sorted.flatMap(session_name => + session_logs.map(_._1).distinct.sorted.flatMap(session_name => read_ml_statistics(store, session_name). map(props => (Build_Log.SESSION_NAME -> session_name) :: props)) File.write_xz(log_path, terminate_lines(Url.read(main_log) :: ml_statistics.map(Build_Log.Log_File.print_props(Build_Log.ML_STATISTICS_MARKER, _))), XZ.options(6)) } } } def build_job_infos(job_name: String): List[Job_Info] = { val Session_Log = new Regex("""^.*/log/([^/]+)\.(db|gz)$""") val infos = for { build <- JSON.array( invoke(root() + "/job/" + job_name, "tree=allBuilds[number,timestamp,artifacts[*]]"), "allBuilds").getOrElse(Nil) number <- JSON.int(build, "number") timestamp <- JSON.long(build, "timestamp") } yield { val job_prefix = root() + "/job/" + job_name + "/" + number val main_log = Url(job_prefix + "/consoleText") val session_logs = for { artifact <- JSON.array(build, "artifacts").getOrElse(Nil) log_path <- JSON.string(artifact, "relativePath") (name, ext) <- (log_path match { case Session_Log(a, b) => Some((a, b)) case _ => None }) } yield (name, ext, Url(job_prefix + "/artifact/" + log_path)) Job_Info(job_name, timestamp, main_log, session_logs) } infos.sortBy(info => - info.timestamp) } } diff --git a/src/Pure/Admin/other_isabelle.scala b/src/Pure/Admin/other_isabelle.scala --- a/src/Pure/Admin/other_isabelle.scala +++ b/src/Pure/Admin/other_isabelle.scala @@ -1,115 +1,115 @@ /* Title: Pure/Admin/other_isabelle.scala Author: Makarius Manage other Isabelle distributions. */ package isabelle object Other_Isabelle { def apply(isabelle_home: Path, isabelle_identifier: String = "", user_home: Path = Path.explode("$USER_HOME"), progress: Progress = No_Progress): Other_Isabelle = new Other_Isabelle(isabelle_home.canonical, isabelle_identifier, user_home, progress) } class Other_Isabelle( val isabelle_home: Path, val isabelle_identifier: String, user_home: Path, progress: Progress) { other_isabelle => override def toString: String = isabelle_home.toString if (proper_string(System.getenv("ISABELLE_SETTINGS_PRESENT")).isDefined) error("Cannot initialize with enclosing ISABELLE_SETTINGS_PRESENT") /* static system */ def bash( script: String, redirect: Boolean = false, echo: Boolean = false, strict: Boolean = true): Process_Result = progress.bash( "export USER_HOME=" + File.bash_path(user_home) + "\n" + Isabelle_System.export_isabelle_identifier(isabelle_identifier) + script, env = null, cwd = isabelle_home.file, redirect = redirect, echo = echo, strict = strict) def apply( cmdline: String, redirect: Boolean = false, echo: Boolean = false, strict: Boolean = true): Process_Result = bash("bin/isabelle " + cmdline, redirect = redirect, echo = echo, strict = strict) def resolve_components(echo: Boolean): Unit = other_isabelle("components -a", redirect = true, echo = echo).check def getenv(name: String): String = other_isabelle("getenv -b " + Bash.string(name)).check.out val isabelle_home_user: Path = Path.explode(getenv("ISABELLE_HOME_USER")) val etc: Path = isabelle_home_user + Path.explode("etc") val etc_settings: Path = etc + Path.explode("settings") val etc_preferences: Path = etc + Path.explode("preferences") def copy_fonts(target_dir: Path): Unit = - Isabelle_Fonts.make_entries(getenv = getenv(_), hidden = true). + Isabelle_Fonts.make_entries(getenv = getenv, hidden = true). foreach(entry => File.copy(entry.path, target_dir)) /* components */ def init_components( components_base: Path = Components.default_components_base, catalogs: List[String] = Nil, components: List[String] = Nil): List[String] = { val dir = Components.admin(isabelle_home) catalogs.map(name => "init_components " + File.bash_path(components_base) + " " + File.bash_path(dir + Path.basic(name))) ::: components.map(name => "init_component " + File.bash_path(components_base + Path.basic(name))) } /* settings */ def clean_settings(): Boolean = if (!etc_settings.is_file) true else if (File.read(etc_settings).startsWith("# generated by Isabelle")) { etc_settings.file.delete; true } else false def init_settings(settings: List[String]) { if (!clean_settings()) error("Cannot proceed with existing user settings file: " + etc_settings) Isabelle_System.mkdirs(etc_settings.dir) File.write(etc_settings, "# generated by Isabelle " + Date.now() + "\n" + "#-*- shell-script -*- :mode=shellscript:\n" + settings.mkString("\n", "\n", "\n")) } /* cleanup */ def cleanup() { clean_settings() etc.file.delete isabelle_home_user.file.delete } } diff --git a/src/Pure/General/csv.scala b/src/Pure/General/csv.scala --- a/src/Pure/General/csv.scala +++ b/src/Pure/General/csv.scala @@ -1,33 +1,33 @@ /* Title: Pure/General/csv.scala Author: Makarius Support for CSV: RFC 4180. */ package isabelle object CSV { def print_field(field: Any): String = { val str = field.toString if (str.exists("\",\r\n".contains(_))) { (for (c <- str) yield { if (c == '"') "\"\"" else c.toString }).mkString("\"", "", "\"") } else str } sealed case class Record(fields: Any*) { - override def toString: String = fields.iterator.map(print_field(_)).mkString(",") + override def toString: String = fields.iterator.map(print_field).mkString(",") } sealed case class File(name: String, header: List[String], records: List[Record]) { override def toString: String = (Record(header:_*) :: records).mkString("\r\n") def file_name: String = name + ".csv" def write(dir: Path) { isabelle.File.write(dir + Path.explode(file_name), toString) } } } diff --git a/src/Pure/General/date.scala b/src/Pure/General/date.scala --- a/src/Pure/General/date.scala +++ b/src/Pure/General/date.scala @@ -1,105 +1,105 @@ /* Title: Pure/General/date.scala Author: Makarius Date and time, with time zone. */ package isabelle import java.util.Locale import java.time.{Instant, ZonedDateTime, LocalTime, ZoneId, OffsetDateTime} import java.time.format.{DateTimeFormatter, DateTimeParseException} import java.time.temporal.TemporalAccessor import scala.annotation.tailrec object Date { /* format */ object Format { def make(fmts: List[DateTimeFormatter], tune: String => String = s => s): Format = { require(fmts.nonEmpty) new Format { def apply(date: Date): String = fmts.head.format(date.rep) def parse(str: String): Date = new Date(ZonedDateTime.from(Formatter.try_variants(fmts, tune(str)))) } } def apply(pats: String*): Format = - make(pats.toList.map(Date.Formatter.pattern(_))) + make(pats.toList.map(Date.Formatter.pattern)) val default: Format = Format("dd-MMM-uuuu HH:mm:ss xx") val date: Format = Format("dd-MMM-uuuu") val time: Format = Format("HH:mm:ss") } abstract class Format private { def apply(date: Date): String def parse(str: String): Date def unapply(str: String): Option[Date] = try { Some(parse(str)) } catch { case _: DateTimeParseException => None } } object Formatter { def pattern(pat: String): DateTimeFormatter = DateTimeFormatter.ofPattern(pat) def variants(pats: List[String], locs: List[Locale] = Nil): List[DateTimeFormatter] = pats.flatMap(pat => { val fmt = pattern(pat) - if (locs.isEmpty) List(fmt) else locs.map(fmt.withLocale(_)) + if (locs.isEmpty) List(fmt) else locs.map(fmt.withLocale) }) @tailrec def try_variants(fmts: List[DateTimeFormatter], str: String, last_exn: Option[DateTimeParseException] = None): TemporalAccessor = { fmts match { case Nil => throw last_exn.getOrElse(new DateTimeParseException("Failed to parse date", str, 0)) case fmt :: rest => try { ZonedDateTime.from(fmt.parse(str)) } catch { case exn: DateTimeParseException => try_variants(rest, str, Some(exn)) } } } } /* date operations */ def timezone_utc: ZoneId = ZoneId.of("UTC") def timezone_berlin: ZoneId = ZoneId.of("Europe/Berlin") def timezone(): ZoneId = ZoneId.systemDefault def now(zone: ZoneId = timezone()): Date = new Date(ZonedDateTime.now(zone)) def instant(t: Instant, zone: ZoneId = timezone()): Date = new Date(ZonedDateTime.ofInstant(t, zone)) def apply(t: Time, zone: ZoneId = timezone()): Date = instant(t.instant, zone) } sealed case class Date(rep: ZonedDateTime) { def midnight: Date = new Date(ZonedDateTime.of(rep.toLocalDate, LocalTime.MIDNIGHT, rep.getZone)) def to(zone: ZoneId): Date = new Date(rep.withZoneSameInstant(zone)) def unix_epoch: Long = rep.toEpochSecond def instant: Instant = Instant.from(rep) def time: Time = Time.instant(instant) def timezone: ZoneId = rep.getZone def format(fmt: Date.Format = Date.Format.default): String = fmt(this) override def toString: String = format() } diff --git a/src/Pure/General/exn.scala b/src/Pure/General/exn.scala --- a/src/Pure/General/exn.scala +++ b/src/Pure/General/exn.scala @@ -1,153 +1,153 @@ /* Title: Pure/General/exn.scala Author: Makarius Support for exceptions (arbitrary throwables). */ package isabelle object Exn { /* user errors */ class User_Error(message: String) extends RuntimeException(message) { override def equals(that: Any): Boolean = that match { case other: User_Error => message == other.getMessage case _ => false } override def hashCode: Int = message.hashCode override def toString: String = "\n" + Output.error_message_text(message) } object ERROR { def apply(message: String): User_Error = new User_Error(message) def unapply(exn: Throwable): Option[String] = user_message(exn) } def error(message: String): Nothing = throw ERROR(message) def cat_message(msgs: String*): String = cat_lines(msgs.iterator.filterNot(_ == "")) def cat_error(msgs: String*): Nothing = error(cat_message(msgs:_*)) /* exceptions as values */ sealed abstract class Result[A] { def user_error: Result[A] = this match { case Exn(ERROR(msg)) => Exn(ERROR(msg)) case _ => this } def map[B](f: A => B): Result[B] = this match { case Res(res) => Res(f(res)) case Exn(exn) => Exn(exn) } } case class Res[A](res: A) extends Result[A] case class Exn[A](exn: Throwable) extends Result[A] def capture[A](e: => A): Result[A] = try { Res(e) } catch { case exn: Throwable => Exn[A](exn) } def release[A](result: Result[A]): A = result match { case Res(x) => x case Exn(exn) => throw exn } def release_first[A](results: List[Result[A]]): List[A] = results.find({ case Exn(exn) => !is_interrupt(exn) case _ => false }) match { case Some(Exn(exn)) => throw exn - case _ => results.map(release(_)) + case _ => results.map(release) } /* interrupts */ def is_interrupt(exn: Throwable): Boolean = { var found_interrupt = false var e = exn while (!found_interrupt && e != null) { found_interrupt |= e.isInstanceOf[InterruptedException] e = e.getCause } found_interrupt } def interruptible_capture[A](e: => A): Result[A] = try { Res(e) } catch { case exn: Throwable if !is_interrupt(exn) => Exn[A](exn) } object Interrupt { def apply(): Throwable = new InterruptedException def unapply(exn: Throwable): Boolean = is_interrupt(exn) def expose() { if (Thread.interrupted) throw apply() } def impose() { Thread.currentThread.interrupt } def postpone[A](body: => A): Option[A] = { val interrupted = Thread.interrupted val result = capture { body } if (interrupted) impose() result match { case Res(x) => Some(x) case Exn(e) => if (is_interrupt(e)) { impose(); None } else throw e } } val return_code = 130 } /* POSIX return code */ def return_code(exn: Throwable, rc: Int): Int = if (is_interrupt(exn)) Interrupt.return_code else rc /* message */ def user_message(exn: Throwable): Option[String] = if (exn.getClass == classOf[RuntimeException] || exn.getClass == classOf[User_Error]) { Some(proper_string(exn.getMessage) getOrElse "Error") } else if (exn.isInstanceOf[java.sql.SQLException]) { Some(proper_string(exn.getMessage) getOrElse "SQL error") } else if (exn.isInstanceOf[java.io.IOException]) { val msg = exn.getMessage Some(if (msg == null || msg == "") "I/O error" else "I/O error: " + msg) } else if (exn.isInstanceOf[RuntimeException]) Some(exn.toString) else None def message(exn: Throwable): String = user_message(exn) getOrElse (if (is_interrupt(exn)) "Interrupt" else exn.toString) /* trace */ def trace(exn: Throwable): String = exn.getStackTrace.mkString("\n") } diff --git a/src/Pure/General/file.scala b/src/Pure/General/file.scala --- a/src/Pure/General/file.scala +++ b/src/Pure/General/file.scala @@ -1,358 +1,358 @@ /* Title: Pure/General/file.scala Author: Makarius File-system operations. */ package isabelle import java.io.{BufferedWriter, OutputStreamWriter, FileOutputStream, BufferedOutputStream, OutputStream, InputStream, FileInputStream, BufferedInputStream, BufferedReader, InputStreamReader, File => JFile, IOException} import java.nio.file.{StandardOpenOption, StandardCopyOption, Path => JPath, Files, SimpleFileVisitor, FileVisitOption, FileVisitResult, FileSystemException} import java.nio.file.attribute.BasicFileAttributes import java.net.{URL, MalformedURLException} import java.util.zip.{GZIPInputStream, GZIPOutputStream} import java.util.regex.Pattern import java.util.EnumSet import org.tukaani.xz.{XZInputStream, XZOutputStream} import scala.collection.mutable import scala.util.matching.Regex object File { /* standard path (Cygwin or Posix) */ def standard_path(path: Path): String = path.expand.implode def standard_path(platform_path: String): String = if (Platform.is_windows) { val Platform_Root = new Regex("(?i)" + Pattern.quote(Isabelle_System.cygwin_root()) + """(?:\\+|\z)(.*)""") val Drive = new Regex("""([a-zA-Z]):\\*(.*)""") platform_path.replace('/', '\\') match { case Platform_Root(rest) => "/" + rest.replace('\\', '/') case Drive(letter, rest) => "/cygdrive/" + Word.lowercase(letter) + (if (rest == "") "" else "/" + rest.replace('\\', '/')) case path => path.replace('\\', '/') } } else platform_path def standard_path(file: JFile): String = standard_path(file.getPath) def standard_url(name: String): String = try { val url = new URL(name) if (url.getProtocol == "file" && Url.is_wellformed_file(name)) standard_path(Url.parse_file(name)) else name } catch { case _: MalformedURLException => standard_path(name) } /* platform path (Windows or Posix) */ private val Cygdrive = new Regex("/cygdrive/([a-zA-Z])($|/.*)") private val Named_Root = new Regex("//+([^/]*)(.*)") def platform_path(standard_path: String): String = if (Platform.is_windows) { val result_path = new StringBuilder val rest = standard_path match { case Cygdrive(drive, rest) => result_path ++= (Word.uppercase(drive) + ":" + JFile.separator) rest case Named_Root(root, rest) => result_path ++= JFile.separator result_path ++= JFile.separator result_path ++= root rest case path if path.startsWith("/") => result_path ++= Isabelle_System.cygwin_root() path case path => path } for (p <- space_explode('/', rest) if p != "") { val len = result_path.length if (len > 0 && result_path(len - 1) != JFile.separatorChar) result_path += JFile.separatorChar result_path ++= p } result_path.toString } else standard_path def platform_path(path: Path): String = platform_path(standard_path(path)) def platform_file(path: Path): JFile = new JFile(platform_path(path)) /* platform files */ def absolute(file: JFile): JFile = file.toPath.toAbsolutePath.normalize.toFile def absolute_name(file: JFile): String = absolute(file).getPath def canonical(file: JFile): JFile = file.getCanonicalFile def canonical_name(file: JFile): String = canonical(file).getPath def path(file: JFile): Path = Path.explode(standard_path(file)) def pwd(): Path = path(Path.current.absolute_file) /* relative paths */ def relative_path(base: Path, other: Path): Option[Path] = { val base_path = base.file.toPath val other_path = other.file.toPath if (other_path.startsWith(base_path)) Some(path(base_path.relativize(other_path).toFile)) else None } /* bash path */ def bash_path(path: Path): String = Bash.string(standard_path(path)) def bash_path(file: JFile): String = Bash.string(standard_path(file)) /* directory entries */ def check_dir(path: Path): Path = if (path.is_dir) path else error("No such directory: " + path) def check_file(path: Path): Path = if (path.is_file) path else error("No such file: " + path) /* directory content */ def read_dir(dir: Path): List[String] = { if (!dir.is_dir) error("No such directory: " + dir.toString) val files = dir.file.listFiles if (files == null) Nil else files.toList.map(_.getName).sorted } def find_files( start: JFile, pred: JFile => Boolean = _ => true, include_dirs: Boolean = false, follow_links: Boolean = false): List[JFile] = { val result = new mutable.ListBuffer[JFile] def check(file: JFile) { if (pred(file)) result += file } if (start.isFile) check(start) else if (start.isDirectory) { val options = if (follow_links) EnumSet.of(FileVisitOption.FOLLOW_LINKS) else EnumSet.noneOf(classOf[FileVisitOption]) Files.walkFileTree(start.toPath, options, Integer.MAX_VALUE, new SimpleFileVisitor[JPath] { override def preVisitDirectory(path: JPath, attrs: BasicFileAttributes): FileVisitResult = { if (include_dirs) check(path.toFile) FileVisitResult.CONTINUE } override def visitFile(path: JPath, attrs: BasicFileAttributes): FileVisitResult = { val file = path.toFile if (include_dirs || !file.isDirectory) check(file) FileVisitResult.CONTINUE } } ) } result.toList } /* read */ def read(file: JFile): String = Bytes.read(file).text def read(path: Path): String = read(path.file) def read_stream(reader: BufferedReader): String = { val output = new StringBuilder(100) var c = -1 while ({ c = reader.read; c != -1 }) output += c.toChar reader.close output.toString } def read_stream(stream: InputStream): String = read_stream(new BufferedReader(new InputStreamReader(stream, UTF8.charset))) def read_gzip(file: JFile): String = read_stream(new GZIPInputStream(new BufferedInputStream(new FileInputStream(file)))) def read_gzip(path: Path): String = read_gzip(path.file) def read_xz(file: JFile): String = read_stream(new XZInputStream(new BufferedInputStream(new FileInputStream(file)))) def read_xz(path: Path): String = read_xz(path.file) /* read lines */ def read_line(reader: BufferedReader): Option[String] = { val line = try { reader.readLine} catch { case _: IOException => null } - if (line == null) None else Some(line) + Option(line) } def read_lines(reader: BufferedReader, progress: String => Unit): List[String] = { val result = new mutable.ListBuffer[String] var line: Option[String] = None while ({ line = read_line(reader); line.isDefined }) { progress(line.get) result += line.get } reader.close result.toList } /* write */ def writer(file: JFile): BufferedWriter = new BufferedWriter(new OutputStreamWriter(new FileOutputStream(file), UTF8.charset)) def write_file(file: JFile, text: CharSequence, make_stream: OutputStream => OutputStream) { val stream = make_stream(new FileOutputStream(file)) using(new BufferedWriter(new OutputStreamWriter(stream, UTF8.charset)))(_.append(text)) } def write(file: JFile, text: CharSequence): Unit = write_file(file, text, s => s) def write(path: Path, text: CharSequence): Unit = write(path.file, text) def write_gzip(file: JFile, text: CharSequence): Unit = write_file(file, text, (s: OutputStream) => new GZIPOutputStream(new BufferedOutputStream(s))) def write_gzip(path: Path, text: CharSequence): Unit = write_gzip(path.file, text) def write_xz(file: JFile, text: CharSequence, options: XZ.Options): Unit = File.write_file(file, text, s => new XZOutputStream(new BufferedOutputStream(s), options)) def write_xz(file: JFile, text: CharSequence): Unit = write_xz(file, text, XZ.options()) def write_xz(path: Path, text: CharSequence, options: XZ.Options): Unit = write_xz(path.file, text, options) def write_xz(path: Path, text: CharSequence): Unit = write_xz(path, text, XZ.options()) def write_backup(path: Path, text: CharSequence) { if (path.is_file) move(path, path.backup) write(path, text) } def write_backup2(path: Path, text: CharSequence) { if (path.is_file) move(path, path.backup2) write(path, text) } /* append */ def append(file: JFile, text: CharSequence): Unit = Files.write(file.toPath, UTF8.bytes(text.toString), StandardOpenOption.APPEND, StandardOpenOption.CREATE) def append(path: Path, text: CharSequence): Unit = append(path.file, text) /* eq */ def eq(file1: JFile, file2: JFile): Boolean = try { java.nio.file.Files.isSameFile(file1.toPath, file2.toPath) } catch { case ERROR(_) => false } def eq(path1: Path, path2: Path): Boolean = eq(path1.file, path2.file) /* eq_content */ def eq_content(file1: JFile, file2: JFile): Boolean = if (eq(file1, file2)) true else if (file1.length != file2.length) false else Bytes.read(file1) == Bytes.read(file2) def eq_content(path1: Path, path2: Path): Boolean = eq_content(path1.file, path2.file) /* copy */ def copy(src: JFile, dst: JFile) { val target = if (dst.isDirectory) new JFile(dst, src.getName) else dst if (!eq(src, target)) Files.copy(src.toPath, target.toPath, StandardCopyOption.COPY_ATTRIBUTES, StandardCopyOption.REPLACE_EXISTING) } def copy(path1: Path, path2: Path): Unit = copy(path1.file, path2.file) /* move */ def move(src: JFile, dst: JFile) { val target = if (dst.isDirectory) new JFile(dst, src.getName) else dst if (!eq(src, target)) Files.move(src.toPath, target.toPath, StandardCopyOption.REPLACE_EXISTING) } def move(path1: Path, path2: Path): Unit = move(path1.file, path2.file) /* symbolic link */ def link(src: Path, dst: Path, force: Boolean = false) { val src_file = src.file val dst_file = dst.file val target = if (dst_file.isDirectory) new JFile(dst_file, src_file.getName) else dst_file if (force) target.delete try { Files.createSymbolicLink(target.toPath, src_file.toPath) } catch { case _: UnsupportedOperationException if Platform.is_windows => Cygwin.link(standard_path(src), target) case _: FileSystemException if Platform.is_windows => Cygwin.link(standard_path(src), target) } } /* permissions */ def is_executable(path: Path): Boolean = { if (Platform.is_windows) Isabelle_System.bash("test -x " + bash_path(path)).check.ok else path.file.canExecute } def set_executable(path: Path, flag: Boolean) { if (Platform.is_windows && flag) Isabelle_System.chmod("a+x", path) else if (Platform.is_windows) Isabelle_System.chmod("a-x", path) else path.file.setExecutable(flag, false) } } 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,330 +1,330 @@ /* 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) } 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) }) } 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)) } def node_height(count: Key => Long): Map[Key, Long] = - reachable_length(count, imm_preds(_), maximals) + reachable_length(count, imm_preds, maximals) def node_depth(count: Key => Long): Map[Key, Long] = - reachable_length(count, imm_succs(_), minimals) + 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 ((n + count(x), rs + x) /: next(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)*/ def reachable(next: Key => Keys, xs: List[Key]): (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) = (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) } /*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(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 } def maximals: List[Key] = (List.empty[Key] /: rep) { 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))) } def restrict(pred: Key => Boolean): Graph[Key, A] = (this /: iterator){ 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)) 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_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)) def topological_order: List[Key] = all_succs(minimals) } diff --git a/src/Pure/General/json.scala b/src/Pure/General/json.scala --- a/src/Pure/General/json.scala +++ b/src/Pure/General/json.scala @@ -1,413 +1,414 @@ /* Title: Pure/General/json.scala Author: Makarius Support for JSON: https://www.json.org/. See also http://seriot.ch/parsing_json.php "Parsing JSON is a Minefield". */ package isabelle +import scala.util.matching.Regex import scala.util.parsing.combinator.Parsers import scala.util.parsing.combinator.lexical.Scanners import scala.util.parsing.input.CharArrayReader.EofCh object JSON { type T = Any type S = String object Object { type Entry = (String, JSON.T) type T = Map[String, JSON.T] val empty: Object.T = Map.empty def apply(entries: Entry*): Object.T = Map(entries:_*) def unapply(obj: T): Option[Object.T] = obj match { case m: Map[_, _] if m.keySet.forall(_.isInstanceOf[String]) => Some(m.asInstanceOf[Object.T]) case _ => None } } /* lexer */ object Kind extends Enumeration { val KEYWORD, STRING, NUMBER, ERROR = Value } sealed case class Token(kind: Kind.Value, text: String) { def is_keyword: Boolean = kind == Kind.KEYWORD def is_keyword(name: String): Boolean = kind == Kind.KEYWORD && text == name def is_string: Boolean = kind == Kind.STRING def is_number: Boolean = kind == Kind.NUMBER def is_error: Boolean = kind == Kind.ERROR } object Lexer extends Scanners with Scan.Parsers { override type Elem = Char type Token = JSON.Token def errorToken(msg: String): Token = Token(Kind.ERROR, msg) val white_space: String = " \t\n\r" - override val whiteSpace = ("[" + white_space + "]+").r + override val whiteSpace: Regex = ("[" + white_space + "]+").r def whitespace: Parser[Any] = many(character(white_space.contains(_))) - val letter: Parser[String] = one(character(Symbol.is_ascii_letter(_))) - val letters1: Parser[String] = many1(character(Symbol.is_ascii_letter(_))) + val letter: Parser[String] = one(character(Symbol.is_ascii_letter)) + val letters1: Parser[String] = many1(character(Symbol.is_ascii_letter)) - def digits: Parser[String] = many(character(Symbol.is_ascii_digit(_))) - def digits1: Parser[String] = many1(character(Symbol.is_ascii_digit(_))) + def digits: Parser[String] = many(character(Symbol.is_ascii_digit)) + def digits1: Parser[String] = many1(character(Symbol.is_ascii_digit)) /* keyword */ def keyword: Parser[Token] = (letters1 | one(character("{}[],:".contains(_)))) ^^ (s => Token(Kind.KEYWORD, s)) /* string */ def string: Parser[Token] = '\"' ~> rep(string_body) <~ '\"' ^^ (cs => Token(Kind.STRING, cs.mkString)) def string_body: Parser[Char] = elem("", c => c > '\u001f' && c != '\"' && c != '\\' && c != EofCh) | '\\' ~> string_escape def string_escape: Parser[Char] = elem("", "\"\\/".contains(_)) | elem("", "bfnrt".contains(_)) ^^ { case 'b' => '\b' case 'f' => '\f' case 'n' => '\n' case 'r' => '\r' case 't' => '\t' } | 'u' ~> repeated(character("0123456789abcdefABCDEF".contains(_)), 4, 4) ^^ (s => Integer.parseInt(s, 16).toChar) def string_failure: Parser[Token] = '\"' ~> failure("Unterminated string") /* number */ def number: Parser[Token] = opt("-") ~ number_body ~ opt(letter) ^^ { case a ~ b ~ None => Token(Kind.NUMBER, a.getOrElse("") + b) case a ~ b ~ Some(c) => errorToken("Invalid number format: " + quote(a.getOrElse("") + b + c)) } def number_body: Parser[String] = (zero | positive) ~ opt(number_fract) ~ opt(number_exp) ^^ { case a ~ b ~ c => a + b.getOrElse("") + c.getOrElse("") } def number_fract: Parser[String] = "." ~ digits1 ^^ { case a ~ b => a + b } def number_exp: Parser[String] = one(character("eE".contains(_))) ~ maybe(character("-+".contains(_))) ~ digits1 ^^ { case a ~ b ~ c => a + b + c } - def zero = one(character(c => c == '0')) - def nonzero = one(character(c => c != '0' && Symbol.is_ascii_digit(c))) + def zero: Parser[String] = one(character(c => c == '0')) + def nonzero: Parser[String] = one(character(c => c != '0' && Symbol.is_ascii_digit(c))) def positive: Parser[String] = nonzero ~ digits ^^ { case a ~ b => a + b } /* token */ def token: Parser[Token] = keyword | (string | (string_failure | (number | failure("Illegal character")))) } /* parser */ trait Parser extends Parsers { type Elem = Token def $$$(name: String): Parser[Token] = elem(name, _.is_keyword(name)) def string: Parser[String] = elem("string", _.is_string) ^^ (_.text) def number: Parser[Double] = elem("number", _.is_number) ^^ (tok => tok.text.toDouble) def json_object: Parser[Object.T] = $$$("{") ~> repsep(string ~ ($$$(":") ~> json_value) ^^ { case a ~ b => (a, b) }, $$$(",")) <~ $$$("}") ^^ (_.toMap) def json_array: Parser[List[T]] = $$$("[") ~> repsep(json_value, $$$(",")) <~ $$$("]") def json_value: Parser[T] = json_object | (json_array | (number | (string | ($$$("true") ^^^ true | ($$$("false") ^^^ false | ($$$("null") ^^^ null)))))) def parse(input: CharSequence, strict: Boolean): T = { val scanner = new Lexer.Scanner(Scan.char_reader(input)) phrase(if (strict) json_object | json_array else json_value)(scanner) match { case Success(json, _) => json case NoSuccess(_, next) => error("Malformed JSON input at " + next.pos) } } } object Parser extends Parser /* standard format */ def parse(s: S, strict: Boolean = true): T = Parser.parse(s, strict) object Format { def unapply(s: S): Option[T] = try { Some(parse(s, strict = false)) } catch { case ERROR(_) => None } def apply(json: T): S = { val result = new StringBuilder def string(s: String) { result += '"' result ++= s.iterator.map { case '"' => "\\\"" case '\\' => "\\\\" case '\b' => "\\b" case '\f' => "\\f" case '\n' => "\\n" case '\r' => "\\r" case '\t' => "\\t" case c => if (c <= '\u001f' || c >= '\u007f' && c <= '\u009f') "\\u%04x".format(c.toInt) else c }.mkString result += '"' } def array(list: List[T]) { result += '[' Library.separate(None, list.map(Some(_))).foreach({ case None => result += ',' case Some(x) => json_format(x) }) result += ']' } def object_(obj: Object.T) { result += '{' Library.separate(None, obj.toList.map(Some(_))).foreach({ case None => result += ',' case Some((x, y)) => string(x) result += ':' json_format(y) }) result += '}' } def json_format(x: T) { x match { case null => result ++= "null" case _: Int | _: Long | _: Boolean => result ++= x.toString case n: Double => val i = n.toLong result ++= (if (i.toDouble == n) i.toString else n.toString) case s: String => string(s) case Object(m) => object_(m) case list: List[T] => array(list) case _ => error("Bad JSON value: " + x.toString) } } json_format(json) result.toString } } /* typed values */ object Value { object UUID { def unapply(json: T): Option[isabelle.UUID.T] = json match { case x: java.lang.String => isabelle.UUID.unapply(x) case _ => None } } object String { def unapply(json: T): Option[java.lang.String] = json match { case x: java.lang.String => Some(x) case _ => None } } object String0 { def unapply(json: T): Option[java.lang.String] = json match { case null => Some("") case x: java.lang.String => Some(x) case _ => None } } object Double { def unapply(json: T): Option[scala.Double] = json match { case x: scala.Double => Some(x) case x: scala.Long => Some(x.toDouble) case x: scala.Int => Some(x.toDouble) case _ => None } } object Long { def unapply(json: T): Option[scala.Long] = json match { case x: scala.Double if x.toLong.toDouble == x => Some(x.toLong) case x: scala.Long => Some(x) case x: scala.Int => Some(x.toLong) case _ => None } } object Int { def unapply(json: T): Option[scala.Int] = json match { case x: scala.Double if x.toInt.toDouble == x => Some(x.toInt) case x: scala.Long if x.toInt.toLong == x => Some(x.toInt) case x: scala.Int => Some(x) case _ => None } } object Boolean { def unapply(json: T): Option[scala.Boolean] = json match { case x: scala.Boolean => Some(x) case _ => None } } object List { def unapply[A](json: T, unapply: T => Option[A]): Option[List[A]] = json match { case xs: List[T] => val ys = xs.map(unapply) if (ys.forall(_.isDefined)) Some(ys.map(_.get)) else None case _ => None } } object Seconds { def unapply(json: T): Option[Time] = - Double.unapply(json).map(Time.seconds(_)) + Double.unapply(json).map(Time.seconds) } } /* object values */ def optional(entry: (String, Option[T])): Object.T = entry match { case (name, Some(x)) => Object(name -> x) case (_, None) => Object.empty } def value(obj: T, name: String): Option[T] = obj match { case Object(m) => m.get(name) case _ => None } def value[A](obj: T, name: String, unapply: T => Option[A]): Option[A] = for { json <- value(obj, name) x <- unapply(json) } yield x def array(obj: T, name: String): Option[List[T]] = value(obj, name) match { case Some(a: List[T]) => Some(a) case _ => None } def value_default[A](obj: T, name: String, unapply: T => Option[A], default: => A): Option[A] = value(obj, name) match { case None => Some(default) case Some(json) => unapply(json) } def uuid(obj: T, name: String): Option[UUID.T] = value(obj, name, Value.UUID.unapply) def string(obj: T, name: String): Option[String] = value(obj, name, Value.String.unapply) def string_default(obj: T, name: String, default: => String = ""): Option[String] = value_default(obj, name, Value.String.unapply, default) def string0(obj: T, name: String): Option[String] = value(obj, name, Value.String0.unapply) def string0_default(obj: T, name: String, default: => String = ""): Option[String] = value_default(obj, name, Value.String0.unapply, default) def double(obj: T, name: String): Option[Double] = value(obj, name, Value.Double.unapply) def double_default(obj: T, name: String, default: => Double = 0.0): Option[Double] = value_default(obj, name, Value.Double.unapply, default) def long(obj: T, name: String): Option[Long] = value(obj, name, Value.Long.unapply) def long_default(obj: T, name: String, default: => Long = 0): Option[Long] = value_default(obj, name, Value.Long.unapply, default) def int(obj: T, name: String): Option[Int] = value(obj, name, Value.Int.unapply) def int_default(obj: T, name: String, default: => Int = 0): Option[Int] = value_default(obj, name, Value.Int.unapply, default) def bool(obj: T, name: String): Option[Boolean] = value(obj, name, Value.Boolean.unapply) def bool_default(obj: T, name: String, default: => Boolean = false): Option[Boolean] = value_default(obj, name, Value.Boolean.unapply, default) def list[A](obj: T, name: String, unapply: T => Option[A]): Option[List[A]] = value(obj, name, Value.List.unapply(_, unapply)) def list_default[A](obj: T, name: String, unapply: T => Option[A], default: => List[A] = Nil) : Option[List[A]] = value_default(obj, name, Value.List.unapply(_, unapply), default) def strings(obj: T, name: String): Option[List[String]] = - list(obj, name, JSON.Value.String.unapply _) + list(obj, name, JSON.Value.String.unapply) def strings_default(obj: T, name: String, default: => List[String] = Nil): Option[List[String]] = - list_default(obj, name, JSON.Value.String.unapply _, default) + list_default(obj, name, JSON.Value.String.unapply, default) def seconds(obj: T, name: String): Option[Time] = value(obj, name, Value.Seconds.unapply) def seconds_default(obj: T, name: String, default: => Time = Time.zero): Option[Time] = value_default(obj, name, Value.Seconds.unapply, default) } 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,159 +1,159 @@ /* 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.SetLike import scala.collection.generic.{SetFactory, CanBuildFrom, GenericSetTemplate, GenericCompanion} import scala.collection.mutable.{Builder, SetBuilder} import scala.language.higherKinds object Linear_Set extends SetFactory[Linear_Set] { private val empty_val: Linear_Set[Nothing] = new Linear_Set[Nothing](None, None, Map(), Map()) - override def empty[A] = empty_val.asInstanceOf[Linear_Set[A]] + override def empty[A]: Linear_Set[A] = empty_val.asInstanceOf[Linear_Set[A]] implicit def canBuildFrom[A]: CanBuildFrom[Coll, A, Linear_Set[A]] = setCanBuildFrom[A] def newBuilder[A]: Builder[A, Linear_Set[A]] = new SetBuilder[A, Linear_Set[A]](empty[A]) 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 scala.collection.immutable.Set[A] with GenericSetTemplate[A, Linear_Set] with SetLike[A, Linear_Set[A]] { override def companion: GenericCompanion[Linear_Set] = 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: Traversable[A]): Linear_Set[A] = // FIXME reverse fold ((hook, this) /: elems) { 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 stringPrefix = "Linear_Set" 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 + (elem: A): Linear_Set[A] = insert_after(end, elem) def - (elem: A): Linear_Set[A] = delete_after(prev(elem)) } 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,335 +1,335 @@ /* 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] = { - def find(root: Path): 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.mkdirs(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(" ")) + 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 = "") { 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 = "") { 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)) { 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) }) 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]) { 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) { 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 = No_Progress) { /* local repository */ Isabelle_System.mkdirs(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 ...") Thread.sleep(500) 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", 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,83 +1,83 @@ /* Title: Pure/General/multi_map.scala Author: Makarius Maps with multiple entries per key. */ package isabelle import scala.collection.GenTraversableOnce import scala.collection.generic.{ImmutableMapFactory, CanBuildFrom} object Multi_Map extends ImmutableMapFactory[Multi_Map] { private val empty_val: Multi_Map[Any, Nothing] = new Multi_Map[Any, Nothing](Map.empty) - override def empty[A, B] = empty_val.asInstanceOf[Multi_Map[A, B]] + override def empty[A, B]: Multi_Map[A, B] = empty_val.asInstanceOf[Multi_Map[A, B]] implicit def canBuildFrom[A, B]: CanBuildFrom[Coll, (A, B), Multi_Map[A, B]] = new MapCanBuildFrom[A, B] } final class Multi_Map[A, +B] private(protected val rep: Map[A, List[B]]) extends scala.collection.immutable.Map[A, B] with scala.collection.immutable.MapLike[A, B, Multi_Map[A, B]] { /* 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) { case (m1, (a, bs)) => (bs :\ m1) { case (b, m2) => m2.insert(a, b) } } /* Map operations */ override def stringPrefix = "Multi_Map" - override def empty = Multi_Map.empty + 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 def + [B1 >: B](p: (A, B1)): Multi_Map[A, B1] = insert(p._1, p._2) override def ++ [B1 >: B](entries: GenTraversableOnce[(A, B1)]): Multi_Map[A, B1] = (this.asInstanceOf[Multi_Map[A, B1]] /: entries)(_ + _) def - (a: A): Multi_Map[A, B] = if (rep.isDefinedAt(a)) new Multi_Map(rep - a) else this } 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,260 +1,260 @@ /* 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 => ".." } /* 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 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]) { val table = (Multi_Map.empty[String, String] /: paths)({ 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(private val elems: List[Path.Elem]) // reversed elements { 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 +(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 backup: Path = { val (prfx, s) = split_path prfx + Path.basic(s + "~") } def backup2: Path = { val (prfx, s) = split_path prfx + Path.basic(s + "~~") } 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 /* 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 /* source position */ def position: Position.T = Position.File(implode) /* 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,187 +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) - def body_length(prts: List[Tree], len: Double): Double = + @tailrec def body_length(prts: List[Tree], len: Double): Double = { val (line, rest) = Library.take_prefix[Tree]({ case Break(true, _, _) => false case _ => true }, prts) val len1 = ((0.0 /: line) { case (l, t) => l + t.length }) max len rest 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_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 = 76.0 - val default_breakgain = default_margin / 20 + 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/properties.scala b/src/Pure/General/properties.scala --- a/src/Pure/General/properties.scala +++ b/src/Pure/General/properties.scala @@ -1,133 +1,133 @@ /* Title: Pure/General/properties.scala Author: Makarius Property lists. */ package isabelle object Properties { /* entries */ type Entry = (java.lang.String, java.lang.String) type T = List[Entry] def defined(props: T, name: java.lang.String): java.lang.Boolean = props.exists({ case (x, _) => x == name }) def get(props: T, name: java.lang.String): Option[java.lang.String] = props.collectFirst({ case (x, y) if x == name => y }) def put(props: T, entry: Entry): T = { val (x, y) = entry def update(ps: T): T = ps match { case (p @ (x1, _)) :: rest => if (x1 == x) (x1, y) :: rest else p :: update(rest) case Nil => Nil } if (defined(props, x)) update(props) else entry :: props } /* external storage */ def encode(ps: T): Bytes = Bytes(YXML.string_of_body(XML.Encode.properties(ps))) def decode(bs: Bytes, xml_cache: Option[XML.Cache] = None): T = { val ps = XML.Decode.properties(YXML.parse_body(bs.text)) xml_cache match { case None => ps case Some(cache) => cache.props(ps) } } def compress(ps: List[T], options: XZ.Options = XZ.options(), cache: XZ.Cache = XZ.cache()): Bytes = { if (ps.isEmpty) Bytes.empty else { Bytes(YXML.string_of_body(XML.Encode.list(XML.Encode.properties)(ps))). compress(options = options, cache = cache) } } def uncompress(bs: Bytes, cache: XZ.Cache = XZ.cache(), xml_cache: Option[XML.Cache] = None): List[T] = { if (bs.isEmpty) Nil else { val ps = XML.Decode.list(XML.Decode.properties)(YXML.parse_body(bs.uncompress(cache = cache).text)) xml_cache match { case None => ps - case Some(cache) => ps.map(cache.props(_)) + case Some(cache) => ps.map(cache.props) } } } /* multi-line entries */ def encode_lines(props: T): T = props.map({ case (a, b) => (a, Library.encode_lines(b)) }) def decode_lines(props: T): T = props.map({ case (a, b) => (a, Library.decode_lines(b)) }) def lines_nonempty(x: java.lang.String, ys: List[java.lang.String]): Properties.T = if (ys.isEmpty) Nil else List((x, cat_lines(ys))) /* entry types */ class String(val name: java.lang.String) { def apply(value: java.lang.String): T = List((name, value)) def unapply(props: T): Option[java.lang.String] = props.find(_._1 == name).map(_._2) } class Boolean(val name: java.lang.String) { def apply(value: scala.Boolean): T = List((name, Value.Boolean(value))) def unapply(props: T): Option[scala.Boolean] = props.find(_._1 == name) match { case None => None case Some((_, value)) => Value.Boolean.unapply(value) } } class Int(val name: java.lang.String) { def apply(value: scala.Int): T = List((name, Value.Int(value))) def unapply(props: T): Option[scala.Int] = props.find(_._1 == name) match { case None => None case Some((_, value)) => Value.Int.unapply(value) } } class Long(val name: java.lang.String) { def apply(value: scala.Long): T = List((name, Value.Long(value))) def unapply(props: T): Option[scala.Long] = props.find(_._1 == name) match { case None => None case Some((_, value)) => Value.Long.unapply(value) } } class Double(val name: java.lang.String) { def apply(value: scala.Double): T = List((name, Value.Double(value))) def unapply(props: T): Option[scala.Double] = props.find(_._1 == name) match { case None => None case Some((_, value)) => Value.Double.unapply(value) } } } 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,512 +1,512 @@ /* Title: Pure/General/scan.scala Author: Makarius Efficient scanning of keywords and tokens. */ package isabelle import scala.annotation.tailrec import scala.collection.{IndexedSeq, Traversable, TraversableOnce} import scala.collection.immutable.PagedSeq +import scala.util.matching.Regex import scala.util.parsing.input.{OffsetPosition, Position => InputPosition, Reader, CharSequenceReader} 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 = "".r + 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) 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) 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) 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) - val comment_text = + 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) 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) }) 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: TraversableOnce[String]): Lexicon = (this /: elems)(_ + _) def ++ (other: Lexicon): Lexicon = if (this eq other) this else if (is_empty) other else this ++ other.raw_iterator def -- (remove: Traversable[String]): Lexicon = - if (remove.exists(contains(_))) + 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 { 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/sql.scala b/src/Pure/General/sql.scala --- a/src/Pure/General/sql.scala +++ b/src/Pure/General/sql.scala @@ -1,514 +1,514 @@ /* Title: Pure/General/sql.scala Author: Makarius Support for SQL databases: SQLite and PostgreSQL. */ package isabelle import java.time.OffsetDateTime import java.sql.{DriverManager, Connection, PreparedStatement, ResultSet} import scala.collection.mutable object SQL { /** SQL language **/ type Source = String /* concrete syntax */ def escape_char(c: Char): String = c match { case '\u0000' => "\\0" case '\'' => "\\'" case '\"' => "\\\"" case '\b' => "\\b" case '\n' => "\\n" case '\r' => "\\r" case '\t' => "\\t" case '\u001a' => "\\Z" case '\\' => "\\\\" case _ => c.toString } def string(s: String): Source = - s.iterator.map(escape_char(_)).mkString("'", "", "'") + s.iterator.map(escape_char).mkString("'", "", "'") def ident(s: String): Source = Long_Name.implode(Long_Name.explode(s).map(a => quote(a.replace("\"", "\"\"")))) def enclose(s: Source): Source = "(" + s + ")" def enclosure(ss: Iterable[Source]): Source = ss.mkString("(", ", ", ")") def select(columns: List[Column] = Nil, distinct: Boolean = false): Source = "SELECT " + (if (distinct) "DISTINCT " else "") + (if (columns.isEmpty) "*" else commas(columns.map(_.ident))) + " FROM " val join_outer: Source = " LEFT OUTER JOIN " val join_inner: Source = " INNER JOIN " def join(outer: Boolean): Source = if (outer) join_outer else join_inner def member(x: Source, set: Iterable[String]): Source = set.iterator.map(a => x + " = " + SQL.string(a)).mkString("(", " OR ", ")") /* types */ object Type extends Enumeration { val Boolean = Value("BOOLEAN") val Int = Value("INTEGER") val Long = Value("BIGINT") val Double = Value("DOUBLE PRECISION") val String = Value("TEXT") val Bytes = Value("BLOB") val Date = Value("TIMESTAMP WITH TIME ZONE") } def sql_type_default(T: Type.Value): Source = T.toString def sql_type_sqlite(T: Type.Value): Source = if (T == Type.Boolean) "INTEGER" else if (T == Type.Date) "TEXT" else sql_type_default(T) def sql_type_postgresql(T: Type.Value): Source = if (T == Type.Bytes) "BYTEA" else sql_type_default(T) /* columns */ object Column { def bool(name: String, strict: Boolean = false, primary_key: Boolean = false): Column = Column(name, Type.Boolean, strict, primary_key) def int(name: String, strict: Boolean = false, primary_key: Boolean = false): Column = Column(name, Type.Int, strict, primary_key) def long(name: String, strict: Boolean = false, primary_key: Boolean = false): Column = Column(name, Type.Long, strict, primary_key) def double(name: String, strict: Boolean = false, primary_key: Boolean = false): Column = Column(name, Type.Double, strict, primary_key) def string(name: String, strict: Boolean = false, primary_key: Boolean = false): Column = Column(name, Type.String, strict, primary_key) def bytes(name: String, strict: Boolean = false, primary_key: Boolean = false): Column = Column(name, Type.Bytes, strict, primary_key) def date(name: String, strict: Boolean = false, primary_key: Boolean = false): Column = Column(name, Type.Date, strict, primary_key) } sealed case class Column( name: String, T: Type.Value, strict: Boolean = false, primary_key: Boolean = false, expr: SQL.Source = "") { def make_primary_key: Column = copy(primary_key = true) def apply(table: Table): Column = Column(Long_Name.qualify(table.name, name), T, strict = strict, primary_key = primary_key) def ident: Source = if (expr == "") SQL.ident(name) else enclose(expr) + " AS " + SQL.ident(name) def decl(sql_type: Type.Value => Source): Source = ident + " " + sql_type(T) + (if (strict || primary_key) " NOT NULL" else "") def defined: String = ident + " IS NOT NULL" def undefined: String = ident + " IS NULL" def equal(s: String): Source = ident + " = " + string(s) def where_equal(s: String): Source = "WHERE " + equal(s) override def toString: Source = ident } /* tables */ sealed case class Table(name: String, columns: List[Column], body: Source = "") { private val columns_index: Map[String, Int] = columns.iterator.map(_.name).zipWithIndex.toMap Library.duplicates(columns.map(_.name)) match { case Nil => case bad => error("Duplicate column names " + commas_quote(bad) + " for table " + quote(name)) } def ident: Source = SQL.ident(name) def query: Source = if (body == "") error("Missing SQL body for table " + quote(name)) else SQL.enclose(body) def query_named: Source = query + " AS " + SQL.ident(name) def create(strict: Boolean = false, sql_type: Type.Value => Source): Source = { val primary_key = columns.filter(_.primary_key).map(_.name) match { case Nil => Nil case keys => List("PRIMARY KEY " + enclosure(keys)) } "CREATE TABLE " + (if (strict) "" else "IF NOT EXISTS ") + ident + " " + enclosure(columns.map(_.decl(sql_type)) ::: primary_key) } def create_index(index_name: String, index_columns: List[Column], strict: Boolean = false, unique: Boolean = false): Source = "CREATE " + (if (unique) "UNIQUE " else "") + "INDEX " + (if (strict) "" else "IF NOT EXISTS ") + SQL.ident(index_name) + " ON " + ident + " " + enclosure(index_columns.map(_.name)) def insert_cmd(cmd: Source, sql: Source = ""): Source = cmd + " INTO " + ident + " VALUES " + enclosure(columns.map(_ => "?")) + (if (sql == "") "" else " " + sql) def insert(sql: Source = ""): Source = insert_cmd("INSERT", sql) def delete(sql: Source = ""): Source = "DELETE FROM " + ident + (if (sql == "") "" else " " + sql) def select( select_columns: List[Column] = Nil, sql: Source = "", distinct: Boolean = false): Source = SQL.select(select_columns, distinct = distinct) + ident + (if (sql == "") "" else " " + sql) override def toString: Source = ident } /** SQL database operations **/ /* statements */ class Statement private[SQL](val db: Database, val rep: PreparedStatement) extends AutoCloseable { stmt => object bool { def update(i: Int, x: Boolean) { rep.setBoolean(i, x) } def update(i: Int, x: Option[Boolean]) { if (x.isDefined) update(i, x.get) else rep.setNull(i, java.sql.Types.BOOLEAN) } } object int { def update(i: Int, x: Int) { rep.setInt(i, x) } def update(i: Int, x: Option[Int]) { if (x.isDefined) update(i, x.get) else rep.setNull(i, java.sql.Types.INTEGER) } } object long { def update(i: Int, x: Long) { rep.setLong(i, x) } def update(i: Int, x: Option[Long]) { if (x.isDefined) update(i, x.get) else rep.setNull(i, java.sql.Types.BIGINT) } } object double { def update(i: Int, x: Double) { rep.setDouble(i, x) } def update(i: Int, x: Option[Double]) { if (x.isDefined) update(i, x.get) else rep.setNull(i, java.sql.Types.DOUBLE) } } object string { def update(i: Int, x: String) { rep.setString(i, x) } def update(i: Int, x: Option[String]): Unit = update(i, x.orNull) } object bytes { def update(i: Int, bytes: Bytes) { if (bytes == null) rep.setBytes(i, null) else rep.setBinaryStream(i, bytes.stream(), bytes.length) } def update(i: Int, bytes: Option[Bytes]): Unit = update(i, bytes.orNull) } object date { def update(i: Int, date: Date): Unit = db.update_date(stmt, i, date) def update(i: Int, date: Option[Date]): Unit = update(i, date.orNull) } def execute(): Boolean = rep.execute() def execute_query(): Result = new Result(this, rep.executeQuery()) def close(): Unit = rep.close } /* results */ class Result private[SQL](val stmt: Statement, val rep: ResultSet) { res => def next(): Boolean = rep.next() def iterator[A](get: Result => A): Iterator[A] = new Iterator[A] { private var _next: Boolean = res.next() def hasNext: Boolean = _next def next: A = { val x = get(res); _next = res.next(); x } } def bool(column: Column): Boolean = rep.getBoolean(column.name) def int(column: Column): Int = rep.getInt(column.name) def long(column: Column): Long = rep.getLong(column.name) def double(column: Column): Double = rep.getDouble(column.name) def string(column: Column): String = { val s = rep.getString(column.name) if (s == null) "" else s } def bytes(column: Column): Bytes = { val bs = rep.getBytes(column.name) if (bs == null) Bytes.empty else Bytes(bs) } def date(column: Column): Date = stmt.db.date(res, column) - def timing(c1: Column, c2: Column, c3: Column) = + def timing(c1: Column, c2: Column, c3: Column): Timing = Timing(Time.ms(long(c1)), Time.ms(long(c2)), Time.ms(long(c3))) def get[A](column: Column, f: Column => A): Option[A] = { val x = f(column) if (rep.wasNull) None else Some(x) } - def get_bool(column: Column): Option[Boolean] = get(column, bool _) - def get_int(column: Column): Option[Int] = get(column, int _) - def get_long(column: Column): Option[Long] = get(column, long _) - def get_double(column: Column): Option[Double] = get(column, double _) - def get_string(column: Column): Option[String] = get(column, string _) - def get_bytes(column: Column): Option[Bytes] = get(column, bytes _) - def get_date(column: Column): Option[Date] = get(column, date _) + def get_bool(column: Column): Option[Boolean] = get(column, bool) + def get_int(column: Column): Option[Int] = get(column, int) + def get_long(column: Column): Option[Long] = get(column, long) + def get_double(column: Column): Option[Double] = get(column, double) + def get_string(column: Column): Option[String] = get(column, string) + def get_bytes(column: Column): Option[Bytes] = get(column, bytes) + def get_date(column: Column): Option[Date] = get(column, date) } /* database */ trait Database extends AutoCloseable { db => /* types */ def sql_type(T: Type.Value): Source /* connection */ def connection: Connection def close() { connection.close } def transaction[A](body: => A): A = { val auto_commit = connection.getAutoCommit try { connection.setAutoCommit(false) val savepoint = connection.setSavepoint try { val result = body connection.commit result } catch { case exn: Throwable => connection.rollback(savepoint); throw exn } } finally { connection.setAutoCommit(auto_commit) } } /* statements and results */ def statement(sql: Source): Statement = new Statement(db, connection.prepareStatement(sql)) def using_statement[A](sql: Source)(f: Statement => A): A = using(statement(sql))(f) def update_date(stmt: Statement, i: Int, date: Date): Unit def date(res: Result, column: Column): Date def insert_permissive(table: Table, sql: Source = ""): Source /* tables and views */ def tables: List[String] = { val result = new mutable.ListBuffer[String] val rs = connection.getMetaData.getTables(null, null, "%", null) while (rs.next) { result += rs.getString(3) } result.toList } def create_table(table: Table, strict: Boolean = false, sql: Source = ""): Unit = using_statement( table.create(strict, sql_type) + (if (sql == "") "" else " " + sql))(_.execute()) def create_index(table: Table, name: String, columns: List[Column], strict: Boolean = false, unique: Boolean = false): Unit = using_statement(table.create_index(name, columns, strict, unique))(_.execute()) def create_view(table: Table, strict: Boolean = false): Unit = { if (strict || !tables.contains(table.name)) { val sql = "CREATE VIEW " + table + " AS " + { table.query; table.body } using_statement(sql)(_.execute()) } } } } /** SQLite **/ object SQLite { // see https://www.sqlite.org/lang_datefunc.html val date_format: Date.Format = Date.Format("uuuu-MM-dd HH:mm:ss.SSS x") lazy val init_jdbc: Unit = { val lib_path = Path.explode("$ISABELLE_SQLITE_HOME/" + Platform.jvm_platform) val lib_name = File.find_files(lib_path.file) match { case List(file) => file.getName case _ => error("Exactly one file expected in directory " + lib_path.expand) } System.setProperty("org.sqlite.lib.path", File.platform_path(lib_path)) System.setProperty("org.sqlite.lib.name", lib_name) Class.forName("org.sqlite.JDBC") } def open_database(path: Path): Database = { init_jdbc val path0 = path.expand val s0 = File.platform_path(path0) val s1 = if (Platform.is_windows) s0.replace('\\', '/') else s0 val connection = DriverManager.getConnection("jdbc:sqlite:" + s1) new Database(path0.toString, connection) } class Database private[SQLite](name: String, val connection: Connection) extends SQL.Database { override def toString: String = name def sql_type(T: SQL.Type.Value): SQL.Source = SQL.sql_type_sqlite(T) def update_date(stmt: SQL.Statement, i: Int, date: Date): Unit = if (date == null) stmt.string(i) = (null: String) else stmt.string(i) = date_format(date) def date(res: SQL.Result, column: SQL.Column): Date = date_format.parse(res.string(column)) def insert_permissive(table: SQL.Table, sql: SQL.Source = ""): SQL.Source = table.insert_cmd("INSERT OR IGNORE", sql = sql) def rebuild { using_statement("VACUUM")(_.execute()) } } } /** PostgreSQL **/ object PostgreSQL { val default_port = 5432 lazy val init_jdbc: Unit = Class.forName("org.postgresql.Driver") def open_database( user: String, password: String, database: String = "", host: String = "", port: Int = 0, ssh: Option[SSH.Session] = None, ssh_close: Boolean = false): Database = { init_jdbc if (user == "") error("Undefined database user") val db_host = proper_string(host) getOrElse "localhost" val db_port = if (port > 0 && port != default_port) ":" + port else "" val db_name = "/" + (proper_string(database) getOrElse user) val (url, name, port_forwarding) = ssh match { case None => val spec = db_host + db_port + db_name val url = "jdbc:postgresql://" + spec val name = user + "@" + spec (url, name, None) case Some(ssh) => val fw = ssh.port_forwarding(remote_host = db_host, remote_port = if (port > 0) port else default_port, ssh_close = ssh_close) val url = "jdbc:postgresql://localhost:" + fw.local_port + db_name val name = user + "@" + fw + db_name + " via ssh " + ssh (url, name, Some(fw)) } try { val connection = DriverManager.getConnection(url, user, password) new Database(name, connection, port_forwarding) } catch { case exn: Throwable => port_forwarding.foreach(_.close); throw exn } } class Database private[PostgreSQL]( name: String, val connection: Connection, port_forwarding: Option[SSH.Port_Forwarding]) extends SQL.Database { override def toString: String = name def sql_type(T: SQL.Type.Value): SQL.Source = SQL.sql_type_postgresql(T) // see https://jdbc.postgresql.org/documentation/head/8-date-time.html def update_date(stmt: SQL.Statement, i: Int, date: Date): Unit = if (date == null) stmt.rep.setObject(i, null) else stmt.rep.setObject(i, OffsetDateTime.from(date.to(Date.timezone_utc).rep)) def date(res: SQL.Result, column: SQL.Column): Date = { val obj = res.rep.getObject(column.name, classOf[OffsetDateTime]) if (obj == null) null else Date.instant(obj.toInstant) } def insert_permissive(table: SQL.Table, sql: SQL.Source = ""): SQL.Source = table.insert_cmd("INSERT", sql = sql + (if (sql == "") "" else " ") + "ON CONFLICT DO NOTHING") override def close() { super.close; port_forwarding.foreach(_.close) } } } diff --git a/src/Pure/General/ssh.scala b/src/Pure/General/ssh.scala --- a/src/Pure/General/ssh.scala +++ b/src/Pure/General/ssh.scala @@ -1,500 +1,501 @@ /* Title: Pure/General/ssh.scala Author: Makarius SSH client based on JSch (see also http://www.jcraft.com/jsch/examples). */ package isabelle import java.io.{InputStream, OutputStream, ByteArrayOutputStream} import scala.collection.mutable +import scala.util.matching.Regex import com.jcraft.jsch.{JSch, Logger => JSch_Logger, Session => JSch_Session, SftpException, OpenSSHConfig, UserInfo, Channel => JSch_Channel, ChannelExec, ChannelSftp, SftpATTRS} object SSH { /* target machine: user@host syntax */ object Target { - val User_Host = "^([^@]+)@(.+)$".r + val User_Host: Regex = "^([^@]+)@(.+)$".r def parse(s: String): (String, String) = s match { case User_Host(user, host) => (user, host) case _ => ("", s) } def unapplySeq(s: String): Option[List[String]] = parse(s) match { case (_, "") => None case (user, host) => Some(List(user, host)) } } val default_port = 22 def make_port(port: Int): Int = if (port > 0) port else default_port def port_suffix(port: Int): String = if (port == default_port) "" else ":" + port def user_prefix(user: String): String = proper_string(user) match { case None => "" case Some(name) => name + "@" } def connect_timeout(options: Options): Int = options.seconds("ssh_connect_timeout").ms.toInt def alive_interval(options: Options): Int = options.seconds("ssh_alive_interval").ms.toInt def alive_count_max(options: Options): Int = options.int("ssh_alive_count_max") /* init context */ def init_context(options: Options): Context = { val config_dir = Path.explode(options.string("ssh_config_dir")) if (!config_dir.is_dir) error("Bad ssh config directory: " + config_dir) val jsch = new JSch val config_file = Path.explode(options.string("ssh_config_file")) if (config_file.is_file) jsch.setConfigRepository(OpenSSHConfig.parseFile(File.platform_path(config_file))) val known_hosts = config_dir + Path.explode("known_hosts") if (!known_hosts.is_file) known_hosts.file.createNewFile jsch.setKnownHosts(File.platform_path(known_hosts)) val identity_files = - space_explode(':', options.string("ssh_identity_files")).map(Path.explode(_)) + space_explode(':', options.string("ssh_identity_files")).map(Path.explode) for (identity_file <- identity_files if identity_file.is_file) jsch.addIdentity(File.platform_path(identity_file)) new Context(options, jsch) } def open_session(options: Options, host: String, user: String = "", port: Int = 0, actual_host: String = "", proxy_host: String = "", proxy_user: String = "", proxy_port: Int = 0, permissive: Boolean = false): Session = init_context(options).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 = permissive) class Context private[SSH](val options: Options, val jsch: JSch) { def update_options(new_options: Options): Context = new Context(new_options, jsch) private def connect_session(host: String, user: String = "", port: Int = 0, host_key_permissive: Boolean = false, nominal_host: String = "", nominal_user: String = "", on_close: () => Unit = () => ()): Session = { val session = jsch.getSession(proper_string(user).orNull, host, make_port(port)) session.setUserInfo(No_User_Info) session.setServerAliveInterval(alive_interval(options)) session.setServerAliveCountMax(alive_count_max(options)) session.setConfig("MaxAuthTries", "3") if (host_key_permissive) session.setConfig("StrictHostKeyChecking", "no") if (nominal_host != "") session.setHostKeyAlias(nominal_host) if (options.bool("ssh_compression")) { session.setConfig("compression.s2c", "zlib@openssh.com,zlib,none") session.setConfig("compression.c2s", "zlib@openssh.com,zlib,none") session.setConfig("compression_level", "9") } session.connect(connect_timeout(options)) new Session(options, session, on_close, proper_string(nominal_host) getOrElse host, proper_string(nominal_user) getOrElse user) } def open_session( host: String, user: String = "", port: Int = 0, actual_host: String = "", proxy_host: String = "", proxy_user: String = "", proxy_port: Int = 0, permissive: Boolean = false): Session = { val connect_host = proper_string(actual_host) getOrElse host if (proxy_host == "") connect_session(host = connect_host, user = user, port = port) else { val proxy = connect_session(host = proxy_host, port = proxy_port, user = proxy_user) val fw = try { proxy.port_forwarding(remote_host = connect_host, remote_port = make_port(port)) } catch { case exn: Throwable => proxy.close; throw exn } try { connect_session(host = fw.local_host, port = fw.local_port, host_key_permissive = permissive, nominal_host = host, nominal_user = user, user = user, on_close = () => { fw.close; proxy.close }) } catch { case exn: Throwable => fw.close; proxy.close; throw exn } } } } /* logging */ def logging(verbose: Boolean = true, debug: Boolean = false) { JSch.setLogger(if (verbose) new Logger(debug) else null) } private class Logger(debug: Boolean) extends JSch_Logger { def isEnabled(level: Int): Boolean = level != JSch_Logger.DEBUG || debug def log(level: Int, msg: String) { level match { case JSch_Logger.ERROR | JSch_Logger.FATAL => Output.error_message(msg) case JSch_Logger.WARN => Output.warning(msg) case _ => Output.writeln(msg) } } } /* user info */ object No_User_Info extends UserInfo { def getPassphrase: String = null def getPassword: String = null def promptPassword(msg: String): Boolean = false def promptPassphrase(msg: String): Boolean = false def promptYesNo(msg: String): Boolean = false def showMessage(msg: String): Unit = Output.writeln(msg) } /* port forwarding */ object Port_Forwarding { def open(ssh: Session, ssh_close: Boolean, local_host: String, local_port: Int, remote_host: String, remote_port: Int): Port_Forwarding = { val port = ssh.session.setPortForwardingL(local_host, local_port, remote_host, remote_port) new Port_Forwarding(ssh, ssh_close, local_host, port, remote_host, remote_port) } } class Port_Forwarding private[SSH]( ssh: SSH.Session, ssh_close: Boolean, val local_host: String, val local_port: Int, val remote_host: String, val remote_port: Int) extends AutoCloseable { override def toString: String = local_host + ":" + local_port + ":" + remote_host + ":" + remote_port def close() { ssh.session.delPortForwardingL(local_host, local_port) if (ssh_close) ssh.close() } } /* Sftp channel */ type Attrs = SftpATTRS sealed case class Dir_Entry(name: String, is_dir: Boolean) { def is_file: Boolean = !is_dir } /* exec channel */ private val exec_wait_delay = Time.seconds(0.3) class Exec private[SSH](session: Session, channel: ChannelExec) extends AutoCloseable { override def toString: String = "exec " + session.toString def close() { channel.disconnect } val exit_status: Future[Int] = Future.thread("ssh_wait") { while (!channel.isClosed) Thread.sleep(exec_wait_delay.ms) channel.getExitStatus } val stdin: OutputStream = channel.getOutputStream val stdout: InputStream = channel.getInputStream val stderr: InputStream = channel.getErrStream // connect after preparing streams channel.connect(connect_timeout(session.options)) def result( progress_stdout: String => Unit = (_: String) => (), progress_stderr: String => Unit = (_: String) => (), strict: Boolean = true): Process_Result = { stdin.close def read_lines(stream: InputStream, progress: String => Unit): List[String] = { val result = new mutable.ListBuffer[String] val line_buffer = new ByteArrayOutputStream(100) def line_flush() { val line = Library.trim_line(line_buffer.toString(UTF8.charset_name)) progress(line) result += line line_buffer.reset } var c = 0 var finished = false while (!finished) { while ({ c = stream.read; c != -1 && c != 10 }) line_buffer.write(c) if (c == 10) line_flush() else if (channel.isClosed) { if (line_buffer.size > 0) line_flush() finished = true } else Thread.sleep(exec_wait_delay.ms) } result.toList } val out_lines = Future.thread("ssh_stdout") { read_lines(stdout, progress_stdout) } val err_lines = Future.thread("ssh_stderr") { read_lines(stderr, progress_stderr) } def terminate() { close out_lines.join err_lines.join exit_status.join } val rc = try { exit_status.join } catch { case Exn.Interrupt() => terminate(); Exn.Interrupt.return_code } close if (strict && rc == Exn.Interrupt.return_code) throw Exn.Interrupt() Process_Result(rc, out_lines.join, err_lines.join) } } /* session */ class Session private[SSH]( val options: Options, val session: JSch_Session, on_close: () => Unit, val nominal_host: String, val nominal_user: String) extends System with AutoCloseable { def update_options(new_options: Options): Session = new Session(new_options, session, on_close, nominal_host, nominal_user) def host: String = if (session.getHost == null) "" else session.getHost override def hg_url: String = "ssh://" + user_prefix(nominal_user) + nominal_host + "/" override def toString: String = user_prefix(session.getUserName) + host + port_suffix(session.getPort) + (if (session.isConnected) "" else " (disconnected)") /* port forwarding */ def port_forwarding( remote_port: Int, remote_host: String = "localhost", local_port: Int = 0, local_host: String = "localhost", ssh_close: Boolean = false): Port_Forwarding = Port_Forwarding.open(this, ssh_close, local_host, local_port, remote_host, remote_port) /* sftp channel */ val sftp: ChannelSftp = session.openChannel("sftp").asInstanceOf[ChannelSftp] sftp.connect(connect_timeout(options)) def close() { sftp.disconnect; session.disconnect; on_close() } val settings: Map[String, String] = { val home = sftp.getHome Map("HOME" -> home, "USER_HOME" -> home) } override def expand_path(path: Path): Path = path.expand_env(settings) def remote_path(path: Path): String = expand_path(path).implode override def bash_path(path: Path): String = Bash.string(remote_path(path)) def chmod(permissions: Int, path: Path): Unit = sftp.chmod(permissions, remote_path(path)) def mv(path1: Path, path2: Path): Unit = sftp.rename(remote_path(path1), remote_path(path2)) def rm(path: Path): Unit = sftp.rm(remote_path(path)) def mkdir(path: Path): Unit = sftp.mkdir(remote_path(path)) def rmdir(path: Path): Unit = sftp.rmdir(remote_path(path)) private def test_entry(path: Path, as_dir: Boolean): Boolean = try { val is_dir = sftp.stat(remote_path(path)).isDir if (as_dir) is_dir else !is_dir } catch { case _: SftpException => false } override def is_dir(path: Path): Boolean = test_entry(path, true) override def is_file(path: Path): Boolean = test_entry(path, false) def is_link(path: Path): Boolean = try { sftp.lstat(remote_path(path)).isLink } catch { case _: SftpException => false } override def mkdirs(path: Path): Unit = if (!is_dir(path)) { execute( "perl -e \"use File::Path make_path; make_path('" + remote_path(path) + "');\"") if (!is_dir(path)) error("Failed to create directory: " + quote(remote_path(path))) } def read_dir(path: Path): List[Dir_Entry] = { if (!is_dir(path)) error("No such directory: " + path.toString) val dir_name = remote_path(path) val dir = sftp.ls(dir_name) (for { i <- (0 until dir.size).iterator a = dir.get(i).asInstanceOf[AnyRef] name = Untyped.get[String](a, "filename") attrs = Untyped.get[Attrs](a, "attrs") if name != "." && name != ".." } yield { Dir_Entry(name, if (attrs.isLink) { try { sftp.stat(dir_name + "/" + name).isDir } catch { case _: SftpException => false } } else attrs.isDir) }).toList.sortBy(_.name) } def find_files( start: Path, pred: Path => Boolean = _ => true, include_dirs: Boolean = false, follow_links: Boolean = false): List[Path] = { val result = new mutable.ListBuffer[Path] def check(path: Path) { if (pred(path)) result += path } def find(dir: Path) { if (include_dirs) check(dir) if (follow_links || !is_link(dir)) { for (entry <- read_dir(dir)) { val path = dir + Path.basic(entry.name) if (entry.is_file) check(path) else find(path) } } } if (is_file(start)) check(start) else find(start) result.toList } def open_input(path: Path): InputStream = sftp.get(remote_path(path)) def open_output(path: Path): OutputStream = sftp.put(remote_path(path)) def read_file(path: Path, local_path: Path): Unit = sftp.get(remote_path(path), File.platform_path(local_path)) def read_bytes(path: Path): Bytes = using(open_input(path))(Bytes.read_stream(_)) - def read(path: Path): String = using(open_input(path))(File.read_stream(_)) + def read(path: Path): String = using(open_input(path))(File.read_stream) def write_file(path: Path, local_path: Path): Unit = sftp.put(File.platform_path(local_path), remote_path(path)) def write_bytes(path: Path, bytes: Bytes): Unit = - using(open_output(path))(bytes.write_stream(_)) + using(open_output(path))(bytes.write_stream) def write(path: Path, text: String): Unit = using(open_output(path))(stream => Bytes(text).write_stream(stream)) /* exec channel */ def exec(command: String): Exec = { val channel = session.openChannel("exec").asInstanceOf[ChannelExec] channel.setCommand("export USER_HOME=\"$HOME\"\n" + command) new Exec(this, channel) } override def execute(command: String, progress_stdout: String => Unit = (_: String) => (), progress_stderr: String => Unit = (_: String) => (), strict: Boolean = true): Process_Result = exec(command).result(progress_stdout, progress_stderr, strict) /* tmp dirs */ def rm_tree(dir: Path): Unit = rm_tree(remote_path(dir)) def rm_tree(remote_dir: String): Unit = execute("rm -r -f " + Bash.string(remote_dir)).check def tmp_dir(): String = execute("mktemp -d -t tmp.XXXXXXXXXX").check.out def with_tmp_dir[A](body: Path => A): A = { val remote_dir = tmp_dir() try { body(Path.explode(remote_dir)) } finally { rm_tree(remote_dir) } } } /* system operations */ trait System { def hg_url: String = "" def expand_path(path: Path): Path = path.expand def bash_path(path: Path): String = File.bash_path(path) def is_dir(path: Path): Boolean = path.is_dir def is_file(path: Path): Boolean = path.is_file def mkdirs(path: Path): Unit = Isabelle_System.mkdirs(path) def execute(command: String, progress_stdout: String => Unit = (_: String) => (), progress_stderr: String => Unit = (_: String) => (), strict: Boolean = true): Process_Result = Isabelle_System.bash(command, progress_stdout = progress_stdout, progress_stderr = progress_stderr, strict = strict) } object Local extends System } 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,657 +1,657 @@ /* 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 = " " private val static_spaces = space * 4000 def spaces(n: Int): String = { require(n >= 0) if (n < static_spaces.length) static_spaces.substring(0, n) else space * n } /* ASCII characters */ 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) 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 = i < text.length - def next = + 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 + Library.trim(is_blank, explode(text)).mkString def all_blank(str: String): Boolean = - iterator(str).forall(is_blank(_)) + 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(_)) + 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 val encode_name: XML.Encode.T[Name] = { import XML.Encode._ variant(List( { case Default => (Nil, Nil) }, { case Id(a) => (List(long_atom(a)), Nil) }, { case File(a) => (List(a), Nil) })) } val decode_name: XML.Decode.T[Name] = { import XML.Decode._ variant(List( { case (Nil, Nil) => Default }, { case (List(a), Nil) => Id(long_atom(a)) }, { case (List(a), Nil) => File(a) })) } 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 case ((list, known), decl) => val (sym, props) = read_decl(decl) if (known(sym)) (list, known) else ((sym, props) :: list, known + sym) })._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 = recode_set( + 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 = recode_set(space, "\t", "\n", "\u000B", "\f", "\r", "\r\n") + val blanks: Set[String] = recode_set(space, "\t", "\n", "\u000B", "\f", "\r", "\r\n") val sym_chars = Set("!", "#", "$", "%", "&", "*", "+", "-", "/", "<", "=", ">", "?", "@", "^", "_", "|", "~") - val symbolic = recode_set((for { (sym, _) <- symbols; if raw_symbolic(sym) } yield sym): _*) + 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): XML.Body = YXML.parse_body(decode(text)) def decode_yxml_failsafe(text: String): XML.Body = YXML.parse_body_failsafe(decode(text)) 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 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 = "\\<^sub>" val sup = "\\<^sup>" val bold = "\\<^bold>" val emph = "\\<^emph>" val bsub = "\\<^bsub>" val esub = "\\<^esub>" val bsup = "\\<^bsup>" val esup = "\\<^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 } diff --git a/src/Pure/General/url.scala b/src/Pure/General/url.scala --- a/src/Pure/General/url.scala +++ b/src/Pure/General/url.scala @@ -1,95 +1,95 @@ /* Title: Pure/General/url.scala Author: Makarius Basic URL operations. */ package isabelle import java.io.{File => JFile} import java.nio.file.{Paths, FileSystemNotFoundException} import java.net.{URI, URISyntaxException, URL, MalformedURLException, URLDecoder, URLEncoder} import java.util.Locale import java.util.zip.GZIPInputStream object Url { /* special characters */ def escape_special(c: Char): String = if ("!#$&'()*+,/:;=?@[]".contains(c)) { String.format(Locale.ROOT, "%%%02X", Integer.valueOf(c.toInt)) } else c.toString - def escape_special(s: String): String = s.iterator.map(escape_special(_)).mkString + def escape_special(s: String): String = s.iterator.map(escape_special).mkString def escape_name(name: String): String = name.iterator.map({ case '\'' => "%27" case c => c.toString }).mkString /* make and check URLs */ def apply(name: String): URL = { try { new URL(name) } catch { case _: MalformedURLException => error("Malformed URL " + quote(name)) } } def is_wellformed(name: String): Boolean = try { Url(name); true } catch { case ERROR(_) => false } def is_readable(name: String): Boolean = try { Url(name).openStream.close; true } catch { case ERROR(_) => false } /* strings */ def decode(s: String): String = URLDecoder.decode(s, UTF8.charset_name) def encode(s: String): String = URLEncoder.encode(s, UTF8.charset_name) /* read */ private def read(url: URL, gzip: Boolean): String = using(url.openStream)(stream => File.read_stream(if (gzip) new GZIPInputStream(stream) else stream)) def read(url: URL): String = read(url, false) def read_gzip(url: URL): String = read(url, true) def read(name: String): String = read(Url(name), false) def read_gzip(name: String): String = read(Url(name), true) def read_bytes(url: URL): Bytes = { val connection = url.openConnection val length = connection.getContentLength using(connection.getInputStream)(Bytes.read_stream(_, hint = length)) } /* file URIs */ def print_file(file: JFile): String = File.absolute(file).toPath.toUri.toString def print_file_name(name: String): String = print_file(new JFile(name)) def parse_file(uri: String): JFile = Paths.get(new URI(uri)).toFile def is_wellformed_file(uri: String): Boolean = try { parse_file(uri); true } catch { case _: URISyntaxException | _: IllegalArgumentException | _: FileSystemNotFoundException => false } def absolute_file(uri: String): JFile = File.absolute(parse_file(uri)) def absolute_file_name(uri: String): String = absolute_file(uri).getPath def canonical_file(uri: String): JFile = File.canonical(parse_file(uri)) def canonical_file_name(uri: String): String = canonical_file(uri).getPath } diff --git a/src/Pure/General/value.scala b/src/Pure/General/value.scala --- a/src/Pure/General/value.scala +++ b/src/Pure/General/value.scala @@ -1,77 +1,77 @@ /* Title: Pure/General/value.scala Author: Makarius Plain values, represented as string. */ package isabelle object Value { object Boolean { def apply(x: scala.Boolean): java.lang.String = x.toString def unapply(s: java.lang.String): Option[scala.Boolean] = s match { case "true" => Some(true) case "false" => Some(false) case _ => None } def parse(s: java.lang.String): scala.Boolean = unapply(s) getOrElse error("Bad boolean: " + quote(s)) } object Nat { def unapply(s: java.lang.String): Option[scala.Int] = s match { case Int(n) if n >= 0 => Some(n) case _ => None } def parse(s: java.lang.String): scala.Int = unapply(s) getOrElse error("Bad natural number: " + quote(s)) } object Int { def apply(x: scala.Int): java.lang.String = Library.signed_string_of_int(x) def unapply(s: java.lang.String): Option[scala.Int] = try { Some(Integer.parseInt(s)) } catch { case _: NumberFormatException => None } def parse(s: java.lang.String): scala.Int = unapply(s) getOrElse error("Bad integer: " + quote(s)) } object Long { def apply(x: scala.Long): java.lang.String = Library.signed_string_of_long(x) def unapply(s: java.lang.String): Option[scala.Long] = try { Some(java.lang.Long.parseLong(s)) } catch { case _: NumberFormatException => None } def parse(s: java.lang.String): scala.Long = unapply(s) getOrElse error("Bad integer: " + quote(s)) } object Double { def apply(x: scala.Double): java.lang.String = x.toString def unapply(s: java.lang.String): Option[scala.Double] = try { Some(java.lang.Double.parseDouble(s)) } catch { case _: NumberFormatException => None } def parse(s: java.lang.String): scala.Double = unapply(s) getOrElse error("Bad real: " + quote(s)) } object Seconds { def apply(t: Time): java.lang.String = { val s = t.seconds if (s.toInt.toDouble == s) s.toInt.toString else t.toString } - def unapply(s: java.lang.String): Option[Time] = Double.unapply(s).map(Time.seconds(_)) + def unapply(s: java.lang.String): Option[Time] = Double.unapply(s).map(Time.seconds) def parse(s: java.lang.String): Time = unapply(s) getOrElse error("Bad real (for seconds): " + quote(s)) } } diff --git a/src/Pure/General/word.scala b/src/Pure/General/word.scala --- a/src/Pure/General/word.scala +++ b/src/Pure/General/word.scala @@ -1,87 +1,87 @@ /* Title: Pure/General/word.scala Author: Makarius Support for words within Unicode text. */ package isabelle import java.text.Bidi import java.util.Locale object Word { /* directionality */ def bidi_detect(str: String): Boolean = str.exists(c => c >= 0x590) && Bidi.requiresBidi(str.toArray, 0, str.length) def bidi_override(str: String): String = if (bidi_detect(str)) "\u200E\u202D" + str + "\u202C" else str /* case */ def lowercase(str: String): String = str.toLowerCase(Locale.ROOT) def uppercase(str: String): String = str.toUpperCase(Locale.ROOT) def capitalize(str: String): String = if (str.length == 0) str else { val n = Character.charCount(str.codePointAt(0)) uppercase(str.substring(0, n)) + lowercase(str.substring(n)) } def perhaps_capitalize(str: String): String = if (Codepoint.iterator(str).forall(c => Character.isLowerCase(c) || Character.isDigit(c))) capitalize(str) else str sealed abstract class Case case object Lowercase extends Case case object Uppercase extends Case case object Capitalized extends Case object Case { def apply(c: Case, str: String): String = c match { case Lowercase => lowercase(str) case Uppercase => uppercase(str) case Capitalized => capitalize(str) } def unapply(str: String): Option[Case] = if (str.nonEmpty) { - if (Codepoint.iterator(str).forall(Character.isLowerCase(_))) Some(Lowercase) - else if (Codepoint.iterator(str).forall(Character.isUpperCase(_))) Some(Uppercase) + if (Codepoint.iterator(str).forall(Character.isLowerCase)) Some(Lowercase) + else if (Codepoint.iterator(str).forall(Character.isUpperCase)) Some(Uppercase) else { val it = Codepoint.iterator(str) - if (Character.isUpperCase(it.next) && it.forall(Character.isLowerCase(_))) + if (Character.isUpperCase(it.next) && it.forall(Character.isLowerCase)) Some(Capitalized) else None } } else None } /* sequence of words */ def implode(words: Iterable[String]): String = words.iterator.mkString(" ") def explode(sep: Char => Boolean, text: String): List[String] = Library.separated_chunks(sep, text).map(_.toString).filter(_ != "").toList def explode(sep: Char, text: String): List[String] = explode(_ == sep, text) def explode(text: String): List[String] = explode(Character.isWhitespace(_), text) /* brackets */ val open_brackets = "([{«‹⟨⌈⌊⦇⟦⦃" val close_brackets = ")]}»›⟩⌉⌋⦈⟧⦄" } 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,228 +1,228 @@ /* 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 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 **/ object Spec { val none: Spec = Spec("") } sealed case class Spec(kind: String, exts: List[String] = Nil, tags: List[String] = Nil) { def is_none: Boolean = kind == "" override def toString: String = kind + (if (exts.isEmpty) "" else " (" + commas_quote(exts) + ")") + (if (tags.isEmpty) "" else tags.map(quote).mkString(" % ", " % ", "")) } object Keywords { def empty: Keywords = new Keywords() } class Keywords private( val kinds: Map[String, String] = Map.empty, val load_commands: Map[String, List[String]] = Map.empty) { override def toString: String = { val entries = for ((name, kind) <- kinds.toList.sortBy(_._1)) yield { val exts = load_commands.getOrElse(name, Nil) val kind_decl = if (kind == "") "" else " :: " + quote(kind) + (if (exts.isEmpty) "" else " (" + commas_quote(exts) + ")") 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 } 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 } new Keywords(kinds1, load_commands1) } /* add keywords */ def + (name: String, kind: String = "", exts: List[String] = Nil): Keywords = { val kinds1 = kinds + (name -> kind) val load_commands1 = if (kind == THY_LOAD) { - if (!Symbol.iterator(name).forall(Symbol.is_ascii(_))) + if (!Symbol.iterator(name).forall(Symbol.is_ascii)) error("Bad theory load command " + quote(name)) load_commands + (name -> exts) } else load_commands new Keywords(kinds1, load_commands1) } def add_keywords(header: Thy_Header.Keywords): Keywords = (this /: header) { case (keywords, (name, spec)) => if (spec.is_none) keywords + Symbol.decode(name) + Symbol.encode(name) else keywords + (Symbol.decode(name), spec.kind, spec.exts) + (Symbol.encode(name), spec.kind, spec.exts) } /* 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) /* load commands */ def load_commands_in(text: String): Boolean = load_commands.exists({ case (cmd, _) => text.containsSlice(cmd) }) /* 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 }) 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() + 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).toStream.headOption 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) { 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,207 +1,207 @@ /* 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)(_ ++ _) /* 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 = "", exts: List[String] = Nil): Outer_Syntax = new Outer_Syntax( keywords + (name, kind, exts), rev_abbrevs, language_context, has_tokens = true) def add_keywords(keywords: Thy_Header.Keywords): Outer_Syntax = (this /: keywords) { case (syntax, (name, spec)) => syntax + (Symbol.decode(name), spec.kind, spec.exts) + (Symbol.encode(name), spec.kind, spec.exts) } /* 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[List[String]] = keywords.load_commands.get(name) def load_commands_in(text: String): Boolean = keywords.load_commands_in(text) /* 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) 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(span: List[Token]) { val kind = if (span.forall(_.is_ignored)) Command_Span.Ignored_Span else if (span.exists(_.is_error)) Command_Span.Malformed_Span else span.find(_.is_command) match { case None => Command_Span.Malformed_Span case Some(cmd) => val name = cmd.source val offset = (0 /: span.takeWhile(_ != cmd)) { case (i, tok) => i + Symbol.length(tok.source) } val end_offset = offset + Symbol.length(name) val pos = Position.Range(Text.Range(offset, end_offset) + 1) Command_Span.Command_Span(name, pos) } result += Command_Span.Span(kind, span) } def flush() { 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))) + (!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/Isar/token.scala b/src/Pure/Isar/token.scala --- a/src/Pure/Isar/token.scala +++ b/src/Pure/Isar/token.scala @@ -1,347 +1,347 @@ /* Title: Pure/Isar/token.scala Author: Makarius Outer token syntax for Isabelle/Isar. */ package isabelle import scala.collection.mutable import scala.util.parsing.input object Token { /* tokens */ object Kind extends Enumeration { /*immediate source*/ val COMMAND = Value("command") val KEYWORD = Value("keyword") val IDENT = Value("identifier") val LONG_IDENT = Value("long identifier") val SYM_IDENT = Value("symbolic identifier") val VAR = Value("schematic variable") val TYPE_IDENT = Value("type variable") val TYPE_VAR = Value("schematic type variable") val NAT = Value("natural number") val FLOAT = Value("floating-point number") val SPACE = Value("white space") /*delimited content*/ val STRING = Value("string") val ALT_STRING = Value("back-quoted string") val VERBATIM = Value("verbatim text") val CARTOUCHE = Value("text cartouche") val INFORMAL_COMMENT = Value("informal comment") val FORMAL_COMMENT = Value("formal comment") /*special content*/ val ERROR = Value("bad input") val UNPARSED = Value("unparsed input") } /* parsers */ object Parsers extends Parsers trait Parsers extends Scan.Parsers with Comment.Parsers { private def delimited_token: Parser[Token] = { val string = quoted("\"") ^^ (x => Token(Token.Kind.STRING, x)) val alt_string = quoted("`") ^^ (x => Token(Token.Kind.ALT_STRING, x)) val verb = verbatim ^^ (x => Token(Token.Kind.VERBATIM, x)) val cart = cartouche ^^ (x => Token(Token.Kind.CARTOUCHE, x)) val cmt = comment ^^ (x => Token(Token.Kind.INFORMAL_COMMENT, x)) val formal_cmt = comment_cartouche ^^ (x => Token(Token.Kind.FORMAL_COMMENT, x)) string | (alt_string | (verb | (cart | (cmt | formal_cmt)))) } private def other_token(keywords: Keyword.Keywords): Parser[Token] = { val letdigs1 = many1(Symbol.is_letdig) val sub = one(s => s == Symbol.sub_decoded || s == Symbol.sub) val id = one(Symbol.is_letter) ~ (rep(letdigs1 | (sub ~ letdigs1 ^^ { case x ~ y => x + y })) ^^ (_.mkString)) ^^ { case x ~ y => x + y } val nat = many1(Symbol.is_digit) val natdot = nat ~ "." ~ nat ^^ { case x ~ y ~ z => x + y + z } val id_nat = id ~ opt("." ~ nat) ^^ { case x ~ Some(y ~ z) => x + y + z case x ~ None => x } val ident = id ~ rep("." ~> id) ^^ { case x ~ Nil => Token(Token.Kind.IDENT, x) case x ~ ys => Token(Token.Kind.LONG_IDENT, (x :: ys).mkString(".")) } val var_ = "?" ~ id_nat ^^ { case x ~ y => Token(Token.Kind.VAR, x + y) } val type_ident = "'" ~ id ^^ { case x ~ y => Token(Token.Kind.TYPE_IDENT, x + y) } val type_var = "?'" ~ id_nat ^^ { case x ~ y => Token(Token.Kind.TYPE_VAR, x + y) } val nat_ = nat ^^ (x => Token(Token.Kind.NAT, x)) val float = ("-" ~ natdot ^^ { case x ~ y => x + y } | natdot) ^^ (x => Token(Token.Kind.FLOAT, x)) val sym_ident = (many1(Symbol.is_symbolic_char) | one(sym => Symbol.is_symbolic(sym))) ^^ (x => Token(Token.Kind.SYM_IDENT, x)) val keyword = literal(keywords.minor) ^^ (x => Token(Token.Kind.KEYWORD, x)) ||| literal(keywords.major) ^^ (x => Token(Token.Kind.COMMAND, x)) val space = many1(Symbol.is_blank) ^^ (x => Token(Token.Kind.SPACE, x)) val recover_delimited = (recover_quoted("\"") | (recover_quoted("`") | (recover_verbatim | (recover_cartouche | recover_comment)))) ^^ (x => Token(Token.Kind.ERROR, x)) val bad = one(_ => true) ^^ (x => Token(Token.Kind.ERROR, x)) space | (recover_delimited | (((ident | (var_ | (type_ident | (type_var | (float | (nat_ | sym_ident)))))) ||| keyword) | bad)) } def token(keywords: Keyword.Keywords): Parser[Token] = delimited_token | other_token(keywords) def token_line(keywords: Keyword.Keywords, ctxt: Scan.Line_Context) : Parser[(Token, Scan.Line_Context)] = { val string = quoted_line("\"", ctxt) ^^ { case (x, c) => (Token(Token.Kind.STRING, x), c) } val alt_string = quoted_line("`", ctxt) ^^ { case (x, c) => (Token(Token.Kind.ALT_STRING, x), c) } val verb = verbatim_line(ctxt) ^^ { case (x, c) => (Token(Token.Kind.VERBATIM, x), c) } val cart = cartouche_line(ctxt) ^^ { case (x, c) => (Token(Token.Kind.CARTOUCHE, x), c) } val cmt = comment_line(ctxt) ^^ { case (x, c) => (Token(Token.Kind.INFORMAL_COMMENT, x), c) } val formal_cmt = comment_cartouche_line(ctxt) ^^ { case (x, c) => (Token(Token.Kind.FORMAL_COMMENT, x), c) } val other = other_token(keywords) ^^ { case x => (x, Scan.Finished) } string | (alt_string | (verb | (cart | (cmt | (formal_cmt | other))))) } } /* explode */ def explode(keywords: Keyword.Keywords, inp: CharSequence): List[Token] = Parsers.parseAll(Parsers.rep(Parsers.token(keywords)), Scan.char_reader(inp)) match { case Parsers.Success(tokens, _) => tokens case _ => error("Unexpected failure of tokenizing input:\n" + inp.toString) } def explode_line(keywords: Keyword.Keywords, inp: CharSequence, context: Scan.Line_Context) : (List[Token], Scan.Line_Context) = { var in: input.Reader[Char] = Scan.char_reader(inp) val toks = new mutable.ListBuffer[Token] var ctxt = context while (!in.atEnd) { Parsers.parse(Parsers.token_line(keywords, ctxt), in) match { case Parsers.Success((x, c), rest) => toks += x; ctxt = c; in = rest case Parsers.NoSuccess(_, rest) => error("Unexpected failure of tokenizing input:\n" + rest.source.toString) } } (toks.toList, ctxt) } val newline: Token = explode(Keyword.Keywords.empty, "\n").head /* embedded */ def read_embedded(keywords: Keyword.Keywords, inp: CharSequence): Option[Token] = explode(keywords, inp) match { case List(tok) if tok.is_embedded => Some(tok) case _ => None } /* names */ def read_name(keywords: Keyword.Keywords, inp: CharSequence): Option[Token] = explode(keywords, inp) match { case List(tok) if tok.is_name => Some(tok) case _ => None } def quote_name(keywords: Keyword.Keywords, name: String): String = if (read_name(keywords, name).isDefined) name else quote(name.replace("\"", "\\\"")) /* plain antiquotation (0 or 1 args) */ def read_antiq_arg(keywords: Keyword.Keywords, inp: CharSequence): Option[(String, Option[String])] = explode(keywords, inp).filter(_.is_proper) match { case List(t) if t.is_name => Some(t.content, None) case List(t1, t2) if t1.is_name && t2.is_embedded => Some(t1.content, Some(t2.content)) case _ => None } /* implode */ def implode(toks: List[Token]): String = toks match { case List(tok) => tok.source case _ => toks.map(_.source).mkString } /* token reader */ object Pos { val none: Pos = new Pos(0, 0, "", "") val start: Pos = new Pos(1, 1, "", "") def file(file: String): Pos = new Pos(1, 1, file, "") def id(id: String): Pos = new Pos(0, 1, "", id) val command: Pos = id(Markup.COMMAND) } final class Pos private[Token]( val line: Int, val offset: Symbol.Offset, val file: String, val id: String) extends input.Position { def column = 0 def lineContents = "" def advance(token: Token): Pos = advance(token.source) def advance(source: String): Pos = { var line1 = line var offset1 = offset for (s <- Symbol.iterator(source)) { if (line1 > 0 && Symbol.is_newline(s)) line1 += 1 if (offset1 > 0) offset1 += 1 } if (line1 == line && offset1 == offset) this else new Pos(line1, offset1, file, id) } private def position(end_offset: Symbol.Offset): Position.T = (if (line > 0) Position.Line(line) else Nil) ::: (if (offset > 0) Position.Offset(offset) else Nil) ::: (if (end_offset > 0) Position.End_Offset(end_offset) else Nil) ::: (if (file != "") Position.File(file) else Nil) ::: (if (id != "") Position.Id_String(id) else Nil) def position(): Position.T = position(0) def position(token: Token): Position.T = position(advance(token).offset) def position(source: String): Position.T = position(advance(source).offset) override def toString: String = Position.here(position(), delimited = false) } abstract class Reader extends input.Reader[Token] private class Token_Reader(tokens: List[Token], val pos: Pos) extends Reader { - def first = tokens.head - def rest = new Token_Reader(tokens.tail, pos.advance(first)) - def atEnd = tokens.isEmpty + def first: Token = tokens.head + def rest: Token_Reader = new Token_Reader(tokens.tail, pos.advance(first)) + def atEnd: Boolean = tokens.isEmpty } def reader(tokens: List[Token], start: Token.Pos): Reader = new Token_Reader(tokens, start) /* XML data representation */ val encode: XML.Encode.T[Token] = (tok: Token) => { import XML.Encode._ pair(int, string)(tok.kind.id, tok.source) } val decode: XML.Decode.T[Token] = (body: XML.Body) => { import XML.Decode._ val (k, s) = pair(int, string)(body) Token(Kind(k), s) } } sealed case class Token(kind: Token.Kind.Value, source: String) { def is_command: Boolean = kind == Token.Kind.COMMAND def is_command(name: String): Boolean = kind == Token.Kind.COMMAND && source == name def is_keyword: Boolean = kind == Token.Kind.KEYWORD def is_keyword(name: String): Boolean = kind == Token.Kind.KEYWORD && source == name def is_keyword(name: Char): Boolean = kind == Token.Kind.KEYWORD && source.length == 1 && source(0) == name def is_delimiter: Boolean = is_keyword && !Symbol.is_ascii_identifier(source) def is_ident: Boolean = kind == Token.Kind.IDENT def is_sym_ident: Boolean = kind == Token.Kind.SYM_IDENT def is_string: Boolean = kind == Token.Kind.STRING def is_nat: Boolean = kind == Token.Kind.NAT def is_float: Boolean = kind == Token.Kind.FLOAT def is_name: Boolean = kind == Token.Kind.IDENT || kind == Token.Kind.LONG_IDENT || kind == Token.Kind.SYM_IDENT || kind == Token.Kind.STRING || kind == Token.Kind.NAT def is_embedded: Boolean = is_name || kind == Token.Kind.CARTOUCHE || kind == Token.Kind.VAR || kind == Token.Kind.TYPE_IDENT || kind == Token.Kind.TYPE_VAR def is_text: Boolean = is_embedded || kind == Token.Kind.VERBATIM def is_space: Boolean = kind == Token.Kind.SPACE def is_informal_comment: Boolean = kind == Token.Kind.INFORMAL_COMMENT def is_formal_comment: Boolean = kind == Token.Kind.FORMAL_COMMENT def is_marker: Boolean = kind == Token.Kind.FORMAL_COMMENT && (source.startsWith(Symbol.marker) || source.startsWith(Symbol.marker_decoded)) def is_comment: Boolean = is_informal_comment || is_formal_comment def is_ignored: Boolean = is_space || is_informal_comment def is_proper: Boolean = !is_space && !is_comment def is_error: Boolean = kind == Token.Kind.ERROR def is_unparsed: Boolean = kind == Token.Kind.UNPARSED def is_unfinished: Boolean = is_error && (source.startsWith("\"") || source.startsWith("`") || source.startsWith("{*") || source.startsWith("(*") || source.startsWith(Symbol.open) || source.startsWith(Symbol.open_decoded)) - def is_open_bracket: Boolean = is_keyword && Word.open_brackets.exists(is_keyword(_)) - def is_close_bracket: Boolean = is_keyword && Word.close_brackets.exists(is_keyword(_)) + def is_open_bracket: Boolean = is_keyword && Word.open_brackets.exists(is_keyword) + def is_close_bracket: Boolean = is_keyword && Word.close_brackets.exists(is_keyword) def is_begin: Boolean = is_keyword("begin") def is_end: Boolean = is_command("end") def is_begin_or_command: Boolean = is_begin || is_command def content: String = if (kind == Token.Kind.STRING) Scan.Parsers.quoted_content("\"", source) else if (kind == Token.Kind.ALT_STRING) Scan.Parsers.quoted_content("`", source) else if (kind == Token.Kind.VERBATIM) Scan.Parsers.verbatim_content(source) else if (kind == Token.Kind.CARTOUCHE) Scan.Parsers.cartouche_content(source) else if (kind == Token.Kind.INFORMAL_COMMENT) Scan.Parsers.comment_content(source) else if (kind == Token.Kind.FORMAL_COMMENT) Comment.content(source) else source def is_system_name: Boolean = { val s = content is_name && Path.is_wellformed(s) && - !s.exists(Symbol.is_ascii_blank(_)) && + !s.exists(Symbol.is_ascii_blank) && !Path.is_reserved(s) } } diff --git a/src/Pure/ML/ml_console.scala b/src/Pure/ML/ml_console.scala --- a/src/Pure/ML/ml_console.scala +++ b/src/Pure/ML/ml_console.scala @@ -1,84 +1,87 @@ /* Title: Pure/ML/ml_console.scala Author: Makarius The raw ML process in interactive mode. */ package isabelle object ML_Console { /* command line entry point */ def main(args: Array[String]) { Command_Line.tool { var dirs: List[Path] = Nil var include_sessions: List[String] = Nil var logic = Isabelle_System.getenv("ISABELLE_LOGIC") var modes: List[String] = Nil var no_build = false var options = Options.init() var raw_ml_system = false val getopts = Getopts(""" Usage: isabelle console [OPTIONS] Options are: -d DIR include session directory -i NAME include session in name-space of theories -l NAME logic session name (default ISABELLE_LOGIC=""" + quote(logic) + """) -m MODE add print mode for output -n no build of session image on startup -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) -r bootstrap from raw Poly/ML Build a logic session image and run the raw Isabelle ML process in interactive mode, with line editor ISABELLE_LINE_EDITOR=""" + quote(Isabelle_System.getenv("ISABELLE_LINE_EDITOR")) + """. """, "d:" -> (arg => dirs = dirs ::: List(Path.explode(arg))), "i:" -> (arg => include_sessions = include_sessions ::: List(arg)), "l:" -> (arg => logic = arg), "m:" -> (arg => modes = arg :: modes), "n" -> (arg => no_build = true), "o:" -> (arg => options = options + arg), "r" -> (_ => raw_ml_system = true)) val more_args = getopts(args) if (more_args.nonEmpty) getopts.usage() + val sessions_structure = Sessions.load_structure(options, dirs = dirs) + val store = Sessions.store(options) + // build logic if (!no_build && !raw_ml_system) { val progress = new Console_Progress() val rc = progress.interrupt_handler { Build.build_logic(options, logic, build_heap = true, progress = progress, dirs = dirs) } if (rc != 0) sys.exit(rc) } // process loop val process = - ML_Process(options, logic = logic, args = List("-i"), dirs = dirs, redirect = true, + ML_Process(options, sessions_structure, store, + logic = logic, args = List("-i"), redirect = true, modes = if (raw_ml_system) Nil else modes ::: List("ASCII"), raw_ml_system = raw_ml_system, - store = Some(Sessions.store(options)), session_base = if (raw_ml_system) None else Some(Sessions.base_info( options, logic, dirs = dirs, include_sessions = include_sessions).check_base)) val tty_loop = new TTY_Loop(process.stdin, process.stdout, Some(process.interrupt _)) val process_result = Future.thread[Int]("process_result") { val rc = process.join tty_loop.cancel // FIXME does not quite work, cannot interrupt blocking read on System.in rc } tty_loop.join process_result.join } } } diff --git a/src/Pure/ML/ml_process.scala b/src/Pure/ML/ml_process.scala --- a/src/Pure/ML/ml_process.scala +++ b/src/Pure/ML/ml_process.scala @@ -1,189 +1,188 @@ /* Title: Pure/ML/ml_process.scala Author: Makarius The raw ML process. */ package isabelle import java.io.{File => JFile} object ML_Process { def apply(options: Options, + sessions_structure: Sessions.Structure, + store: Sessions.Store, logic: String = "", args: List[String] = Nil, - dirs: List[Path] = Nil, modes: List[String] = Nil, raw_ml_system: Boolean = false, cwd: JFile = null, env: Map[String, String] = Isabelle_System.settings(), redirect: Boolean = false, cleanup: () => Unit = () => (), - sessions_structure: Option[Sessions.Structure] = None, - session_base: Option[Sessions.Base] = None, - store: Option[Sessions.Store] = None): Bash.Process = + session_base: Option[Sessions.Base] = None): Bash.Process = { val logic_name = Isabelle_System.default_logic(logic) - val _store = store.getOrElse(Sessions.store(options)) - - val sessions_structure1 = - sessions_structure.getOrElse(Sessions.load_structure(options, dirs = dirs)) val heaps: List[String] = if (raw_ml_system) Nil else { - sessions_structure1.selection(Sessions.Selection.session(logic_name)). + sessions_structure.selection(Sessions.Selection.session(logic_name)). build_requirements(List(logic_name)). - map(a => File.platform_path(_store.the_heap(a))) + map(a => File.platform_path(store.the_heap(a))) } val eval_init = if (heaps.isEmpty) { List( """ fun chapter (_: string) = (); fun section (_: string) = (); fun subsection (_: string) = (); fun subsubsection (_: string) = (); fun paragraph (_: string) = (); fun subparagraph (_: string) = (); val ML_file = PolyML.use; """, if (Platform.is_windows) "fun exit 0 = OS.Process.exit OS.Process.success" + " | exit 1 = OS.Process.exit OS.Process.failure" + " | exit rc = OS.Process.exit (RunCall.unsafeCast (Word8.fromInt rc))" else "fun exit rc = Posix.Process.exit (Word8.fromInt rc)", "PolyML.Compiler.prompt1 := \"Poly/ML> \"", "PolyML.Compiler.prompt2 := \"Poly/ML# \"") } else List( "(PolyML.SaveState.loadHierarchy " + ML_Syntax.print_list(ML_Syntax.print_string_bytes)(heaps) + "; PolyML.print_depth 0) handle exn => (TextIO.output (TextIO.stdErr, General.exnMessage exn ^ " + ML_Syntax.print_string_bytes(": " + logic_name + "\n") + "); OS.Process.exit OS.Process.failure)") val eval_modes = if (modes.isEmpty) Nil else List("Print_Mode.add_modes " + ML_Syntax.print_list(ML_Syntax.print_string_bytes)(modes)) // options val isabelle_process_options = Isabelle_System.tmp_file("options") Isabelle_System.chmod("600", File.path(isabelle_process_options)) File.write(isabelle_process_options, YXML.string_of_body(options.encode)) val env_options = Map("ISABELLE_PROCESS_OPTIONS" -> File.standard_path(isabelle_process_options)) val eval_options = if (heaps.isEmpty) Nil else List("Options.load_default ()") // session base val eval_session_base = session_base match { case None => Nil case Some(base) => def print_table(table: List[(String, String)]): String = ML_Syntax.print_list( ML_Syntax.print_pair( ML_Syntax.print_string_bytes, ML_Syntax.print_string_bytes))(table) def print_list(list: List[String]): String = ML_Syntax.print_list(ML_Syntax.print_string_bytes)(list) def print_sessions(list: List[(String, Position.T)]): String = ML_Syntax.print_list( ML_Syntax.print_pair(ML_Syntax.print_string_bytes, ML_Syntax.print_properties))(list) List("Resources.init_session_base" + - " {session_positions = " + print_sessions(sessions_structure1.session_positions) + - ", session_directories = " + print_table(sessions_structure1.dest_session_directories) + + " {session_positions = " + print_sessions(sessions_structure.session_positions) + + ", session_directories = " + print_table(sessions_structure.dest_session_directories) + ", docs = " + print_list(base.doc_names) + ", global_theories = " + print_table(base.global_theories.toList) + ", loaded_theories = " + print_list(base.loaded_theories.keys) + "}") } // process val eval_process = if (heaps.isEmpty) List("PolyML.print_depth " + ML_Syntax.print_int(options.int("ML_print_depth"))) else List("Isabelle_Process.init ()") // ISABELLE_TMP val isabelle_tmp = Isabelle_System.tmp_dir("process") val env_tmp = Map("ISABELLE_TMP" -> File.standard_path(isabelle_tmp)) val ml_runtime_options = { val ml_options = Word.explode(Isabelle_System.getenv("ML_OPTIONS")) val ml_options1 = if (ml_options.exists(_.containsSlice("gcthreads"))) ml_options else ml_options ::: List("--gcthreads", options.int("threads").toString) val ml_options2 = if (!Platform.is_windows || ml_options.exists(_.containsSlice("codepage"))) ml_options1 else ml_options1 ::: List("--codepage", "utf8") ml_options2 } // bash val bash_args = ml_runtime_options ::: (eval_init ::: eval_modes ::: eval_options ::: eval_session_base ::: eval_process) .flatMap(eval => List("--eval", eval)) ::: args Bash.process( "exec " + options.string("ML_process_policy") + """ "$ML_HOME/poly" -q """ + Bash.strings(bash_args), cwd = cwd, env = env ++ env_options ++ env_tmp, redirect = redirect, cleanup = () => { isabelle_process_options.delete Isabelle_System.rm_tree(isabelle_tmp) cleanup() }) } /* Isabelle tool wrapper */ val isabelle_tool = Isabelle_Tool("process", "raw ML process (batch mode)", args => { var dirs: List[Path] = Nil var eval_args: List[String] = Nil var logic = Isabelle_System.getenv("ISABELLE_LOGIC") var modes: List[String] = Nil var options = Options.init() val getopts = Getopts(""" Usage: isabelle process [OPTIONS] Options are: -T THEORY load theory -d DIR include session directory -e ML_EXPR evaluate ML expression on startup -f ML_FILE evaluate ML file on startup -l NAME logic session name (default ISABELLE_LOGIC=""" + quote(logic) + """) -m MODE add print mode for output -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) Run the raw Isabelle ML process in batch mode. """, "T:" -> (arg => eval_args = eval_args ::: List("--eval", "use_thy " + ML_Syntax.print_string_bytes(arg))), "d:" -> (arg => dirs = dirs ::: List(Path.explode(arg))), "e:" -> (arg => eval_args = eval_args ::: List("--eval", arg)), "f:" -> (arg => eval_args = eval_args ::: List("--use", arg)), "l:" -> (arg => logic = arg), "m:" -> (arg => modes = arg :: modes), "o:" -> (arg => options = options + arg)) val more_args = getopts(args) if (args.isEmpty || more_args.nonEmpty) getopts.usage() - val rc = ML_Process(options, logic = logic, args = eval_args, dirs = dirs, modes = modes). - result().print_stdout.rc + val sessions_structure = Sessions.load_structure(options, dirs = dirs) + val store = Sessions.store(options) + + val rc = + ML_Process(options, sessions_structure, store, logic = logic, args = eval_args, modes = modes) + .result().print_stdout.rc sys.exit(rc) }) } 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,220 +1,220 @@ /* 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 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 all_fields: List[Fields] = List(tasks_fields, workers_fields, GC_fields, heap_fields, program_fields, threads_fields, time_fields, speed_fields) val main_fields: List[Fields] = List(tasks_fields, workers_fields, heap_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_statistics: List[Properties.T], heading: String = "", domain: String => Boolean = (key: String) => true): ML_Statistics = { require(ml_statistics.forall(props => Now.unapply(props).isDefined)) 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 require(time >= 0.0) // 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) }) 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]) { 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 => + 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/ML/ml_syntax.scala b/src/Pure/ML/ml_syntax.scala --- a/src/Pure/ML/ml_syntax.scala +++ b/src/Pure/ML/ml_syntax.scala @@ -1,65 +1,65 @@ /* Title: Pure/ML/ml_syntax.scala Author: Makarius Concrete ML syntax for basic values. */ package isabelle object ML_Syntax { /* int */ private def signed_int(s: String): String = if (s(0) == '-') "~" + s.substring(1) else s def print_int(i: Int): String = signed_int(Value.Int(i)) def print_long(i: Long): String = signed_int(Value.Long(i)) /* string */ private def print_byte(c: Byte): String = c match { case 34 => "\\\"" case 92 => "\\\\" case 9 => "\\t" case 10 => "\\n" case 12 => "\\f" case 13 => "\\r" case _ => if (c < 0) "\\" + Library.signed_string_of_int(256 + c) else if (c < 32) new String(Array[Char](92, 94, (c + 64).toChar)) else if (c < 127) Symbol.ascii(c.toChar) else "\\" + Library.signed_string_of_int(c) } private def print_symbol(s: Symbol.Symbol): String = if (s.startsWith("\\<")) s - else UTF8.bytes(s).iterator.map(print_byte(_)).mkString + else UTF8.bytes(s).iterator.map(print_byte).mkString def print_string_bytes(str: String): String = - quote(UTF8.bytes(str).iterator.map(print_byte(_)).mkString) + quote(UTF8.bytes(str).iterator.map(print_byte).mkString) def print_string_symbols(str: String): String = - quote(Symbol.iterator(str).map(print_symbol(_)).mkString) + quote(Symbol.iterator(str).map(print_symbol).mkString) /* pair */ def print_pair[A, B](f: A => String, g: B => String)(pair: (A, B)): String = "(" + f(pair._1) + ", " + g(pair._2) + ")" /* list */ def print_list[A](f: A => String)(list: List[A]): String = "[" + commas(list.map(f)) + "]" /* properties */ def print_properties(props: Properties.T): String = - print_list(print_pair(print_string_bytes(_), print_string_bytes(_))(_))(props) + print_list(print_pair(print_string_bytes, print_string_bytes))(props) } 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,1199 +1,1199 @@ /* 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: 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() + val no_header: Header = Header() def bad_header(msg: String): Header = Header(errors = List(msg)) object Name { - val empty = 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 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) { 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(_)) (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 { + graph.keys_iterator.filter(is_suppressed).toList match { case Nil => None case del => Some(new Nodes((graph /: del)(_.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)) 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(names).reverse 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)))(_ + _) } } 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 */ object Snapshot { - val init = State.init.snapshot() + val init: Snapshot = State.init.snapshot() } abstract class Snapshot { def state: State def version: Version def is_outdated: Boolean def convert(i: Text.Offset): Text.Offset def revert(i: Text.Offset): Text.Offset def convert(range: Text.Range): Text.Range def revert(range: Text.Range): Text.Range def node_name: Node.Name def node: Node def nodes: List[(Node.Name, Node)] def commands_loading: List[Command] def commands_loading_ranges(pred: Node.Name => Boolean): List[Text.Range] def current_command(other_node_name: Node.Name, offset: Text.Offset): Option[Command] def markup_to_XML(range: Text.Range, elements: Markup.Elements): XML.Body def messages: List[(XML.Tree, Position.T)] def exports: List[Export.Entry] def exports_map: Map[String, Export.Entry] def find_command(id: Document_ID.Generic): Option[(Node, Command)] def find_command_position(id: Document_ID.Generic, offset: Symbol.Offset) : Option[Line.Node_Position] 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]] 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 command_results(range: Text.Range): Command.Results def command_results(command: Command): Command.Results def command_id_map: Map[Document_ID.Generic, Command] } /* 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); this } def unfinished: Assignment = new Assignment(command_execs, false) def assign(update: Assign_Update): Assignment = { require(!is_finished) val command_execs1 = (command_execs /: update) { 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, /*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 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] = 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(id: Document_ID.Generic) : Option[(Symbol.Text_Chunk.Id, Symbol.Text_Chunk)] = lookup_id(id).map(st => ((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) }) def accumulate(id: Document_ID.Generic, message: XML.Elem, xml_cache: XML.Cache) : (Command.State, State) = { execs.get(id) match { case Some(st) => - val new_st = st.accumulate(self_id(st), other_id _, message, xml_cache) + val new_st = st.accumulate(self_id(st), other_id, message, xml_cache) val execs1 = execs + (id -> new_st) (new_st, copy(execs = execs1, commands_redirection = redirection(new_st))) case None => commands.get(id) match { case Some(st) => - val new_st = st.accumulate(self_id(st), other_id _, message, xml_cache) + val new_st = st.accumulate(self_id(st), other_id, message, xml_cache) val commands1 = commands + (id -> new_st) (new_st, copy(commands = commands1, commands_redirection = redirection(new_st))) 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 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) { 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)) 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)) 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)) 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 markup_to_XML( version: Version, node_name: Node.Name, range: Text.Range, elements: Markup.Elements): 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) } // persistent user-view def snapshot(name: Node.Name = Node.Name.empty, pending_edits: List[Text.Edit] = Nil) : Snapshot = { val stable = recent_stable val latest = history.tip /* pending edits and unstable changes */ val rev_pending_changes = for { change <- history.undo_list.takeWhile(_ != stable) (a, edits) <- change.rev_edits if a == name } yield edits val edits = (pending_edits /: rev_pending_changes)({ case (edits, Node.Edits(es)) => es ::: edits case (edits, _) => edits }) lazy val reverse_edits = edits.reverse new Snapshot { /* global information */ val state: State = State.this val version: Version = stable.version.get_finished val is_outdated: Boolean = pending_edits.nonEmpty || latest != stable /* local node content */ - def convert(offset: Text.Offset) = (offset /: edits)((i, edit) => edit.convert(i)) - def revert(offset: Text.Offset) = (offset /: reverse_edits)((i, edit) => edit.revert(i)) + def convert(offset: Text.Offset): Text.Offset = (offset /: edits)((i, edit) => edit.convert(i)) + def revert(offset: Text.Offset): Text.Offset = (offset /: reverse_edits)((i, edit) => edit.revert(i)) - def convert(range: Text.Range) = range.map(convert(_)) - def revert(range: Text.Range) = range.map(revert(_)) + def convert(range: Text.Range): Text.Range = range.map(convert) + def revert(range: Text.Range): Text.Range = range.map(revert) val node_name: Node.Name = name val node: Node = version.nodes(name) def nodes: List[(Node.Name, Node)] = (node_name :: node.load_commands.flatMap(_.blobs_names)). map(name => (name, version.nodes(name))) 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 def current_command(other_node_name: Node.Name, offset: Text.Offset): Option[Command] = if (other_node_name.is_theory) { val other_node = version.nodes(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 def markup_to_XML(range: Text.Range, elements: Markup.Elements): XML.Body = state.markup_to_XML(version, node_name, range, elements) lazy val messages: List[(XML.Tree, 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) (_, tree) <- state.command_results(version, command).iterator } yield (tree, pos)).toList 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 node = version.nodes(command.node_name) if (node.commands.contains(command)) Some((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(_)) Line.Node_Position(name, pos) } /* 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)) + for (Text.Info(r, Some(x)) <- cumulate(range, None, elements, result1, status)) yield Text.Info(r, x) } /* 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, version.nodes(node_name).commands) /* output */ override def toString: String = "Snapshot(node = " + node_name.node + ", version = " + version.id + (if (is_outdated) ", outdated" else "") + ")" } } } } diff --git a/src/Pure/PIDE/document_id.scala b/src/Pure/PIDE/document_id.scala --- a/src/Pure/PIDE/document_id.scala +++ b/src/Pure/PIDE/document_id.scala @@ -1,24 +1,24 @@ /* Title: Pure/PIDE/document_id.scala Author: Makarius Unique identifiers for document structure. NB: ML ticks forwards > 0, JVM ticks backwards < 0. */ package isabelle object Document_ID { type Generic = Long type Version = Generic type Command = Generic type Exec = Generic val none: Generic = 0 - val make = Counter.make() + val make: Counter = Counter.make() def apply(id: Generic): String = Value.Long.apply(id) def unapply(s: String): Option[Generic] = Value.Long.unapply(s) } 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 = make(Iterator.empty) + 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)(_ + _) } 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)({ 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,677 +1,678 @@ /* 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(pending1).reverse.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) 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 }) } 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(_)) + 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 = No_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 = dependencies.loaded_files(false).flatMap(_._2). map(path => 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) { 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 = Standard_Thread.delay_first(nodes_status_delay max Time.zero) { progress.nodes_status(use_theories_state.value.nodes_status) } val delay_commit_clean = Standard_Thread.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(_)) + 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(base_info: Sessions.Base_Info, log: Logger = No_Logger): Resources = - new Resources(base_info, log = log) + 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 = No_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(base_info, log = log) + 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) }) (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(_)))) + 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))) def remove_required(id: UUID.T, names: List[Document.Node.Name]): State = copy(required = (required /: names)(_.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) } })) def remove_theories(remove: List[Document.Node.Name]): State = { require(remove.forall(name => !is_required(name))) 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 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 => - def options: Options = session_base_info.options + val store: Sessions.Store = Sessions.store(options) /* session */ def start_session(print_mode: List[String] = Nil, progress: Progress = No_Progress): Session = { val session = new Session(session_base_info.session, options, resources) val session_error = Future.promise[String] var session_phase: Session.Consumer[Session.Phase] = null session_phase = Session.Consumer(getClass.getName) { case Session.Ready => session.phase_changed -= session_phase session_error.fulfill("") case Session.Terminated(result) if !result.ok => session.phase_changed -= session_phase session_error.fulfill("Session start failed: return code " + result.rc) case _ => } session.phase_changed += session_phase progress.echo("Starting session " + session_base_info.session + " ...") - Isabelle_Process.start(session, options, - logic = session_base_info.session, dirs = session_base_info.dirs, modes = print_mode) + Isabelle_Process(session, options, session_base_info.sessions_structure, store, + logic = session_base_info.session, modes = print_mode) session_error.join match { case "" => session case msg => session.stop(); error(msg) } } /* 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) { 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_reader(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]) { 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]) { 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/markup.scala b/src/Pure/PIDE/markup.scala --- a/src/Pure/PIDE/markup.scala +++ b/src/Pure/PIDE/markup.scala @@ -1,753 +1,753 @@ /* Title: Pure/PIDE/markup.scala Author: Makarius Quasi-abstract markup elements. */ package isabelle object Markup { /* elements */ object Elements { def apply(elems: Set[String]): Elements = new Elements(elems) def apply(elems: String*): Elements = apply(Set(elems: _*)) val empty: Elements = apply() val full: Elements = new Elements(Set.empty) { override def apply(elem: String): Boolean = true override def toString: String = "Elements.full" } } sealed class Elements private[Markup](private val rep: Set[String]) { def apply(elem: String): Boolean = rep.contains(elem) def + (elem: String): Elements = new Elements(rep + elem) def ++ (elems: Elements): Elements = new Elements(rep ++ elems.rep) def - (elem: String): Elements = new Elements(rep - elem) def -- (elems: Elements): Elements = new Elements(rep -- elems.rep) override def toString: String = rep.mkString("Elements(", ",", ")") } /* properties */ val NAME = "name" val Name = new Properties.String(NAME) val XNAME = "xname" val XName = new Properties.String(XNAME) val KIND = "kind" val Kind = new Properties.String(KIND) val CONTENT = "content" val Content = new Properties.String(CONTENT) val SERIAL = "serial" val Serial = new Properties.Long(SERIAL) val INSTANCE = "instance" val Instance = new Properties.String(INSTANCE) /* basic markup */ - val Empty = Markup("", Nil) - val Broken = Markup("broken", Nil) + val Empty: Markup = Markup("", Nil) + val Broken: Markup = Markup("broken", Nil) class Markup_String(val name: String, prop: String) { private val Prop = new Properties.String(prop) def apply(s: String): Markup = Markup(name, Prop(s)) def unapply(markup: Markup): Option[String] = if (markup.name == name) Prop.unapply(markup.properties) else None } class Markup_Int(val name: String, prop: String) { private val Prop = new Properties.Int(prop) def apply(i: Int): Markup = Markup(name, Prop(i)) def unapply(markup: Markup): Option[Int] = if (markup.name == name) Prop.unapply(markup.properties) else None } class Markup_Long(val name: String, prop: String) { private val Prop = new Properties.Long(prop) def apply(i: Long): Markup = Markup(name, Prop(i)) def unapply(markup: Markup): Option[Long] = if (markup.name == name) Prop.unapply(markup.properties) else None } /* meta data */ val META_TITLE = "meta_title" val META_CREATOR = "meta_creator" val META_CONTRIBUTOR = "meta_contributor" val META_DATE = "meta_date" val META_LICENSE = "meta_license" val META_DESCRIPTION = "meta_description" /* formal entities */ val BINDING = "binding" val ENTITY = "entity" object Entity { val Def = new Properties.Long("def") val Ref = new Properties.Long("ref") def unapply(markup: Markup): Option[(String, String)] = markup match { case Markup(ENTITY, props) => val kind = Kind.unapply(props).getOrElse("") val name = Name.unapply(props).getOrElse("") Some((kind, name)) case _ => None } } /* completion */ val COMPLETION = "completion" val NO_COMPLETION = "no_completion" val UPDATE = "update" /* position */ val LINE = "line" val END_LINE = "line" val OFFSET = "offset" val END_OFFSET = "end_offset" val FILE = "file" val ID = "id" val DEF_LINE = "def_line" val DEF_OFFSET = "def_offset" val DEF_END_OFFSET = "def_end_offset" val DEF_FILE = "def_file" val DEF_ID = "def_id" val POSITION_PROPERTIES = Set(LINE, OFFSET, END_OFFSET, FILE, ID) val POSITION = "position" /* expression */ val EXPRESSION = "expression" object Expression { def unapply(markup: Markup): Option[String] = markup match { case Markup(EXPRESSION, Kind(kind)) => Some(kind) case Markup(EXPRESSION, _) => Some("") case _ => None } } /* citation */ val CITATION = "citation" val Citation = new Markup_String(CITATION, NAME) /* embedded languages */ val Symbols = new Properties.Boolean("symbols") val Antiquotes = new Properties.Boolean("antiquotes") val Delimited = new Properties.Boolean("delimited") val LANGUAGE = "language" object Language { val DOCUMENT = "document" val ML = "ML" val SML = "SML" val PATH = "path" val UNKNOWN = "unknown" def unapply(markup: Markup): Option[(String, Boolean, Boolean, Boolean)] = markup match { case Markup(LANGUAGE, props) => (props, props, props, props) match { case (Name(name), Symbols(symbols), Antiquotes(antiquotes), Delimited(delimited)) => Some((name, symbols, antiquotes, delimited)) case _ => None } case _ => None } } /* external resources */ val PATH = "path" val Path = new Markup_String(PATH, NAME) val EXPORT_PATH = "export_path" val Export_Path = new Markup_String(EXPORT_PATH, NAME) val URL = "url" val Url = new Markup_String(URL, NAME) val DOC = "doc" val Doc = new Markup_String(DOC, NAME) /* pretty printing */ val Consistent = new Properties.Boolean("consistent") val Indent = new Properties.Int("indent") val Width = new Properties.Int("width") object Block { val name = "block" def apply(c: Boolean, i: Int): Markup = Markup(name, (if (c) Consistent(c) else Nil) ::: (if (i != 0) Indent(i) else Nil)) def unapply(markup: Markup): Option[(Boolean, Int)] = if (markup.name == name) { val c = Consistent.unapply(markup.properties).getOrElse(false) val i = Indent.unapply(markup.properties).getOrElse(0) Some((c, i)) } else None } object Break { val name = "break" def apply(w: Int, i: Int): Markup = Markup(name, (if (w != 0) Width(w) else Nil) ::: (if (i != 0) Indent(i) else Nil)) def unapply(markup: Markup): Option[(Int, Int)] = if (markup.name == name) { val w = Width.unapply(markup.properties).getOrElse(0) val i = Indent.unapply(markup.properties).getOrElse(0) Some((w, i)) } else None } val ITEM = "item" val BULLET = "bullet" val SEPARATOR = "separator" /* text properties */ val WORDS = "words" val HIDDEN = "hidden" val DELETE = "delete" /* misc entities */ val CLASS = "class" val TYPE_NAME = "type_name" val FIXED = "fixed" val CASE = "case" val CONSTANT = "constant" val DYNAMIC_FACT = "dynamic_fact" /* inner syntax */ val TFREE = "tfree" val TVAR = "tvar" val FREE = "free" val SKOLEM = "skolem" val BOUND = "bound" val VAR = "var" val NUMERAL = "numeral" val LITERAL = "literal" val DELIMITER = "delimiter" val INNER_STRING = "inner_string" val INNER_CARTOUCHE = "inner_cartouche" val TOKEN_RANGE = "token_range" val SORTING = "sorting" val TYPING = "typing" val CLASS_PARAMETER = "class_parameter" val ATTRIBUTE = "attribute" val METHOD = "method" /* antiquotations */ val ANTIQUOTED = "antiquoted" val ANTIQUOTE = "antiquote" val ML_ANTIQUOTATION = "ML_antiquotation" val DOCUMENT_ANTIQUOTATION = "document_antiquotation" val DOCUMENT_ANTIQUOTATION_OPTION = "document_antiquotation_option" /* document text */ val RAW_TEXT = "raw_text" val PLAIN_TEXT = "plain_text" val PARAGRAPH = "paragraph" val TEXT_FOLD = "text_fold" object Document_Tag { val ELEMENT = "document_tag" val IMPORTANT = "important" val UNIMPORTANT = "unimportant" def unapply(markup: Markup): Option[String] = markup match { case Markup(ELEMENT, Name(name)) => Some(name) case _ => None } } /* Markdown document structure */ val MARKDOWN_PARAGRAPH = "markdown_paragraph" val MARKDOWN_ITEM = "markdown_item" val Markdown_Bullet = new Markup_Int("markdown_bullet", "depth") val Markdown_List = new Markup_String("markdown_list", "kind") val ITEMIZE = "itemize" val ENUMERATE = "enumerate" val DESCRIPTION = "description" /* ML */ val ML_KEYWORD1 = "ML_keyword1" val ML_KEYWORD2 = "ML_keyword2" val ML_KEYWORD3 = "ML_keyword3" val ML_DELIMITER = "ML_delimiter" val ML_TVAR = "ML_tvar" val ML_NUMERAL = "ML_numeral" val ML_CHAR = "ML_char" val ML_STRING = "ML_string" val ML_COMMENT = "ML_comment" val ML_DEF = "ML_def" val ML_OPEN = "ML_open" val ML_STRUCTURE = "ML_structure" val ML_TYPING = "ML_typing" val ML_BREAKPOINT = "ML_breakpoint" /* outer syntax */ val COMMAND = "command" val KEYWORD = "keyword" val KEYWORD1 = "keyword1" val KEYWORD2 = "keyword2" val KEYWORD3 = "keyword3" val QUASI_KEYWORD = "quasi_keyword" val IMPROPER = "improper" val OPERATOR = "operator" val STRING = "string" val ALT_STRING = "alt_string" val VERBATIM = "verbatim" val CARTOUCHE = "cartouche" val COMMENT = "comment" /* comments */ val COMMENT1 = "comment1" val COMMENT2 = "comment2" val COMMENT3 = "comment3" /* timing */ val Elapsed = new Properties.Double("elapsed") val CPU = new Properties.Double("cpu") val GC = new Properties.Double("gc") object Timing_Properties { def apply(timing: isabelle.Timing): Properties.T = Elapsed(timing.elapsed.seconds) ::: CPU(timing.cpu.seconds) ::: GC(timing.gc.seconds) def unapply(props: Properties.T): Option[isabelle.Timing] = (props, props, props) match { case (Elapsed(elapsed), CPU(cpu), GC(gc)) => Some(new isabelle.Timing(Time.seconds(elapsed), Time.seconds(cpu), Time.seconds(gc))) case _ => None } } val TIMING = "timing" object Timing { def apply(timing: isabelle.Timing): Markup = Markup(TIMING, Timing_Properties(timing)) def unapply(markup: Markup): Option[isabelle.Timing] = markup match { case Markup(TIMING, Timing_Properties(timing)) => Some(timing) case _ => None } } /* process result */ val Return_Code = new Properties.Int("return_code") object Process_Result { def apply(result: Process_Result): Properties.T = Return_Code(result.rc) ::: (if (result.timing.is_zero) Nil else Timing_Properties(result.timing)) def unapply(props: Properties.T): Option[Process_Result] = props match { case Return_Code(rc) => val timing = Timing_Properties.unapply(props).getOrElse(isabelle.Timing.zero) Some(isabelle.Process_Result(rc, timing = timing)) case _ => None } } /* command indentation */ object Command_Indent { val name = "command_indent" def unapply(markup: Markup): Option[Int] = if (markup.name == name) Indent.unapply(markup.properties) else None } /* goals */ val GOAL = "goal" val SUBGOAL = "subgoal" /* command status */ val TASK = "task" val ACCEPTED = "accepted" val FORKED = "forked" val JOINED = "joined" val RUNNING = "running" val FINISHED = "finished" val FAILED = "failed" val CANCELED = "canceled" val INITIALIZED = "initialized" val FINALIZED = "finalized" val CONSOLIDATING = "consolidating" val CONSOLIDATED = "consolidated" /* interactive documents */ val VERSION = "version" val ASSIGN = "assign" /* prover process */ val PROVER_COMMAND = "prover_command" val PROVER_ARG = "prover_arg" /* messages */ val INIT = "init" val STATUS = "status" val REPORT = "report" val RESULT = "result" val WRITELN = "writeln" val STATE = "state" val INFORMATION = "information" val TRACING = "tracing" val WARNING = "warning" val LEGACY = "legacy" val ERROR = "error" val NODES_STATUS = "nodes_status" val PROTOCOL = "protocol" val SYSTEM = "system" val STDOUT = "stdout" val STDERR = "stderr" val EXIT = "exit" val WRITELN_MESSAGE = "writeln_message" val STATE_MESSAGE = "state_message" val INFORMATION_MESSAGE = "information_message" val TRACING_MESSAGE = "tracing_message" val WARNING_MESSAGE = "warning_message" val LEGACY_MESSAGE = "legacy_message" val ERROR_MESSAGE = "error_message" val messages = Map( WRITELN -> WRITELN_MESSAGE, STATE -> STATE_MESSAGE, INFORMATION -> INFORMATION_MESSAGE, TRACING -> TRACING_MESSAGE, WARNING -> WARNING_MESSAGE, LEGACY -> LEGACY_MESSAGE, ERROR -> ERROR_MESSAGE) val message: String => String = messages.withDefault((s: String) => s) val NO_REPORT = "no_report" val BAD = "bad" val INTENSIFY = "intensify" /* active areas */ val BROWSER = "browser" val GRAPHVIEW = "graphview" val THEORY_EXPORTS = "theory_exports" val SENDBACK = "sendback" val PADDING = "padding" val PADDING_LINE = (PADDING, "line") val PADDING_COMMAND = (PADDING, "command") val DIALOG = "dialog" val Result = new Properties.String(RESULT) val JEDIT_ACTION = "jedit_action" /* protocol message functions */ val FUNCTION = "function" val Function = new Properties.String(FUNCTION) val COMMAND_TIMING: Properties.Entry = (FUNCTION, "command_timing") val THEORY_TIMING: Properties.Entry = (FUNCTION, "theory_timing") val Commands_Accepted: Properties.T = List((FUNCTION, "commands_accepted")) val Assign_Update: Properties.T = List((FUNCTION, "assign_update")) val Removed_Versions: Properties.T = List((FUNCTION, "removed_versions")) object Protocol_Handler { def unapply(props: Properties.T): Option[(String)] = props match { case List((FUNCTION, "protocol_handler"), (NAME, name)) => Some(name) case _ => None } } val INVOKE_SCALA = "invoke_scala" object Invoke_Scala { def unapply(props: Properties.T): Option[(String, String)] = props match { case List((FUNCTION, INVOKE_SCALA), (NAME, name), (ID, id)) => Some((name, id)) case _ => None } } val CANCEL_SCALA = "cancel_scala" object Cancel_Scala { def unapply(props: Properties.T): Option[String] = props match { case List((FUNCTION, CANCEL_SCALA), (ID, id)) => Some(id) case _ => None } } object ML_Statistics { def unapply(props: Properties.T): Option[Properties.T] = props match { case (FUNCTION, "ML_statistics") :: props => Some(props) case _ => None } } object Task_Statistics { def unapply(props: Properties.T): Option[Properties.T] = props match { case (FUNCTION, "task_statistics") :: props => Some(props) case _ => None } } val LOADING_THEORY = "loading_theory" object Loading_Theory { def unapply(props: Properties.T): Option[String] = props match { case List((FUNCTION, LOADING_THEORY), (NAME, name)) => Some(name) case _ => None } } val BUILD_SESSION_FINISHED = "build_session_finished" val Build_Session_Finished: Properties.T = List((FUNCTION, BUILD_SESSION_FINISHED)) val PRINT_OPERATIONS = "print_operations" /* export */ val EXPORT = "export" object Export { sealed case class Args( id: Option[String], serial: Long, theory_name: String, name: String, executable: Boolean, compress: Boolean, strict: Boolean) { def compound_name: String = isabelle.Export.compound_name(theory_name, name) } val THEORY_NAME = "theory_name" val EXECUTABLE = "executable" val COMPRESS = "compress" val STRICT = "strict" def dest_inline(props: Properties.T): Option[(Args, Path)] = props match { case List( (SERIAL, Value.Long(serial)), (THEORY_NAME, theory_name), (NAME, name), (EXECUTABLE, Value.Boolean(executable)), (COMPRESS, Value.Boolean(compress)), (STRICT, Value.Boolean(strict)), (FILE, file)) if isabelle.Path.is_valid(file) => val args = Args(None, serial, theory_name, name, executable, compress, strict) Some((args, isabelle.Path.explode(file))) case _ => None } def unapply(props: Properties.T): Option[Args] = props match { case List( (FUNCTION, EXPORT), (ID, id), (SERIAL, Value.Long(serial)), (THEORY_NAME, theory_name), (NAME, name), (EXECUTABLE, Value.Boolean(executable)), (COMPRESS, Value.Boolean(compress)), (STRICT, Value.Boolean(strict))) => Some(Args(proper_string(id), serial, theory_name, name, executable, compress, strict)) case _ => None } } /* debugger output */ val DEBUGGER_STATE = "debugger_state" object Debugger_State { def unapply(props: Properties.T): Option[String] = props match { case List((FUNCTION, DEBUGGER_STATE), (NAME, name)) => Some(name) case _ => None } } val DEBUGGER_OUTPUT = "debugger_output" object Debugger_Output { def unapply(props: Properties.T): Option[String] = props match { case List((FUNCTION, DEBUGGER_OUTPUT), (NAME, name)) => Some(name) case _ => None } } /* simplifier trace */ val SIMP_TRACE_PANEL = "simp_trace_panel" val SIMP_TRACE_LOG = "simp_trace_log" val SIMP_TRACE_STEP = "simp_trace_step" val SIMP_TRACE_RECURSE = "simp_trace_recurse" val SIMP_TRACE_HINT = "simp_trace_hint" val SIMP_TRACE_IGNORE = "simp_trace_ignore" val SIMP_TRACE_CANCEL = "simp_trace_cancel" object Simp_Trace_Cancel { def unapply(props: Properties.T): Option[Long] = props match { case (FUNCTION, SIMP_TRACE_CANCEL) :: Serial(i) => Some(i) case _ => None } } /* XML data representation */ def encode: XML.Encode.T[Markup] = (markup: Markup) => { import XML.Encode._ pair(string, properties)((markup.name, markup.properties)) } def decode: XML.Decode.T[Markup] = (body: XML.Body) => { import XML.Decode._ val (name, props) = pair(string, properties)(body) Markup(name, props) } } sealed case class Markup(name: String, properties: Properties.T) { def markup(s: String): String = YXML.string_of_tree(XML.Elem(this, List(XML.Text(s)))) def update_properties(more_props: Properties.T): Markup = if (more_props.isEmpty) this else Markup(name, (more_props :\ properties) { case (p, ps) => Properties.put(ps, p) }) def + (entry: Properties.Entry): Markup = Markup(name, Properties.put(properties, entry)) } diff --git a/src/Pure/PIDE/protocol.scala b/src/Pure/PIDE/protocol.scala --- a/src/Pure/PIDE/protocol.scala +++ b/src/Pure/PIDE/protocol.scala @@ -1,381 +1,381 @@ /* Title: Pure/PIDE/protocol.scala Author: Makarius Protocol message formats for interactive proof documents. */ package isabelle object Protocol { /* document editing */ object Commands_Accepted { def unapply(text: String): Option[List[Document_ID.Command]] = try { Some(space_explode(',', text).map(Value.Long.parse)) } catch { case ERROR(_) => None } val message: XML.Elem = XML.elem(Markup.STATUS, List(XML.elem(Markup.ACCEPTED))) } object Assign_Update { def unapply(text: String) : Option[(Document_ID.Version, List[String], Document.Assign_Update)] = { try { import XML.Decode._ def decode_upd(body: XML.Body): (Long, List[Long]) = space_explode(',', string(body)).map(Value.Long.parse) match { case a :: bs => (a, bs) case _ => throw new XML.XML_Body(body) } - Some(triple(long, list(string), list(decode_upd _))(Symbol.decode_yxml(text))) + Some(triple(long, list(string), list(decode_upd))(Symbol.decode_yxml(text))) } catch { case ERROR(_) => None case _: XML.Error => None } } } object Removed { def unapply(text: String): Option[List[Document_ID.Version]] = try { import XML.Decode._ Some(list(long)(Symbol.decode_yxml(text))) } catch { case ERROR(_) => None case _: XML.Error => None } } /* command timing */ object Command_Timing { def unapply(props: Properties.T): Option[(Document_ID.Generic, isabelle.Timing)] = props match { case Markup.COMMAND_TIMING :: args => (args, args) match { case (Position.Id(id), Markup.Timing_Properties(timing)) => Some((id, timing)) case _ => None } case _ => None } } /* theory timing */ object Theory_Timing { def unapply(props: Properties.T): Option[(String, isabelle.Timing)] = props match { case Markup.THEORY_TIMING :: args => (args, args) match { case (Markup.Name(name), Markup.Timing_Properties(timing)) => Some((name, timing)) case _ => None } case _ => None } } /* result messages */ def is_result(msg: XML.Tree): Boolean = msg match { case XML.Elem(Markup(Markup.RESULT, _), _) => true case _ => false } def is_tracing(msg: XML.Tree): Boolean = msg match { case XML.Elem(Markup(Markup.TRACING, _), _) => true case XML.Elem(Markup(Markup.TRACING_MESSAGE, _), _) => true case _ => false } def is_state(msg: XML.Tree): Boolean = msg match { case XML.Elem(Markup(Markup.STATE, _), _) => true case XML.Elem(Markup(Markup.STATE_MESSAGE, _), _) => true case _ => false } def is_information(msg: XML.Tree): Boolean = msg match { case XML.Elem(Markup(Markup.INFORMATION, _), _) => true case XML.Elem(Markup(Markup.INFORMATION_MESSAGE, _), _) => true case _ => false } def is_writeln(msg: XML.Tree): Boolean = msg match { case XML.Elem(Markup(Markup.WRITELN, _), _) => true case XML.Elem(Markup(Markup.WRITELN_MESSAGE, _), _) => true case _ => false } def is_warning(msg: XML.Tree): Boolean = msg match { case XML.Elem(Markup(Markup.WARNING, _), _) => true case XML.Elem(Markup(Markup.WARNING_MESSAGE, _), _) => true case _ => false } def is_legacy(msg: XML.Tree): Boolean = msg match { case XML.Elem(Markup(Markup.LEGACY, _), _) => true case XML.Elem(Markup(Markup.LEGACY_MESSAGE, _), _) => true case _ => false } def is_error(msg: XML.Tree): Boolean = msg match { case XML.Elem(Markup(Markup.ERROR, _), _) => true case XML.Elem(Markup(Markup.ERROR_MESSAGE, _), _) => true case _ => false } def is_inlined(msg: XML.Tree): Boolean = !(is_result(msg) || is_tracing(msg) || is_state(msg)) def is_exported(msg: XML.Tree): Boolean = is_writeln(msg) || is_warning(msg) || is_legacy(msg) || is_error(msg) def message_text(msg: XML.Tree): String = { val text = Pretty.string_of(List(msg)) if (is_warning(msg) || is_legacy(msg)) Library.prefix_lines("### ", text) else if (is_error(msg)) Library.prefix_lines("*** ", text) else text } /* breakpoints */ object ML_Breakpoint { def unapply(tree: XML.Tree): Option[Long] = tree match { case XML.Elem(Markup(Markup.ML_BREAKPOINT, Markup.Serial(breakpoint)), _) => Some(breakpoint) case _ => None } } /* dialogs */ object Dialog_Args { def unapply(props: Properties.T): Option[(Document_ID.Generic, Long, String)] = (props, props, props) match { case (Position.Id(id), Markup.Serial(serial), Markup.Result(result)) => Some((id, serial, result)) case _ => None } } object Dialog { def unapply(tree: XML.Tree): Option[(Document_ID.Generic, Long, String)] = tree match { case XML.Elem(Markup(Markup.DIALOG, Dialog_Args(id, serial, result)), _) => Some((id, serial, result)) case _ => None } } object Dialog_Result { def apply(id: Document_ID.Generic, serial: Long, result: String): XML.Elem = { val props = Position.Id(id) ::: Markup.Serial(serial) XML.Elem(Markup(Markup.RESULT, props), List(XML.Text(result))) } def unapply(tree: XML.Tree): Option[String] = tree match { case XML.Elem(Markup(Markup.RESULT, _), List(XML.Text(result))) => Some(result) case _ => None } } } trait Protocol { /* protocol commands */ def protocol_command_raw(name: String, args: List[Bytes]): Unit def protocol_command_args(name: String, args: List[String]) def protocol_command(name: String, args: String*): Unit /* options */ def options(opts: Options): Unit = protocol_command("Prover.options", Symbol.encode_yxml(opts.encode)) /* session base */ private def encode_table(table: List[(String, String)]): String = { import XML.Encode._ Symbol.encode_yxml(list(pair(string, string))(table)) } private def encode_list(lst: List[String]): String = { import XML.Encode._ Symbol.encode_yxml(list(string)(lst)) } private def encode_sessions(lst: List[(String, Position.T)]): String = { import XML.Encode._ Symbol.encode_yxml(list(pair(string, properties))(lst)) } def session_base(resources: Resources) { protocol_command("Prover.init_session_base", encode_sessions(resources.sessions_structure.session_positions), encode_table(resources.sessions_structure.dest_session_directories), encode_list(resources.session_base.doc_names), encode_table(resources.session_base.global_theories.toList), encode_list(resources.session_base.loaded_theories.keys)) } /* interned items */ def define_blob(digest: SHA1.Digest, bytes: Bytes): Unit = protocol_command_raw("Document.define_blob", List(Bytes(digest.toString), bytes)) private def encode_command(command: Command): (String, String, String, String, List[String]) = { import XML.Encode._ val blobs_yxml = { val encode_blob: T[Command.Blob] = variant(List( { case Exn.Res((a, b)) => (Nil, pair(string, option(string))((a.node, b.map(p => p._1.toString)))) }, { case Exn.Exn(e) => (Nil, string(Exn.message(e))) })) Symbol.encode_yxml(pair(list(encode_blob), int)(command.blobs, command.blobs_index)) } val toks_yxml = { val encode_tok: T[Token] = (tok => pair(int, int)((tok.kind.id, Symbol.length(tok.source)))) Symbol.encode_yxml(list(encode_tok)(command.span.content)) } val toks_sources = command.span.content.map(tok => Symbol.encode(tok.source)) (Document_ID(command.id), Symbol.encode(command.span.name), blobs_yxml, toks_yxml, toks_sources) } def define_command(command: Command) { val (command_id, name, blobs_yxml, toks_yxml, toks_sources) = encode_command(command) protocol_command_args( "Document.define_command", command_id :: name :: blobs_yxml :: toks_yxml :: toks_sources) } def define_commands(commands: List[Command]) { protocol_command_args("Document.define_commands", commands.map(command => { import XML.Encode._ val (command_id, name, blobs_yxml, toks_yxml, toks_sources) = encode_command(command) val body = pair(string, pair(string, pair(string, pair(string, list(string)))))( command_id, (name, (blobs_yxml, (toks_yxml, toks_sources)))) YXML.string_of_body(body) })) } def define_commands_bulk(commands: List[Command]) { val (irregular, regular) = commands.partition(command => YXML.detect(command.source)) - irregular.foreach(define_command(_)) + irregular.foreach(define_command) regular match { case Nil => case List(command) => define_command(command) case _ => define_commands(regular) } } /* execution */ def discontinue_execution(): Unit = protocol_command("Document.discontinue_execution") def cancel_exec(id: Document_ID.Exec): Unit = protocol_command("Document.cancel_exec", Document_ID(id)) /* document versions */ def update(old_id: Document_ID.Version, new_id: Document_ID.Version, edits: List[Document.Edit_Command], consolidate: List[Document.Node.Name]) { val consolidate_yxml = { import XML.Encode._ Symbol.encode_yxml(list(string)(consolidate.map(_.node))) } val edits_yxml = { import XML.Encode._ def id: T[Command] = (cmd => long(cmd.id)) def encode_edit(name: Document.Node.Name) : T[Document.Node.Edit[Command.Edit, Command.Perspective]] = variant(List( { case Document.Node.Edits(a) => (Nil, list(pair(option(id), option(id)))(a)) }, { case Document.Node.Deps(header) => val master_dir = File.standard_url(name.master_dir) val imports = header.imports.map(_.node) val keywords = header.keywords.map({ case (a, Keyword.Spec(b, c, d)) => (a, ((b, c), d)) }) (Nil, pair(string, pair(string, pair(list(string), pair(list(pair(string, pair(pair(string, list(string)), list(string)))), list(string)))))( (master_dir, (name.theory, (imports, (keywords, header.errors)))))) }, { case Document.Node.Perspective(a, b, c) => (bool_atom(a) :: b.commands.map(cmd => long_atom(cmd.id)), list(pair(id, pair(string, list(string))))(c.dest)) })) edits.map({ case (name, edit) => Symbol.encode_yxml(pair(string, encode_edit(name))(name.node, edit)) }) } protocol_command_args("Document.update", Document_ID(old_id) :: Document_ID(new_id) :: consolidate_yxml :: edits_yxml) } def remove_versions(versions: List[Document.Version]) { val versions_yxml = { import XML.Encode._ Symbol.encode_yxml(list(long)(versions.map(_.id))) } protocol_command("Document.remove_versions", versions_yxml) } /* dialog via document content */ def dialog_result(serial: Long, result: String): Unit = protocol_command("Document.dialog_result", Value.Long(serial), result) } 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,366 +1,366 @@ /* 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 = kind == Markup.INIT - def is_exit = kind == Markup.EXIT - def is_stdout = kind == Markup.STDOUT - def is_stderr = kind == Markup.STDERR - def is_system = kind == Markup.SYSTEM - def is_status = kind == Markup.STATUS - def is_report = kind == Markup.REPORT - def is_syslog = is_init || is_exit || is_system || is_stderr + 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) 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, xml_cache: XML.Cache, channel: System_Channel, process: Bash.Process) extends Protocol { /** receiver output **/ private def system_output(text: String) { receiver(new Prover.Output(XML.Elem(Markup(Markup.SYSTEM, Nil), List(XML.Text(text))))) } private def protocol_output(props: Properties.T, bytes: Bytes) { receiver(new Prover.Protocol_Output(props, bytes)) } private def output(kind: String, props: Properties.T, body: XML.Body) { 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(xml_cache.elem(msg))) } private def exit_message(result: Process_Result) { output(Markup.EXIT, Markup.Process_Result(result), List(XML.Text("Return code: " + result.rc.toString))) } /** 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() { try { process.terminate } catch { case exn @ ERROR(_) => system_output("Failed to terminate prover process: " + exn.getMessage) } } private val process_manager = Standard_Thread.fork("process_manager") { 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) } } Thread.sleep(10) } (finished.isEmpty || !finished.get, result.toString.trim) } if (startup_errors != "") system_output(startup_errors) if (startup_failed) { terminate_process() process_result.join exit_message(Process_Result(127)) } else { val (command_stream, message_stream) = channel.rendezvous() command_input_init(command_stream) val stdout = physical_output(false) 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() { process_manager.join() } def terminate() { system_output("Terminating prover process") command_input_close() var count = 10 while (!process_result.is_finished && count > 0) { Thread.sleep(100) 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) { 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) Standard_Thread.fork(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" Standard_Thread.fork(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 }) 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]) { 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,730 +1,730 @@ /* 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 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 = values + val background_colors: ValueSet = values // foreground val quoted, antiquoted = Value - val foreground_colors = values -- background_colors + val foreground_colors: ValueSet = values -- background_colors // message underline val writeln, information, warning, legacy, error = Value - val message_underline_colors = values -- background_colors -- foreground_colors + 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 = + 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 = + 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.Tree] = { 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") /* markup elements */ 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 warning_elements = Markup.Elements(Markup.WARNING, Markup.LEGACY) val error_elements = Markup.Elements(Markup.ERROR) val caret_focus_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) } abstract class Rendering( val snapshot: Document.Snapshot, val options: Options, val session: Session) { override def toString: String = "Rendering(" + snapshot.toString + ")" def model: Document.Model /* 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)) => model.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 citation(range: Text.Range): Option[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 }).headOption.map(_.info) /* file-system path completion */ def language_path(range: Text.Range): Option[Text.Range] = snapshot.select(range, Rendering.language_elements, _ => { case Text.Info(info_range, XML.Elem(Markup.Language(Markup.Language.PATH, _, _, _), _)) => Some(snapshot.convert(info_range)) case _ => None }).headOption.map(_.info) 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 { r1 <- language_path(before_caret_range(caret)) s1 <- model.get_text(r1) if is_wrapped(s1) r2 = Text.Range(r1.start + 1, r1.stop - 1) s2 = s1.substring(1, s1.length - 1) 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: Set[Long]) : List[Text.Info[Rendering.Color.Value]] = { for { Text.Info(r, result) <- snapshot.cumulate[(List[Markup], Option[Rendering.Color.Value])]( range, (List(Markup.Empty), None), Rendering.background_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(Markup.ENTITY, props), _))) => props match { case Markup.Entity.Def(i) if focus(i) => Some((Nil, Some(Rendering.Color.entity))) case Markup.Entity.Ref(i) if focus(i) => Some((Nil, Some(Rendering.Color.entity))) case _ => None } 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 (acc, 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)) }) /* caret focus */ private def entity_focus(range: Text.Range, check: (Boolean, Long) => Boolean = (is_def: Boolean, i: Long) => is_def): Set[Long] = { val results = snapshot.cumulate[Set[Long]](range, Set.empty, Rendering.caret_focus_elements, _ => { case (serials, Text.Info(_, XML.Elem(Markup(Markup.ENTITY, props), _))) => props match { case Markup.Entity.Def(i) if check(true, i) => Some(serials + i) case Markup.Entity.Ref(i) if check(false, i) => Some(serials + i) case _ => None } case _ => None }) (Set.empty[Long] /: results){ case (s1, Text.Info(_, s2)) => s1 ++ s2 } } def caret_focus(caret_range: Text.Range, visible_range: Text.Range): Set[Long] = { val focus_defs = entity_focus(caret_range) if (focus_defs.nonEmpty) focus_defs else { val visible_defs = entity_focus(visible_range) entity_focus(caret_range, (is_def: Boolean, i: Long) => !is_def && visible_defs.contains(i)) } } def caret_focus_ranges(caret_range: Text.Range, visible_range: Text.Range): List[Text.Range] = { val focus = caret_focus(caret_range, visible_range) if (focus.nonEmpty) { val results = snapshot.cumulate[Boolean](visible_range, false, Rendering.caret_focus_elements, _ => { case (_, Text.Info(_, XML.Elem(Markup(Markup.ENTITY, props), _))) => props match { case Markup.Entity.Def(i) if focus(i) => Some(true) case Markup.Entity.Ref(i) if focus(i) => Some(true) case _ => None } }) for (info <- results if info.info) yield info.range } else Nil } /* message underline color */ 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) } /* tooltips */ def timing_threshold: Double = 0.0 private sealed case class Tooltip_Info( range: Text.Range, timing: Timing = Timing.zero, messages: List[Command.Results.Entry] = 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 = Command.Results.make(results.flatMap(_.messages)).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/session.scala b/src/Pure/PIDE/session.scala --- a/src/Pure/PIDE/session.scala +++ b/src/Pure/PIDE/session.scala @@ -1,746 +1,746 @@ /* Title: Pure/PIDE/session.scala Author: Makarius Options: :folding=explicit: PIDE editor session, potentially with running prover process. */ package isabelle import scala.collection.immutable.Queue import scala.collection.mutable import scala.annotation.tailrec object Session { /* outlets */ object Consumer { def apply[A](name: String)(consume: A => Unit): Consumer[A] = new Consumer[A](name, consume) } final class Consumer[-A] private(val name: String, val consume: A => Unit) class Outlet[A](dispatcher: Consumer_Thread[() => Unit]) { private val consumers = Synchronized[List[Consumer[A]]](Nil) def += (c: Consumer[A]) { consumers.change(Library.update(c)) } def -= (c: Consumer[A]) { consumers.change(Library.remove(c)) } def post(a: A) { for (c <- consumers.value.iterator) { dispatcher.send(() => try { c.consume(a) } catch { case exn: Throwable => Output.error_message("Consumer failed: " + quote(c.name) + "\n" + Exn.message(exn)) }) } } } /* change */ sealed case class Change( previous: Document.Version, syntax_changed: List[Document.Node.Name], deps_changed: Boolean, doc_edits: List[Document.Edit_Command], consolidate: List[Document.Node.Name], version: Document.Version) case object Change_Flush /* events */ //{{{ case class Statistics(props: Properties.T) case class Global_Options(options: Options) case object Caret_Focus case class Raw_Edits(doc_blobs: Document.Blobs, edits: List[Document.Edit_Text]) case class Dialog_Result(id: Document_ID.Generic, serial: Long, result: String) case class Build_Theories(id: String, master_dir: Path, theories: List[(Options, List[Path])]) case class Commands_Changed( assignment: Boolean, nodes: Set[Document.Node.Name], commands: Set[Command]) sealed abstract class Phase { def print: String = this match { case Terminated(result) => if (result.ok) "finished" else "failed" case _ => Word.lowercase(this.toString) } } case object Inactive extends Phase // stable case object Startup extends Phase // transient case object Ready extends Phase // metastable case object Shutdown extends Phase // transient case class Terminated(result: Process_Result) extends Phase // stable //}}} /* syslog */ private[Session] class Syslog(limit: Int) { private var queue = Queue.empty[XML.Elem] private var length = 0 def += (msg: XML.Elem): Unit = synchronized { queue = queue.enqueue(msg) length += 1 if (length > limit) queue = queue.dequeue._2 } def content: String = synchronized { cat_lines(queue.iterator.map(XML.content)) + (if (length > limit) "\n(A total of " + length + " messages...)" else "") } } /* protocol handlers */ abstract class Protocol_Handler { def init(session: Session): Unit = {} def exit(): Unit = {} val functions: List[(String, Prover.Protocol_Output => Boolean)] } } class Session(_session_options: => Options, val resources: Resources) extends Document.Session { session => val xml_cache: XML.Cache = XML.make_cache() val xz_cache: XZ.Cache = XZ.make_cache() /* global flags */ @volatile var timing: Boolean = false @volatile var verbose: Boolean = false /* dynamic session options */ def session_options: Options = _session_options def output_delay: Time = session_options.seconds("editor_output_delay") def consolidate_delay: Time = session_options.seconds("editor_consolidate_delay") def prune_delay: Time = session_options.seconds("editor_prune_delay") def prune_size: Int = session_options.int("editor_prune_size") def syslog_limit: Int = session_options.int("editor_syslog_limit") def reparse_limit: Int = session_options.int("editor_reparse_limit") /* dispatcher */ private val dispatcher = Consumer_Thread.fork[() => Unit]("Session.dispatcher", daemon = true) { case e => e(); true } def assert_dispatcher[A](body: => A): A = { assert(dispatcher.check_thread) body } def require_dispatcher[A](body: => A): A = { require(dispatcher.check_thread) body } def send_dispatcher(body: => Unit): Unit = { if (dispatcher.check_thread) body else dispatcher.send(() => body) } def send_wait_dispatcher(body: => Unit): Unit = { if (dispatcher.check_thread) body else dispatcher.send_wait(() => body) } /* outlets */ val statistics = new Session.Outlet[Session.Statistics](dispatcher) val global_options = new Session.Outlet[Session.Global_Options](dispatcher) val caret_focus = new Session.Outlet[Session.Caret_Focus.type](dispatcher) val raw_edits = new Session.Outlet[Session.Raw_Edits](dispatcher) val commands_changed = new Session.Outlet[Session.Commands_Changed](dispatcher) val phase_changed = new Session.Outlet[Session.Phase](dispatcher) val syslog_messages = new Session.Outlet[Prover.Output](dispatcher) val raw_output_messages = new Session.Outlet[Prover.Output](dispatcher) val trace_events = new Session.Outlet[Simplifier_Trace.Event.type](dispatcher) val debugger_updates = new Session.Outlet[Debugger.Update.type](dispatcher) val all_messages = new Session.Outlet[Prover.Message](dispatcher) // potential bottle-neck! /** main protocol manager **/ /* internal messages */ private case class Start(start_prover: Prover.Receiver => Prover) private case object Stop private case class Get_State(promise: Promise[Document.State]) private case class Cancel_Exec(exec_id: Document_ID.Exec) private case class Protocol_Command(name: String, args: List[String]) private case class Update_Options(options: Options) private case object Consolidate_Execution private case object Prune_History /* phase */ private def post_phase(new_phase: Session.Phase): Session.Phase = { phase_changed.post(new_phase) new_phase } private val _phase = Synchronized[Session.Phase](Session.Inactive) private def phase_=(new_phase: Session.Phase): Unit = _phase.change(_ => post_phase(new_phase)) - def phase = _phase.value + def phase: Session.Phase = _phase.value def is_ready: Boolean = phase == Session.Ready /* syslog */ private val syslog = new Session.Syslog(syslog_limit) def syslog_content(): String = syslog.content /* pipelined change parsing */ private case class Text_Edits( previous: Future[Document.Version], doc_blobs: Document.Blobs, text_edits: List[Document.Edit_Text], consolidate: List[Document.Node.Name], version_result: Promise[Document.Version]) private val change_parser = Consumer_Thread.fork[Text_Edits]("change_parser", daemon = true) { case Text_Edits(previous, doc_blobs, text_edits, consolidate, version_result) => val prev = previous.get_finished val change = Timing.timeit("parse_change", timing) { resources.parse_change(reparse_limit, prev, doc_blobs, text_edits, consolidate) } version_result.fulfill(change.version) manager.send(change) true } /* buffered changes */ private object change_buffer { private var assignment: Boolean = false private var nodes: Set[Document.Node.Name] = Set.empty private var commands: Set[Command] = Set.empty def flush(): Unit = synchronized { if (assignment || nodes.nonEmpty || commands.nonEmpty) commands_changed.post(Session.Commands_Changed(assignment, nodes, commands)) if (nodes.nonEmpty) consolidation.update(nodes) assignment = false nodes = Set.empty commands = Set.empty } private val delay_flush = Standard_Thread.delay_first(output_delay) { flush() } def invoke(assign: Boolean, edited_nodes: List[Document.Node.Name], cmds: List[Command]): Unit = synchronized { assignment |= assign for (node <- edited_nodes) { nodes += node } for (command <- cmds) { nodes += command.node_name command.blobs_names.foreach(nodes += _) commands += command } delay_flush.invoke() } def shutdown() { delay_flush.revoke() flush() } } /* postponed changes */ private object postponed_changes { private var postponed: List[Session.Change] = Nil def store(change: Session.Change): Unit = synchronized { postponed ::= change } def flush(state: Document.State): List[Session.Change] = synchronized { val (assigned, unassigned) = postponed.partition(change => state.is_assigned(change.previous)) postponed = unassigned assigned.reverse } } /* node consolidation */ private object consolidation { private val delay = Standard_Thread.delay_first(consolidate_delay) { manager.send(Consolidate_Execution) } private val init_state: Option[Set[Document.Node.Name]] = Some(Set.empty) private val state = Synchronized(init_state) def exit() { delay.revoke() state.change(_ => None) } def update(new_nodes: Set[Document.Node.Name] = Set.empty) { val active = state.change_result(st => (st.isDefined, st.map(nodes => if (nodes.isEmpty) new_nodes else nodes ++ new_nodes))) if (active) delay.invoke() } def flush(): Set[Document.Node.Name] = state.change_result(st => if (st.isDefined) (st.get, init_state) else (Set.empty, None)) } /* prover process */ private object prover { private val variable = Synchronized[Option[Prover]](None) def defined: Boolean = variable.value.isDefined def get: Prover = variable.value.get def set(p: Prover) { variable.change(_ => Some(p)) } def reset { variable.change(_ => None) } def await_reset() { variable.guarded_access({ case None => Some((), None) case _ => None }) } } /* file formats */ lazy val file_formats: File_Format.Session = resources.file_formats.start_session(session) /* protocol handlers */ private val protocol_handlers = Protocol_Handlers.init(session) def get_protocol_handler(name: String): Option[Session.Protocol_Handler] = protocol_handlers.get(name) def init_protocol_handler(handler: Session.Protocol_Handler): Unit = protocol_handlers.init(handler) def init_protocol_handler(name: String): Unit = protocol_handlers.init(name) /* debugger */ private val debugger_handler = new Debugger.Handler(this) init_protocol_handler(debugger_handler) def debugger: Debugger = debugger_handler.debugger /* manager thread */ private val delay_prune = Standard_Thread.delay_first(prune_delay) { manager.send(Prune_History) } private val manager: Consumer_Thread[Any] = { /* global state */ val global_state = Synchronized(Document.State.init) /* raw edits */ def handle_raw_edits( doc_blobs: Document.Blobs = Document.Blobs.empty, edits: List[Document.Edit_Text] = Nil, consolidate: List[Document.Node.Name] = Nil) //{{{ { require(prover.defined) if (edits.nonEmpty) prover.get.discontinue_execution() val previous = global_state.value.history.tip.version val version = Future.promise[Document.Version] global_state.change(_.continue_history(previous, edits, version)) raw_edits.post(Session.Raw_Edits(doc_blobs, edits)) change_parser.send(Text_Edits(previous, doc_blobs, edits, consolidate, version)) } //}}} /* resulting changes */ def handle_change(change: Session.Change) //{{{ { require(prover.defined) // define commands { val id_commands = new mutable.ListBuffer[Command] def id_command(command: Command) { for { (name, digest) <- command.blobs_defined if !global_state.value.defined_blob(digest) } { change.version.nodes(name).get_blob match { case Some(blob) => global_state.change(_.define_blob(digest)) prover.get.define_blob(digest, blob.bytes) case None => Output.error_message("Missing blob " + quote(name.toString)) } } if (!global_state.value.defined_command(command.id)) { global_state.change(_.define_command(command)) id_commands += command } } for { (_, edit) <- change.doc_edits } { edit.foreach({ case (c1, c2) => c1.foreach(id_command); c2.foreach(id_command) }) } if (id_commands.nonEmpty) prover.get.define_commands_bulk(id_commands.toList) } val assignment = global_state.value.the_assignment(change.previous).check_finished global_state.change(_.define_version(change.version, assignment)) prover.get.update(change.previous.id, change.version.id, change.doc_edits, change.consolidate) resources.commit(change) } //}}} /* prover output */ def handle_output(output: Prover.Output) //{{{ { def bad_output() { if (verbose) Output.warning("Ignoring bad prover output: " + output.message.toString) } def change_command(f: Document.State => (Command.State, Document.State)) { try { val st = global_state.change_result(f) change_buffer.invoke(false, Nil, List(st.command)) } catch { case _: Document.State.Fail => bad_output() } } output match { case msg: Prover.Protocol_Output => val handled = protocol_handlers.invoke(msg) if (!handled) { msg.properties match { case Markup.Protocol_Handler(name) if prover.defined => init_protocol_handler(name) case Protocol.Command_Timing(state_id, timing) if prover.defined => val message = XML.elem(Markup.STATUS, List(XML.Elem(Markup.Timing(timing), Nil))) change_command(_.accumulate(state_id, xml_cache.elem(message), xml_cache)) case Protocol.Theory_Timing(_, _) => // FIXME case Markup.Export(args) if args.id.isDefined && Value.Long.unapply(args.id.get).isDefined => val id = Value.Long.unapply(args.id.get).get val export = Export.make_entry("", args, msg.bytes, cache = xz_cache) change_command(_.add_export(id, (args.serial, export))) case Markup.Commands_Accepted => msg.text match { case Protocol.Commands_Accepted(ids) => ids.foreach(id => change_command(_.accumulate(id, Protocol.Commands_Accepted.message, xml_cache))) case _ => bad_output() } case Markup.Assign_Update => msg.text match { case Protocol.Assign_Update(id, edited, update) => try { val (edited_nodes, cmds) = global_state.change_result(_.assign(id, edited, update)) change_buffer.invoke(true, edited_nodes, cmds) manager.send(Session.Change_Flush) } catch { case _: Document.State.Fail => bad_output() } case _ => bad_output() } delay_prune.invoke() case Markup.Removed_Versions => msg.text match { case Protocol.Removed(removed) => try { global_state.change(_.removed_versions(removed)) manager.send(Session.Change_Flush) } catch { case _: Document.State.Fail => bad_output() } case _ => bad_output() } case Markup.ML_Statistics(props) => statistics.post(Session.Statistics(props)) case Markup.Task_Statistics(props) => // FIXME case _ => bad_output() } } case _ => output.properties match { case Position.Id(state_id) => change_command(_.accumulate(state_id, output.message, xml_cache)) case _ if output.is_init => prover.get.options(file_formats.prover_options(session_options)) prover.get.session_base(resources) phase = Session.Ready debugger.ready() case Markup.Process_Result(result) if output.is_exit => file_formats.stop_session phase = Session.Terminated(result) prover.reset case _ => raw_output_messages.post(output) } } } //}}} /* main thread */ Consumer_Thread.fork[Any]("Session.manager", daemon = true) { case arg: Any => //{{{ arg match { case output: Prover.Output => if (output.is_stdout || output.is_stderr) raw_output_messages.post(output) else handle_output(output) if (output.is_syslog) { syslog += output.message syslog_messages.post(output) } all_messages.post(output) case input: Prover.Input => all_messages.post(input) case Start(start_prover) if !prover.defined => prover.set(start_prover(manager.send(_))) case Stop => consolidation.exit() delay_prune.revoke() if (prover.defined) { protocol_handlers.exit() global_state.change(_ => Document.State.init) prover.get.terminate } case Get_State(promise) => promise.fulfill(global_state.value) case Consolidate_Execution => if (prover.defined) { val state = global_state.value state.stable_tip_version match { case None => consolidation.update() case Some(version) => val consolidate = consolidation.flush().iterator.filter(name => !resources.session_base.loaded_theory(name) && !state.node_consolidated(version, name) && state.node_maybe_consolidated(version, name)).toList if (consolidate.nonEmpty) handle_raw_edits(consolidate = consolidate) } } case Prune_History => if (prover.defined) { val old_versions = global_state.change_result(_.remove_versions(prune_size)) if (old_versions.nonEmpty) prover.get.remove_versions(old_versions) } case Update_Options(options) => if (prover.defined && is_ready) { prover.get.options(file_formats.prover_options(options)) handle_raw_edits() } global_options.post(Session.Global_Options(options)) case Cancel_Exec(exec_id) if prover.defined => prover.get.cancel_exec(exec_id) case Session.Raw_Edits(doc_blobs, edits) if prover.defined => handle_raw_edits(doc_blobs = doc_blobs, edits = edits) case Session.Dialog_Result(id, serial, result) if prover.defined => prover.get.dialog_result(serial, result) handle_output(new Prover.Output(Protocol.Dialog_Result(id, serial, result))) case Protocol_Command(name, args) if prover.defined => prover.get.protocol_command_args(name, args) case change: Session.Change if prover.defined => val state = global_state.value if (!state.removing_versions && state.is_assigned(change.previous)) handle_change(change) else postponed_changes.store(change) case Session.Change_Flush if prover.defined => val state = global_state.value if (!state.removing_versions) - postponed_changes.flush(state).foreach(handle_change(_)) + postponed_changes.flush(state).foreach(handle_change) case bad => if (verbose) Output.warning("Ignoring bad message: " + bad.toString) } true //}}} } } /* main operations */ def get_state(): Document.State = { if (manager.is_active) { val promise = Future.promise[Document.State] manager.send_wait(Get_State(promise)) promise.join } else Document.State.init } def snapshot(name: Document.Node.Name = Document.Node.Name.empty, pending_edits: List[Text.Edit] = Nil): Document.Snapshot = get_state().snapshot(name, pending_edits) def recent_syntax(name: Document.Node.Name): Outer_Syntax = get_state().recent_finished.version.get_finished.nodes(name).syntax getOrElse resources.session_base.overall_syntax @tailrec final def await_stable_snapshot(): Document.Snapshot = { val snapshot = this.snapshot() if (snapshot.is_outdated) { Thread.sleep(output_delay.ms) await_stable_snapshot() } else snapshot } def start(start_prover: Prover.Receiver => Prover) { file_formats _phase.change( { case Session.Inactive => manager.send(Start(start_prover)) post_phase(Session.Startup) case phase => error("Cannot start prover in phase " + quote(phase.print)) }) } def send_stop() { val was_ready = _phase.guarded_access(phase => phase match { case Session.Startup | Session.Shutdown => None case Session.Terminated(_) => Some((false, phase)) case Session.Inactive => Some((false, post_phase(Session.Terminated(Process_Result(0))))) case Session.Ready => Some((true, post_phase(Session.Shutdown))) }) if (was_ready) manager.send(Stop) } def stop(): Process_Result = { send_stop() prover.await_reset() change_parser.shutdown() change_buffer.shutdown() manager.shutdown() dispatcher.shutdown() phase match { case Session.Terminated(result) => result case phase => error("Bad session phase after shutdown: " + quote(phase.print)) } } def protocol_command(name: String, args: String*) { manager.send(Protocol_Command(name, args.toList)) } def cancel_exec(exec_id: Document_ID.Exec) { manager.send(Cancel_Exec(exec_id)) } def update(doc_blobs: Document.Blobs, edits: List[Document.Edit_Text]) { if (edits.nonEmpty) manager.send_wait(Session.Raw_Edits(doc_blobs, edits)) } def update_options(options: Options) { manager.send_wait(Update_Options(options)) } def dialog_result(id: Document_ID.Generic, serial: Long, result: String) { manager.send(Session.Dialog_Result(id, serial, result)) } } 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,416 +1,416 @@ /* 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 { 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 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 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.Text(s) => op(x, s) } (a /: body)(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(c: Char, s: StringBuilder) { c match { case '<' => s ++= "<" case '>' => s ++= ">" case '&' => s ++= "&" case '"' => s ++= """ case '\'' => s ++= "'" case _ => s += c } } def output_string(str: String, s: StringBuilder) { if (str == null) s ++= str else str.iterator.foreach(c => output_char(c, s)) } def string_of_body(body: Body): String = { val s = new StringBuilder def text(txt: String) { output_string(txt, s) } def elem(markup: Markup) { s ++= markup.name for ((a, b) <- markup.properties) { s += ' '; s ++= a; s += '='; s += '"'; text(b); s += '"' } } def tree(t: Tree): Unit = t match { case XML.Elem(markup, Nil) => s += '<'; elem(markup); s ++= "/>" case XML.Elem(markup, ts) => s += '<'; elem(markup); s += '>' ts.foreach(tree) s ++= "' case XML.Text(txt) => text(txt) } body.foreach(tree) s.toString } def string_of_tree(tree: XML.Tree): String = string_of_body(List(tree)) /** cache **/ def make_cache(initial_size: Int = 131071, max_string: Int = 100): Cache = new Cache(initial_size, max_string) class Cache private[XML](initial_size: Int, max_string: Int) extends isabelle.Cache(initial_size, max_string) { 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(_)) + case None => x.map(cache_tree) } } // main methods def props(x: Properties.T): Properties.T = synchronized { cache_props(x) } def markup(x: Markup): Markup = synchronized { cache_markup(x) } def tree(x: XML.Tree): XML.Tree = synchronized { cache_tree(x) } def body(x: XML.Body): XML.Body = synchronized { cache_body(x) } def elem(x: XML.Elem): XML.Elem = 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/PIDE/yxml.scala b/src/Pure/PIDE/yxml.scala --- a/src/Pure/PIDE/yxml.scala +++ b/src/Pure/PIDE/yxml.scala @@ -1,155 +1,155 @@ /* Title: Pure/PIDE/yxml.scala Author: Makarius Efficient text representation of XML trees. Suitable for direct inlining into plain text. */ package isabelle import scala.collection.mutable object YXML { /* chunk markers */ val X = '\u0005' val Y = '\u0006' val is_X = (c: Char) => c == X val is_Y = (c: Char) => c == Y - val X_string = X.toString - val Y_string = Y.toString - val XY_string = X_string + Y_string - val XYX_string = XY_string + X_string + val X_string: String = X.toString + val Y_string: String = Y.toString + val XY_string: String = X_string + Y_string + val XYX_string: String = XY_string + X_string def detect(s: String): Boolean = s.exists(c => c == X || c == Y) def detect_elem(s: String): Boolean = s.startsWith(XY_string) /* string representation */ // FIXME byte array version with pseudo-utf-8 (!?) def traversal(string: String => Unit, body: XML.Body): Unit = { def tree(t: XML.Tree): Unit = t match { case XML.Elem(Markup(name, atts), ts) => string(XY_string) string(name) for ((a, x) <- atts) { string(Y_string); string(a); string("="); string(x) } string(X_string) ts.foreach(tree) string(XYX_string) case XML.Text(text) => string(text) } body.foreach(tree) } def string_of_body(body: XML.Body): String = { val s = new StringBuilder traversal(str => s ++= str, body) s.toString } def string_of_tree(tree: XML.Tree): String = string_of_body(List(tree)) /* parsing */ private def err(msg: String) = error("Malformed YXML: " + msg) private def err_attribute() = err("bad attribute") private def err_element() = err("bad element") private def err_unbalanced(name: String) = if (name == "") err("unbalanced element") else err("unbalanced element " + quote(name)) private def parse_attrib(source: CharSequence) = { val s = source.toString val i = s.indexOf('=') if (i <= 0) err_attribute() (s.substring(0, i), s.substring(i + 1)) } def parse_body(source: CharSequence): XML.Body = { /* stack operations */ def buffer(): mutable.ListBuffer[XML.Tree] = new mutable.ListBuffer[XML.Tree] var stack: List[(Markup, mutable.ListBuffer[XML.Tree])] = List((Markup.Empty, buffer())) def add(x: XML.Tree) { (stack: @unchecked) match { case ((_, body) :: _) => body += x } } def push(name: String, atts: XML.Attributes) { if (name == "") err_element() else stack = (Markup(name, atts), buffer()) :: stack } def pop() { (stack: @unchecked) match { case ((Markup.Empty, _) :: _) => err_unbalanced("") case ((markup, body) :: pending) => stack = pending add(XML.Elem(markup, body.toList)) } } /* parse chunks */ for (chunk <- Library.separated_chunks(is_X, source) if chunk.length != 0) { if (chunk.length == 1 && chunk.charAt(0) == Y) pop() else { Library.separated_chunks(is_Y, chunk).toList match { case ch :: name :: atts if ch.length == 0 => push(name.toString, atts.map(parse_attrib)) case txts => for (txt <- txts) add(XML.Text(txt.toString)) } } } (stack: @unchecked) match { case List((Markup.Empty, body)) => body.toList case (Markup(name, _), _) :: _ => err_unbalanced(name) } } def parse(source: CharSequence): XML.Tree = parse_body(source) match { case List(result) => result case Nil => XML.Text("") case _ => err("multiple XML trees") } def parse_elem(source: CharSequence): XML.Tree = parse_body(source) match { case List(elem: XML.Elem) => elem case _ => err("single XML element expected") } /* failsafe parsing */ private def markup_broken(source: CharSequence) = XML.Elem(Markup.Broken, List(XML.Text(source.toString))) def parse_body_failsafe(source: CharSequence): XML.Body = { try { parse_body(source) } catch { case ERROR(_) => List(markup_broken(source)) } } def parse_failsafe(source: CharSequence): XML.Tree = { try { parse(source) } catch { case ERROR(_) => markup_broken(source) } } } diff --git a/src/Pure/System/bash.scala b/src/Pure/System/bash.scala --- a/src/Pure/System/bash.scala +++ b/src/Pure/System/bash.scala @@ -1,206 +1,206 @@ /* Title: Pure/System/bash.scala Author: Makarius GNU bash processes, with propagation of interrupts. */ package isabelle import java.io.{File => JFile, BufferedReader, InputStreamReader, BufferedWriter, OutputStreamWriter} object Bash { /* concrete syntax */ private def bash_chr(c: Byte): String = { val ch = c.toChar ch match { case '\t' => "$'\\t'" case '\n' => "$'\\n'" case '\f' => "$'\\f'" case '\r' => "$'\\r'" case _ => if (Symbol.is_ascii_letter(ch) || Symbol.is_ascii_digit(ch) || "+,-./:_".contains(ch)) Symbol.ascii(ch) else if (c < 0) "$'\\x" + Integer.toHexString(256 + c) + "'" else if (c < 16) "$'\\x0" + Integer.toHexString(c) + "'" else if (c < 32 || c >= 127) "$'\\x" + Integer.toHexString(c) + "'" else "\\" + ch } } def string(s: String): String = if (s == "") "\"\"" - else UTF8.bytes(s).iterator.map(bash_chr(_)).mkString + else UTF8.bytes(s).iterator.map(bash_chr).mkString def strings(ss: List[String]): String = - ss.iterator.map(Bash.string(_)).mkString(" ") + ss.iterator.map(Bash.string).mkString(" ") /* process and result */ private class Limited_Progress(proc: Process, progress_limit: Option[Long]) { private var count = 0L def apply(progress: String => Unit)(line: String): Unit = synchronized { progress(line) count = count + line.length + 1 progress_limit match { case Some(limit) if count > limit => proc.terminate case _ => } } } def process(script: String, cwd: JFile = null, env: Map[String, String] = Isabelle_System.settings(), redirect: Boolean = false, cleanup: () => Unit = () => ()): Process = new Process(script, cwd, env, redirect, cleanup) class Process private[Bash]( script: String, cwd: JFile, env: Map[String, String], redirect: Boolean, cleanup: () => Unit) { private val timing_file = Isabelle_System.tmp_file("bash_timing") private val timing = Synchronized[Option[Timing]](None) def get_timing: Timing = timing.value getOrElse Timing.zero private val script_file = Isabelle_System.tmp_file("bash_script") File.write(script_file, script) private val proc = Isabelle_System.process( List(File.platform_path(Path.variable("ISABELLE_BASH_PROCESS")), "-", File.standard_path(timing_file), "bash", File.standard_path(script_file)), cwd = cwd, env = env, redirect = redirect) // channels val stdin: BufferedWriter = new BufferedWriter(new OutputStreamWriter(proc.getOutputStream, UTF8.charset)) val stdout: BufferedReader = new BufferedReader(new InputStreamReader(proc.getInputStream, UTF8.charset)) val stderr: BufferedReader = new BufferedReader(new InputStreamReader(proc.getErrorStream, UTF8.charset)) // signals private val pid = stdout.readLine def interrupt() { Exn.Interrupt.postpone { Isabelle_System.kill("INT", pid) } } private def kill(signal: String): Boolean = Exn.Interrupt.postpone { Isabelle_System.kill(signal, pid) Isabelle_System.kill("0", pid)._2 == 0 } getOrElse true private def multi_kill(signal: String): Boolean = { var running = true var count = 10 while (running && count > 0) { if (kill(signal)) { Exn.Interrupt.postpone { Thread.sleep(100) count -= 1 } } else running = false } running } def terminate() { multi_kill("INT") && multi_kill("TERM") && kill("KILL") proc.destroy do_cleanup() } // JVM shutdown hook private val shutdown_hook = new Thread { override def run = terminate() } try { Runtime.getRuntime.addShutdownHook(shutdown_hook) } catch { case _: IllegalStateException => } // cleanup private def do_cleanup() { try { Runtime.getRuntime.removeShutdownHook(shutdown_hook) } catch { case _: IllegalStateException => } script_file.delete timing.change { case None => if (timing_file.isFile) { val t = Word.explode(File.read(timing_file)) match { case List(Value.Long(elapsed), Value.Long(cpu)) => Timing(Time.ms(elapsed), Time.ms(cpu), Time.zero) case _ => Timing.zero } timing_file.delete Some(t) } else Some(Timing.zero) case some => some } cleanup() } // join def join: Int = { val rc = proc.waitFor do_cleanup() rc } // result def result( progress_stdout: String => Unit = (_: String) => (), progress_stderr: String => Unit = (_: String) => (), progress_limit: Option[Long] = None, strict: Boolean = true): Process_Result = { stdin.close val limited = new Limited_Progress(this, progress_limit) val out_lines = Future.thread("bash_stdout") { File.read_lines(stdout, limited(progress_stdout)) } val err_lines = Future.thread("bash_stderr") { File.read_lines(stderr, limited(progress_stderr)) } val rc = try { join } catch { case Exn.Interrupt() => terminate(); Exn.Interrupt.return_code } if (strict && rc == Exn.Interrupt.return_code) throw Exn.Interrupt() Process_Result(rc, out_lines.join, err_lines.join, false, get_timing) } } } 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,69 +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) { 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)) } - private def getopts(args: List[List[Char]]): List[List[Char]] = + @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/invoke_scala.scala b/src/Pure/System/invoke_scala.scala --- a/src/Pure/System/invoke_scala.scala +++ b/src/Pure/System/invoke_scala.scala @@ -1,128 +1,128 @@ /* Title: Pure/System/invoke_scala.scala Author: Makarius JVM method invocation service via Isabelle/Scala. */ package isabelle import java.lang.reflect.{Method, Modifier, InvocationTargetException} import scala.util.matching.Regex object Invoke_Scala { /* method reflection */ private val Ext = new Regex("(.*)\\.([^.]*)") private val STRING = Class.forName("java.lang.String") private def get_method(name: String): String => String = name match { case Ext(class_name, method_name) => val m = try { Class.forName(class_name).getMethod(method_name, STRING) } catch { case _: ClassNotFoundException | _: NoSuchMethodException => error("No such method: " + quote(name)) } if (!Modifier.isStatic(m.getModifiers)) error("Not at static method: " + m.toString) if (m.getReturnType != STRING) error("Bad method return type: " + m.toString) (arg: String) => { try { m.invoke(null, arg).asInstanceOf[String] } catch { case e: InvocationTargetException if e.getCause != null => throw e.getCause } } case _ => error("Malformed method name: " + quote(name)) } /* method invocation */ object Tag extends Enumeration { val NULL, OK, ERROR, FAIL, INTERRUPT = Value } def method(name: String, arg: String): (Tag.Value, String) = Exn.capture { get_method(name) } match { case Exn.Res(f) => Exn.capture { f(arg) } match { case Exn.Res(null) => (Tag.NULL, "") case Exn.Res(res) => (Tag.OK, res) case Exn.Exn(Exn.Interrupt()) => (Tag.INTERRUPT, "") case Exn.Exn(e) => (Tag.ERROR, Exn.message(e)) } case Exn.Exn(e) => (Tag.FAIL, Exn.message(e)) } } /* protocol handler */ class Invoke_Scala extends Session.Protocol_Handler { private var session: Session = null private var futures = Map.empty[String, Future[Unit]] override def init(init_session: Session): Unit = synchronized { session = init_session } override def exit(): Unit = synchronized { for ((id, future) <- futures) cancel(id, future) futures = Map.empty } private def fulfill(id: String, tag: Invoke_Scala.Tag.Value, res: String): Unit = synchronized { if (futures.isDefinedAt(id)) { session.protocol_command("Invoke_Scala.fulfill", id, tag.id.toString, res) futures -= id } } private def cancel(id: String, future: Future[Unit]) { future.cancel fulfill(id, Invoke_Scala.Tag.INTERRUPT, "") } private def invoke_scala(msg: Prover.Protocol_Output): Boolean = synchronized { msg.properties match { case Markup.Invoke_Scala(name, id) => futures += (id -> Future.fork { val (tag, result) = Invoke_Scala.method(name, msg.text) fulfill(id, tag, result) }) true case _ => false } } private def cancel_scala(msg: Prover.Protocol_Output): Boolean = synchronized { msg.properties match { case Markup.Cancel_Scala(id) => futures.get(id) match { case Some(future) => cancel(id, future) case None => } true case _ => false } } val functions = List( - Markup.INVOKE_SCALA -> invoke_scala _, - Markup.CANCEL_SCALA -> cancel_scala _) + Markup.INVOKE_SCALA -> invoke_scala, + Markup.CANCEL_SCALA -> cancel_scala) } diff --git a/src/Pure/System/isabelle_process.scala b/src/Pure/System/isabelle_process.scala --- a/src/Pure/System/isabelle_process.scala +++ b/src/Pure/System/isabelle_process.scala @@ -1,63 +1,44 @@ /* Title: Pure/System/isabelle_process.scala Author: Makarius Isabelle process wrapper. */ package isabelle import java.io.{File => JFile} object Isabelle_Process { - def start(session: Session, + def apply( + session: Session, options: Options, + sessions_structure: Sessions.Structure, + store: Sessions.Store, logic: String = "", args: List[String] = Nil, - dirs: List[Path] = Nil, modes: List[String] = Nil, cwd: JFile = null, env: Map[String, String] = Isabelle_System.settings(), - sessions_structure: Option[Sessions.Structure] = None, - store: Option[Sessions.Store] = None, phase_changed: Session.Phase => Unit = null) { - if (phase_changed != null) - session.phase_changed += Session.Consumer("Isabelle_Process")(phase_changed) - - session.start(receiver => - Isabelle_Process(options, logic = logic, args = args, dirs = dirs, modes = modes, - cwd = cwd, env = env, receiver = receiver, xml_cache = session.xml_cache, - sessions_structure = sessions_structure, store = store)) - } - - def apply( - options: Options, - logic: String = "", - args: List[String] = Nil, - dirs: List[Path] = Nil, - modes: List[String] = Nil, - cwd: JFile = null, - env: Map[String, String] = Isabelle_System.settings(), - receiver: Prover.Receiver = (msg: Prover.Message) => Output.writeln(msg.toString, stdout = true), - xml_cache: XML.Cache = XML.make_cache(), - sessions_structure: Option[Sessions.Structure] = None, - store: Option[Sessions.Store] = None): Prover = - { val channel = System_Channel() val process = try { val channel_options = options.string.update("system_channel_address", channel.address). string.update("system_channel_password", channel.password) - ML_Process(channel_options, logic = logic, args = args, dirs = dirs, modes = modes, - cwd = cwd, env = env, sessions_structure = sessions_structure, store = store) + ML_Process(channel_options, sessions_structure, store, + logic = logic, args = args, modes = modes, cwd = cwd, env = env) } catch { case exn @ ERROR(_) => channel.shutdown(); throw exn } process.stdin.close - new Prover(receiver, xml_cache, channel, process) + if (phase_changed != null) + session.phase_changed += Session.Consumer("Isabelle_Process")(phase_changed) + + session.start(receiver => new Prover(receiver, session.xml_cache, channel, process)) } } diff --git a/src/Pure/System/numa.scala b/src/Pure/System/numa.scala --- a/src/Pure/System/numa.scala +++ b/src/Pure/System/numa.scala @@ -1,86 +1,86 @@ /* Title: Pure/System/numa.scala Author: Makarius Support for Non-Uniform Memory Access of separate CPU nodes. */ package isabelle object NUMA { /* available nodes */ def nodes(): List[Int] = { val numa_nodes_linux = Path.explode("/sys/devices/system/node/online") val Single = """^(\d+)$""".r val Multiple = """^(\d+)-(\d+)$""".r def read(s: String): List[Int] = s match { case Single(Value.Int(i)) => List(i) case Multiple(Value.Int(i), Value.Int(j)) => (i to j).toList case _ => error("Cannot parse CPU node specification: " + quote(s)) } if (numa_nodes_linux.is_file) { - space_explode(',', File.read(numa_nodes_linux).trim).flatMap(read(_)) + space_explode(',', File.read(numa_nodes_linux).trim).flatMap(read) } else Nil } /* CPU policy via numactl tool */ lazy val numactl_available: Boolean = Isabelle_System.bash("numactl -m0 -N0 true").ok def policy(node: Int): String = if (numactl_available) "numactl -m" + node + " -N" + node else "" def policy_options(options: Options, numa_node: Option[Int] = Some(0)): Options = numa_node match { case None => options case Some(n) => options.string("ML_process_policy") = policy(n) } /* shuffling of CPU nodes */ def enabled: Boolean = try { nodes().length >= 2 && numactl_available } catch { case ERROR(_) => false } def enabled_warning(progress: Progress, enabled: Boolean): Boolean = { def warning = if (nodes().length < 2) Some("no NUMA nodes available") else if (!numactl_available) Some("bad numactl tool") else None enabled && (warning match { case Some(s) => progress.echo_warning("Shuffling of CPU nodes is disabled: " + s); false case _ => true }) } class Nodes(enabled: Boolean = true) { private val available = nodes().zipWithIndex private var next_index = 0 def next(used: Int => Boolean = _ => false): Option[Int] = synchronized { if (!enabled || available.isEmpty) None else { val candidates = available.drop(next_index) ::: available.take(next_index) val (n, i) = candidates.find({ case (n, i) => i == next_index && !used(n) }) orElse candidates.find({ case (n, _) => !used(n) }) getOrElse candidates.head next_index = (i + 1) % available.length Some(n) } } } } 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,453 +1,453 @@ /* 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(_))) + 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 = + 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.empty + "=" + val prefs_syntax: Outer_Syntax = Outer_Syntax.empty + "=" trait Parser extends Parse.Parser { - val option_name = atom("option name", _.is_name) - val option_type = atom("option type", _.is_name) - val option_value = + 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) } } catch { case ERROR(msg) => error(msg + 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)(_ + _) } /* encode */ val encode: XML.Encode.T[Options] = (options => options.encode) /* Isabelle tool wrapper */ val isabelle_tool = Isabelle_Tool("options", "print Isabelle system options", 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")))(_ + _) else options0 (options1 /: more_options)(_ + _) } 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) }) /* 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) { 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.mkdirs(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/System/progress.scala b/src/Pure/System/progress.scala --- a/src/Pure/System/progress.scala +++ b/src/Pure/System/progress.scala @@ -1,87 +1,87 @@ /* Title: Pure/System/progress.scala Author: Makarius Progress context for system processes. */ package isabelle import java.io.{File => JFile} object Progress { sealed case class Theory(theory: String, session: String = "", percentage: Option[Int] = None) { def message: String = print_session + print_theory + print_percentage def print_session: String = if (session == "") "" else session + ": " def print_theory: String = "theory " + theory def print_percentage: String = percentage match { case None => "" case Some(p) => " " + p + "%" } } } class Progress { def echo(msg: String) {} def echo_if(cond: Boolean, msg: String) { if (cond) echo(msg) } def theory(theory: Progress.Theory) {} def nodes_status(nodes_status: Document_Status.Nodes_Status) {} def echo_warning(msg: String) { echo(Output.warning_text(msg)) } def echo_error_message(msg: String) { echo(Output.error_message_text(msg)) } def timeit[A](message: String = "", enabled: Boolean = true)(e: => A): A = - Timing.timeit(message, enabled, echo(_))(e) + Timing.timeit(message, enabled, echo)(e) def stopped: Boolean = false def interrupt_handler[A](e: => A): A = e def expose_interrupt() { if (stopped) throw Exn.Interrupt() } override def toString: String = if (stopped) "Progress(stopped)" else "Progress" def bash(script: String, cwd: JFile = null, env: Map[String, String] = Isabelle_System.settings(), redirect: Boolean = false, echo: Boolean = false, strict: Boolean = true): Process_Result = { Isabelle_System.bash(script, cwd = cwd, env = env, redirect = redirect, progress_stdout = echo_if(echo, _), progress_stderr = echo_if(echo, _), strict = strict) } } object No_Progress extends Progress class Console_Progress(verbose: Boolean = false, stderr: Boolean = false) extends Progress { override def echo(msg: String): Unit = Output.writeln(msg, stdout = !stderr, include_empty = true) override def theory(theory: Progress.Theory): Unit = if (verbose) echo(theory.message) @volatile private var is_stopped = false override def interrupt_handler[A](e: => A): A = POSIX_Interrupt.handler { is_stopped = true } { e } override def stopped: Boolean = { if (Thread.interrupted) is_stopped = true is_stopped } } class File_Progress(path: Path, verbose: Boolean = false) extends Progress { override def echo(msg: String): Unit = File.append(path, msg + "\n") override def theory(theory: Progress.Theory): Unit = if (verbose) echo(theory.message) override def toString: String = path.toString } diff --git a/src/Pure/Thy/export.scala b/src/Pure/Thy/export.scala --- a/src/Pure/Thy/export.scala +++ b/src/Pure/Thy/export.scala @@ -1,419 +1,419 @@ /* Title: Pure/Thy/export.scala Author: Makarius Manage theory exports: compressed blobs. */ package isabelle import scala.annotation.tailrec import scala.util.matching.Regex object Export { /* name structure */ def explode_name(s: String): List[String] = space_explode('/', s) def implode_name(elems: Iterable[String]): String = elems.mkString("/") /* SQL data model */ object Data { val session_name = SQL.Column.string("session_name").make_primary_key val theory_name = SQL.Column.string("theory_name").make_primary_key val name = SQL.Column.string("name").make_primary_key val executable = SQL.Column.bool("executable") val compressed = SQL.Column.bool("compressed") val body = SQL.Column.bytes("body") val table = SQL.Table("isabelle_exports", List(session_name, theory_name, name, executable, compressed, body)) def where_equal(session_name: String, theory_name: String = "", name: String = ""): SQL.Source = "WHERE " + Data.session_name.equal(session_name) + (if (theory_name == "") "" else " AND " + Data.theory_name.equal(theory_name)) + (if (name == "") "" else " AND " + Data.name.equal(name)) } def read_name(db: SQL.Database, session_name: String, theory_name: String, name: String): Boolean = { val select = Data.table.select(List(Data.name), Data.where_equal(session_name, theory_name, name)) db.using_statement(select)(stmt => stmt.execute_query().next()) } def read_names(db: SQL.Database, session_name: String, theory_name: String): List[String] = { val select = Data.table.select(List(Data.name), Data.where_equal(session_name, theory_name)) db.using_statement(select)(stmt => stmt.execute_query().iterator(res => res.string(Data.name)).toList) } def read_theory_names(db: SQL.Database, session_name: String): List[String] = { val select = Data.table.select(List(Data.theory_name), Data.where_equal(session_name), distinct = true) db.using_statement(select)(stmt => stmt.execute_query().iterator(_.string(Data.theory_name)).toList) } def read_theory_exports(db: SQL.Database, session_name: String): List[(String, String)] = { val select = Data.table.select(List(Data.theory_name, Data.name), Data.where_equal(session_name)) db.using_statement(select)(stmt => stmt.execute_query().iterator(res => (res.string(Data.theory_name), res.string(Data.name))).toList) } def message(msg: String, theory_name: String, name: String): String = msg + " " + quote(name) + " for theory " + quote(theory_name) def compound_name(a: String, b: String): String = a + ":" + b sealed case class Entry( session_name: String, theory_name: String, name: String, executable: Boolean, body: Future[(Boolean, Bytes)]) { override def toString: String = name def compound_name: String = Export.compound_name(theory_name, name) val name_elems: List[String] = explode_name(name) def name_extends(elems: List[String]): Boolean = name_elems.startsWith(elems) && name_elems != elems def text: String = uncompressed().text def uncompressed(cache: XZ.Cache = XZ.cache()): Bytes = { val (compressed, bytes) = body.join if (compressed) bytes.uncompress(cache = cache) else bytes } def uncompressed_yxml(cache: XZ.Cache = XZ.cache()): XML.Body = YXML.parse_body(UTF8.decode_permissive(uncompressed(cache = cache))) def write(db: SQL.Database) { val (compressed, bytes) = body.join db.using_statement(Data.table.insert())(stmt => { stmt.string(1) = session_name stmt.string(2) = theory_name stmt.string(3) = name stmt.bool(4) = executable stmt.bool(5) = compressed stmt.bytes(6) = bytes stmt.execute() }) } } def make_regex(pattern: String): Regex = { @tailrec def make(result: List[String], depth: Int, chs: List[Char]): Regex = chs match { case '*' :: '*' :: rest => make("[^:]*" :: result, depth, rest) case '*' :: rest => make("[^:/]*" :: result, depth, rest) case '?' :: rest => make("[^:/]" :: result, depth, rest) case '\\' :: c :: rest => make(("\\" + c) :: result, depth, rest) case '{' :: rest => make("(" :: result, depth + 1, rest) case ',' :: rest if depth > 0 => make("|" :: result, depth, rest) case '}' :: rest if depth > 0 => make(")" :: result, depth - 1, rest) case c :: rest if ".+()".contains(c) => make(("\\" + c) :: result, depth, rest) case c :: rest => make(c.toString :: result, depth, rest) case Nil => result.reverse.mkString.r } make(Nil, 0, pattern.toList) } def make_matcher(pattern: String): (String, String) => Boolean = { val regex = make_regex(pattern) (theory_name: String, name: String) => regex.pattern.matcher(compound_name(theory_name, name)).matches } def make_entry(session_name: String, args: Markup.Export.Args, body: Bytes, cache: XZ.Cache = XZ.cache()): Entry = { Entry(session_name, args.theory_name, args.name, args.executable, if (args.compress) Future.fork(body.maybe_compress(cache = cache)) else Future.value((false, body))) } def read_entry(db: SQL.Database, session_name: String, theory_name: String, name: String) : Option[Entry] = { val select = Data.table.select(List(Data.executable, Data.compressed, Data.body), Data.where_equal(session_name, theory_name, name)) db.using_statement(select)(stmt => { val res = stmt.execute_query() if (res.next()) { val executable = res.bool(Data.executable) val compressed = res.bool(Data.compressed) val body = res.bytes(Data.body) Some(Entry(session_name, theory_name, name, executable, Future.value(compressed, body))) } else None }) } def read_entry(dir: Path, session_name: String, theory_name: String, name: String): Option[Entry] = { val path = dir + Path.basic(theory_name) + Path.explode(name) if (path.is_file) { val executable = File.is_executable(path) val uncompressed = Bytes.read(path) Some(Entry(session_name, theory_name, name, executable, Future.value((false, uncompressed)))) } else None } /* database consumer thread */ def consumer(db: SQL.Database, cache: XZ.Cache = XZ.cache()): Consumer = new Consumer(db, cache) class Consumer private[Export](db: SQL.Database, cache: XZ.Cache) { private val errors = Synchronized[List[String]](Nil) private val consumer = Consumer_Thread.fork_bulk[(Entry, Boolean)](name = "export")( bulk = { case (entry, _) => entry.body.is_finished }, consume = (args: List[(Entry, Boolean)]) => { val results = db.transaction { for ((entry, strict) <- args) yield { if (read_name(db, entry.session_name, entry.theory_name, entry.name)) { if (strict) { val msg = message("Duplicate export", entry.theory_name, entry.name) errors.change(msg :: _) } Exn.Res(()) } else Exn.capture { entry.write(db) } } } (results, true) }) def apply(session_name: String, args: Markup.Export.Args, body: Bytes): Unit = consumer.send(make_entry(session_name, args, body, cache = cache) -> args.strict) def shutdown(close: Boolean = false): List[String] = { consumer.shutdown() if (close) db.close() errors.value.reverse } } /* abstract provider */ object Provider { def none: Provider = new Provider { def apply(export_name: String): Option[Entry] = None def focus(other_theory: String): Provider = this override def toString: String = "none" } def database(db: SQL.Database, session_name: String, theory_name: String): Provider = new Provider { def apply(export_name: String): Option[Entry] = read_entry(db, session_name, theory_name, export_name) def focus(other_theory: String): Provider = if (other_theory == theory_name) this else Provider.database(db, session_name, other_theory) override def toString: String = db.toString } def snapshot(snapshot: Document.Snapshot): Provider = new Provider { def apply(export_name: String): Option[Entry] = snapshot.exports_map.get(export_name) def focus(other_theory: String): Provider = if (other_theory == snapshot.node_name.theory) this else { val node_name = snapshot.version.nodes.theory_name(other_theory) getOrElse error("Bad theory " + quote(other_theory)) Provider.snapshot(snapshot.state.snapshot(node_name)) } override def toString: String = snapshot.toString } def directory(dir: Path, session_name: String, theory_name: String): Provider = new Provider { def apply(export_name: String): Option[Entry] = read_entry(dir, session_name, theory_name, export_name) def focus(other_theory: String): Provider = if (other_theory == theory_name) this else Provider.directory(dir, session_name, other_theory) override def toString: String = dir.toString } } trait Provider { def apply(export_name: String): Option[Entry] def uncompressed_yxml(export_name: String, cache: XZ.Cache = XZ.cache()): XML.Body = apply(export_name) match { case Some(entry) => entry.uncompressed_yxml(cache = cache) case None => Nil } def focus(other_theory: String): Provider } /* export to file-system */ def export_files( store: Sessions.Store, session_name: String, export_dir: Path, progress: Progress = No_Progress, export_prune: Int = 0, export_list: Boolean = false, export_patterns: List[String] = Nil) { using(store.open_database(session_name))(db => { db.transaction { val export_names = read_theory_exports(db, session_name) // list if (export_list) { (for ((theory_name, name) <- export_names) yield compound_name(theory_name, name)). - sorted.foreach(progress.echo(_)) + sorted.foreach(progress.echo) } // export if (export_patterns.nonEmpty) { val exports = (for { export_pattern <- export_patterns.iterator matcher = make_matcher(export_pattern) (theory_name, name) <- export_names if matcher(theory_name, name) } yield (theory_name, name)).toSet for { (theory_name, group) <- exports.toList.groupBy(_._1).toList.sortBy(_._1) name <- group.map(_._2).sorted entry <- read_entry(db, session_name, theory_name, name) } { val elems = theory_name :: space_explode('/', name) val path = if (elems.length < export_prune + 1) { error("Cannot prune path by " + export_prune + " element(s): " + Path.make(elems)) } else export_dir + Path.make(elems.drop(export_prune)) progress.echo("export " + path + (if (entry.executable) " (executable)" else "")) Isabelle_System.mkdirs(path.dir) Bytes.write(path, entry.uncompressed(cache = store.xz_cache)) File.set_executable(path, entry.executable) } } } }) } /* Isabelle tool wrapper */ - val default_export_dir = Path.explode("export") + val default_export_dir: Path = Path.explode("export") val isabelle_tool = Isabelle_Tool("export", "retrieve theory exports", args => { /* arguments */ var export_dir = default_export_dir var dirs: List[Path] = Nil var export_list = false var no_build = false var options = Options.init() var export_prune = 0 var export_patterns: List[String] = Nil val getopts = Getopts(""" Usage: isabelle export [OPTIONS] SESSION Options are: -O DIR output directory for exported files (default: """ + default_export_dir + """) -d DIR include session directory -l list exports -n no build of session -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) -p NUM prune path of exported files by NUM elements -x PATTERN extract files matching pattern (e.g. "*:**" for all) List or export theory exports for SESSION: named blobs produced by isabelle build. Option -l or -x is required; option -x may be repeated. The PATTERN language resembles glob patterns in the shell, with ? and * (both excluding ":" and "/"), ** (excluding ":"), and [abc] or [^abc], and variants {pattern1,pattern2,pattern3}. """, "O:" -> (arg => export_dir = Path.explode(arg)), "d:" -> (arg => dirs = dirs ::: List(Path.explode(arg))), "l" -> (_ => export_list = true), "n" -> (_ => no_build = true), "o:" -> (arg => options = options + arg), "p:" -> (arg => export_prune = Value.Int.parse(arg)), "x:" -> (arg => export_patterns ::= arg)) val more_args = getopts(args) val session_name = more_args match { case List(session_name) if export_list || export_patterns.nonEmpty => session_name case _ => getopts.usage() } val progress = new Console_Progress() /* build */ if (!no_build) { val rc = progress.interrupt_handler { Build.build_logic(options, session_name, progress = progress, dirs = dirs) } if (rc != 0) sys.exit(rc) } /* export files */ val store = Sessions.store(options) export_files(store, session_name, export_dir, progress = progress, export_prune = export_prune, export_list = export_list, export_patterns = export_patterns) }) } diff --git a/src/Pure/Thy/export_theory.scala b/src/Pure/Thy/export_theory.scala --- a/src/Pure/Thy/export_theory.scala +++ b/src/Pure/Thy/export_theory.scala @@ -1,757 +1,757 @@ /* Title: Pure/Thy/export_theory.scala Author: Makarius Export foundational theory content. */ package isabelle import scala.collection.immutable.SortedMap object Export_Theory { /** session content **/ sealed case class Session(name: String, theory_graph: Graph[String, Option[Theory]]) { override def toString: String = name def theory(theory_name: String): Option[Theory] = if (theory_graph.defined(theory_name)) theory_graph.get_node(theory_name) else None def theories: List[Theory] = - theory_graph.topological_order.flatMap(theory(_)) + theory_graph.topological_order.flatMap(theory) } def read_session( store: Sessions.Store, sessions_structure: Sessions.Structure, session_name: String, progress: Progress = No_Progress, cache: Term.Cache = Term.make_cache()): Session = { val thys = sessions_structure.build_requirements(List(session_name)).flatMap(session => using(store.open_database(session))(db => { db.transaction { for (theory <- Export.read_theory_names(db, session)) yield { progress.echo("Reading theory " + theory) read_theory(Export.Provider.database(db, session, theory), session, theory, cache = Some(cache)) } } })) val graph0 = (Graph.string[Option[Theory]] /: thys) { case (g, thy) => g.default_node(thy.name, Some(thy)) } val graph1 = (graph0 /: thys) { case (g0, thy) => (g0 /: thy.parents) { case (g1, parent) => g1.default_node(parent, None).add_edge_acyclic(parent, thy.name) } } Session(session_name, graph1) } /** theory content **/ val export_prefix: String = "theory/" val export_prefix_proofs: String = "proofs/" sealed case class Theory(name: String, parents: List[String], types: List[Type], consts: List[Const], axioms: List[Axiom], thms: List[Thm], classes: List[Class], locales: List[Locale], locale_dependencies: List[Locale_Dependency], classrel: List[Classrel], arities: List[Arity], constdefs: List[Constdef], typedefs: List[Typedef], datatypes: List[Datatype], spec_rules: List[Spec_Rule]) { override def toString: String = name def entity_iterator: Iterator[Entity] = types.iterator.map(_.entity) ++ consts.iterator.map(_.entity) ++ axioms.iterator.map(_.entity) ++ thms.iterator.map(_.entity) ++ classes.iterator.map(_.entity) ++ locales.iterator.map(_.entity) ++ locale_dependencies.iterator.map(_.entity) def cache(cache: Term.Cache): Theory = Theory(cache.string(name), parents.map(cache.string(_)), types.map(_.cache(cache)), consts.map(_.cache(cache)), axioms.map(_.cache(cache)), thms.map(_.cache(cache)), classes.map(_.cache(cache)), locales.map(_.cache(cache)), locale_dependencies.map(_.cache(cache)), classrel.map(_.cache(cache)), arities.map(_.cache(cache)), constdefs.map(_.cache(cache)), typedefs.map(_.cache(cache)), datatypes.map(_.cache(cache)), spec_rules.map(_.cache(cache))) } def read_theory(provider: Export.Provider, session_name: String, theory_name: String, cache: Option[Term.Cache] = None): Theory = { val parents = if (theory_name == Thy_Header.PURE) Nil else { provider(export_prefix + "parents") match { case Some(entry) => split_lines(entry.uncompressed().text) case None => error("Missing theory export in session " + quote(session_name) + ": " + quote(theory_name)) } } val theory = Theory(theory_name, parents, read_types(provider), read_consts(provider), read_axioms(provider), read_thms(provider), read_classes(provider), read_locales(provider), read_locale_dependencies(provider), read_classrel(provider), read_arities(provider), read_constdefs(provider), read_typedefs(provider), read_datatypes(provider), read_spec_rules(provider)) if (cache.isDefined) theory.cache(cache.get) else theory } def read_pure[A](store: Sessions.Store, read: (Export.Provider, String, String) => A): A = { val session_name = Thy_Header.PURE val theory_name = Thy_Header.PURE using(store.open_database(session_name))(db => { db.transaction { read(Export.Provider.database(db, session_name, theory_name), session_name, theory_name) } }) } def read_pure_theory(store: Sessions.Store, cache: Option[Term.Cache] = None): Theory = read_pure(store, read_theory(_, _, _, cache = cache)) def read_pure_proof( store: Sessions.Store, id: Thm_Id, cache: Option[Term.Cache] = None): Option[Proof] = read_pure(store, (provider, _, _) => read_proof(provider, id, cache = cache)) /* entities */ object Kind extends Enumeration { val TYPE = Value("type") val CONST = Value("const") val AXIOM = Value("axiom") val THM = Value("thm") val PROOF = Value("proof") val CLASS = Value("class") val LOCALE = Value("locale") val LOCALE_DEPENDENCY = Value("locale_dependency") val DOCUMENT_HEADING = Value("document_heading") val DOCUMENT_TEXT = Value("document_text") val PROOF_TEXT = Value("proof_text") } sealed case class Entity( kind: Kind.Value, name: String, xname: String, pos: Position.T, id: Option[Long], serial: Long) { override def toString: String = kind.toString + " " + quote(name) def cache(cache: Term.Cache): Entity = Entity(kind, cache.string(name), cache.string(xname), cache.position(pos), id, serial) } def decode_entity(kind: Kind.Value, tree: XML.Tree): (Entity, XML.Body) = { def err(): Nothing = throw new XML.XML_Body(List(tree)) tree match { case XML.Elem(Markup(Markup.ENTITY, props), body) => val name = Markup.Name.unapply(props) getOrElse err() val xname = Markup.XName.unapply(props) getOrElse err() val pos = props.filter({ case (a, _) => Markup.POSITION_PROPERTIES(a) && a != Markup.ID }) val id = Position.Id.unapply(props) val serial = Markup.Serial.unapply(props) getOrElse err() (Entity(kind, name, xname, pos, id, serial), body) case _ => err() } } /* approximative syntax */ object Assoc extends Enumeration { val NO_ASSOC, LEFT_ASSOC, RIGHT_ASSOC = Value } sealed abstract class Syntax case object No_Syntax extends Syntax case class Prefix(delim: String) extends Syntax case class Infix(assoc: Assoc.Value, delim: String, pri: Int) extends Syntax def decode_syntax: XML.Decode.T[Syntax] = XML.Decode.variant(List( { case (Nil, Nil) => No_Syntax }, { case (List(delim), Nil) => Prefix(delim) }, { case (Nil, body) => import XML.Decode._ val (ass, delim, pri) = triple(int, string, int)(body) Infix(Assoc(ass), delim, pri) })) /* types */ sealed case class Type( entity: Entity, syntax: Syntax, args: List[String], abbrev: Option[Term.Typ]) { def cache(cache: Term.Cache): Type = Type(entity.cache(cache), syntax, args.map(cache.string(_)), abbrev.map(cache.typ(_))) } def read_types(provider: Export.Provider): List[Type] = provider.uncompressed_yxml(export_prefix + "types").map((tree: XML.Tree) => { val (entity, body) = decode_entity(Kind.TYPE, tree) val (syntax, args, abbrev) = { import XML.Decode._ triple(decode_syntax, list(string), option(Term_XML.Decode.typ))(body) } Type(entity, syntax, args, abbrev) }) /* consts */ sealed case class Const( entity: Entity, syntax: Syntax, typargs: List[String], typ: Term.Typ, abbrev: Option[Term.Term], propositional: Boolean) { def cache(cache: Term.Cache): Const = Const(entity.cache(cache), syntax, typargs.map(cache.string(_)), cache.typ(typ), abbrev.map(cache.term(_)), propositional) } def read_consts(provider: Export.Provider): List[Const] = provider.uncompressed_yxml(export_prefix + "consts").map((tree: XML.Tree) => { val (entity, body) = decode_entity(Kind.CONST, tree) val (syntax, (typargs, (typ, (abbrev, propositional)))) = { import XML.Decode._ pair(decode_syntax, pair(list(string), pair(Term_XML.Decode.typ, pair(option(Term_XML.Decode.term), bool))))(body) } Const(entity, syntax, typargs, typ, abbrev, propositional) }) /* axioms */ sealed case class Prop( typargs: List[(String, Term.Sort)], args: List[(String, Term.Typ)], term: Term.Term) { def cache(cache: Term.Cache): Prop = Prop( typargs.map({ case (name, sort) => (cache.string(name), cache.sort(sort)) }), args.map({ case (name, typ) => (cache.string(name), cache.typ(typ)) }), cache.term(term)) } def decode_prop(body: XML.Body): Prop = { val (typargs, args, t) = { import XML.Decode._ import Term_XML.Decode._ triple(list(pair(string, sort)), list(pair(string, typ)), term)(body) } Prop(typargs, args, t) } sealed case class Axiom(entity: Entity, prop: Prop) { def cache(cache: Term.Cache): Axiom = Axiom(entity.cache(cache), prop.cache(cache)) } def read_axioms(provider: Export.Provider): List[Axiom] = provider.uncompressed_yxml(export_prefix + "axioms").map((tree: XML.Tree) => { val (entity, body) = decode_entity(Kind.AXIOM, tree) val prop = decode_prop(body) Axiom(entity, prop) }) /* theorems */ sealed case class Thm_Id(serial: Long, theory_name: String) { def pure: Boolean = theory_name == Thy_Header.PURE } sealed case class Thm( entity: Entity, prop: Prop, deps: List[String], proof: Term.Proof) { def cache(cache: Term.Cache): Thm = Thm( entity.cache(cache), prop.cache(cache), deps.map(cache.string _), cache.proof(proof)) } def read_thms(provider: Export.Provider): List[Thm] = provider.uncompressed_yxml(export_prefix + "thms").map((tree: XML.Tree) => { val (entity, body) = decode_entity(Kind.THM, tree) val (prop, deps, prf) = { import XML.Decode._ import Term_XML.Decode._ triple(decode_prop, list(string), proof)(body) } Thm(entity, prop, deps, prf) }) sealed case class Proof( typargs: List[(String, Term.Sort)], args: List[(String, Term.Typ)], term: Term.Term, proof: Term.Proof) { def prop: Prop = Prop(typargs, args, term) def cache(cache: Term.Cache): Proof = Proof( typargs.map({ case (name, sort) => (cache.string(name), cache.sort(sort)) }), args.map({ case (name, typ) => (cache.string(name), cache.typ(typ)) }), cache.term(term), cache.proof(proof)) } def read_proof( provider: Export.Provider, id: Thm_Id, cache: Option[Term.Cache] = None): Option[Proof] = { for { entry <- provider.focus(id.theory_name)(export_prefix_proofs + id.serial) } yield { val body = entry.uncompressed_yxml() val (typargs, (args, (prop_body, proof_body))) = { import XML.Decode._ import Term_XML.Decode._ pair(list(pair(string, sort)), pair(list(pair(string, typ)), pair(x => x, x => x)))(body) } val env = args.toMap val prop = Term_XML.Decode.term_env(env)(prop_body) val proof = Term_XML.Decode.proof_env(env)(proof_body) val result = Proof(typargs, args, prop, proof) cache.map(result.cache(_)) getOrElse result } } def read_proof_boxes( store: Sessions.Store, provider: Export.Provider, proof: Term.Proof, suppress: Thm_Id => Boolean = _ => false, cache: Option[Term.Cache] = None): List[(Thm_Id, Proof)] = { var seen = Set.empty[Long] var result = SortedMap.empty[Long, (Thm_Id, Proof)] def boxes(context: Option[(Long, Term.Proof)], prf: Term.Proof) { prf match { case Term.Abst(_, _, p) => boxes(context, p) case Term.AbsP(_, _, p) => boxes(context, p) case Term.Appt(p, _) => boxes(context, p) case Term.AppP(p, q) => boxes(context, p); boxes(context, q) case thm: Term.PThm if !seen(thm.serial) => seen += thm.serial val id = Thm_Id(thm.serial, thm.theory_name) if (!suppress(id)) { val read = if (id.pure) Export_Theory.read_pure_proof(store, id, cache = cache) else Export_Theory.read_proof(provider, id, cache = cache) read match { case Some(p) => result += (thm.serial -> (id -> p)) boxes(Some((thm.serial, p.proof)), p.proof) case None => error("Missing proof " + thm.serial + " (theory " + quote (thm.theory_name) + ")" + (context match { case None => "" case Some((i, p)) => " in proof " + i + ":\n" + p })) } } case _ => } } boxes(None, proof) result.iterator.map(_._2).toList } /* type classes */ sealed case class Class( entity: Entity, params: List[(String, Term.Typ)], axioms: List[Prop]) { def cache(cache: Term.Cache): Class = Class(entity.cache(cache), params.map({ case (name, typ) => (cache.string(name), cache.typ(typ)) }), axioms.map(_.cache(cache))) } def read_classes(provider: Export.Provider): List[Class] = provider.uncompressed_yxml(export_prefix + "classes").map((tree: XML.Tree) => { val (entity, body) = decode_entity(Kind.CLASS, tree) val (params, axioms) = { import XML.Decode._ import Term_XML.Decode._ pair(list(pair(string, typ)), list(decode_prop))(body) } Class(entity, params, axioms) }) /* locales */ sealed case class Locale( entity: Entity, typargs: List[(String, Term.Sort)], args: List[((String, Term.Typ), Syntax)], axioms: List[Prop]) { def cache(cache: Term.Cache): Locale = Locale(entity.cache(cache), typargs.map({ case (name, sort) => (cache.string(name), cache.sort(sort)) }), args.map({ case ((name, typ), syntax) => ((cache.string(name), cache.typ(typ)), syntax) }), axioms.map(_.cache(cache))) } def read_locales(provider: Export.Provider): List[Locale] = provider.uncompressed_yxml(export_prefix + "locales").map((tree: XML.Tree) => { val (entity, body) = decode_entity(Kind.LOCALE, tree) val (typargs, args, axioms) = { import XML.Decode._ import Term_XML.Decode._ triple(list(pair(string, sort)), list(pair(pair(string, typ), decode_syntax)), list(decode_prop))(body) } Locale(entity, typargs, args, axioms) }) /* locale dependencies */ sealed case class Locale_Dependency( entity: Entity, source: String, target: String, prefix: List[(String, Boolean)], subst_types: List[((String, Term.Sort), Term.Typ)], subst_terms: List[((String, Term.Typ), Term.Term)]) { def cache(cache: Term.Cache): Locale_Dependency = Locale_Dependency(entity.cache(cache), cache.string(source), cache.string(target), prefix.map({ case (name, mandatory) => (cache.string(name), mandatory) }), subst_types.map({ case ((a, s), ty) => ((cache.string(a), cache.sort(s)), cache.typ(ty)) }), subst_terms.map({ case ((x, ty), t) => ((cache.string(x), cache.typ(ty)), cache.term(t)) })) def is_inclusion: Boolean = subst_types.isEmpty && subst_terms.isEmpty } def read_locale_dependencies(provider: Export.Provider): List[Locale_Dependency] = provider.uncompressed_yxml(export_prefix + "locale_dependencies").map((tree: XML.Tree) => { val (entity, body) = decode_entity(Kind.LOCALE_DEPENDENCY, tree) val (source, (target, (prefix, (subst_types, subst_terms)))) = { import XML.Decode._ import Term_XML.Decode._ pair(string, pair(string, pair(list(pair(string, bool)), pair(list(pair(pair(string, sort), typ)), list(pair(pair(string, typ), term))))))(body) } Locale_Dependency(entity, source, target, prefix, subst_types, subst_terms) }) /* sort algebra */ sealed case class Classrel(class1: String, class2: String, prop: Prop) { def cache(cache: Term.Cache): Classrel = Classrel(cache.string(class1), cache.string(class2), prop.cache(cache)) } def read_classrel(provider: Export.Provider): List[Classrel] = { val body = provider.uncompressed_yxml(export_prefix + "classrel") val classrel = { import XML.Decode._ list(pair(decode_prop, pair(string, string)))(body) } for ((prop, (c1, c2)) <- classrel) yield Classrel(c1, c2, prop) } sealed case class Arity(type_name: String, domain: List[Term.Sort], codomain: String, prop: Prop) { def cache(cache: Term.Cache): Arity = Arity(cache.string(type_name), domain.map(cache.sort(_)), cache.string(codomain), prop.cache(cache)) } def read_arities(provider: Export.Provider): List[Arity] = { val body = provider.uncompressed_yxml(export_prefix + "arities") val arities = { import XML.Decode._ import Term_XML.Decode._ list(pair(decode_prop, triple(string, list(sort), string)))(body) } for ((prop, (a, b, c)) <- arities) yield Arity(a, b, c, prop) } /* Pure constdefs */ sealed case class Constdef(name: String, axiom_name: String) { def cache(cache: Term.Cache): Constdef = Constdef(cache.string(name), cache.string(axiom_name)) } def read_constdefs(provider: Export.Provider): List[Constdef] = { val body = provider.uncompressed_yxml(export_prefix + "constdefs") val constdefs = { import XML.Decode._ list(pair(string, string))(body) } for ((name, axiom_name) <- constdefs) yield Constdef(name, axiom_name) } /* HOL typedefs */ sealed case class Typedef(name: String, rep_type: Term.Typ, abs_type: Term.Typ, rep_name: String, abs_name: String, axiom_name: String) { def cache(cache: Term.Cache): Typedef = Typedef(cache.string(name), cache.typ(rep_type), cache.typ(abs_type), cache.string(rep_name), cache.string(abs_name), cache.string(axiom_name)) } def read_typedefs(provider: Export.Provider): List[Typedef] = { val body = provider.uncompressed_yxml(export_prefix + "typedefs") val typedefs = { import XML.Decode._ import Term_XML.Decode._ list(pair(string, pair(typ, pair(typ, pair(string, pair(string, string))))))(body) } for { (name, (rep_type, (abs_type, (rep_name, (abs_name, axiom_name))))) <- typedefs } yield Typedef(name, rep_type, abs_type, rep_name, abs_name, axiom_name) } /* HOL datatypes */ sealed case class Datatype( pos: Position.T, name: String, co: Boolean, typargs: List[(String, Term.Sort)], typ: Term.Typ, constructors: List[(Term.Term, Term.Typ)]) { def id: Option[Long] = Position.Id.unapply(pos) def cache(cache: Term.Cache): Datatype = Datatype( cache.position(pos), cache.string(name), co, typargs.map({ case (name, sort) => (cache.string(name), cache.sort(sort)) }), cache.typ(typ), constructors.map({ case (term, typ) => (cache.term(term), cache.typ(typ)) })) } def read_datatypes(provider: Export.Provider): List[Datatype] = { val body = provider.uncompressed_yxml(export_prefix + "datatypes") val datatypes = { import XML.Decode._ import Term_XML.Decode._ list(pair(properties, pair(string, pair(bool, pair(list(pair(string, sort)), pair(typ, list(pair(term, typ))))))))(body) } for ((pos, (name, (co, (typargs, (typ, constructors))))) <- datatypes) yield Datatype(pos, name, co, typargs, typ, constructors) } /* Pure spec rules */ sealed abstract class Recursion { def cache(cache: Term.Cache): Recursion = this match { case Primrec(types) => Primrec(types.map(cache.string)) case Primcorec(types) => Primcorec(types.map(cache.string)) case _ => this } } case class Primrec(types: List[String]) extends Recursion case object Recdef extends Recursion case class Primcorec(types: List[String]) extends Recursion case object Corec extends Recursion case object Unknown_Recursion extends Recursion val decode_recursion: XML.Decode.T[Recursion] = { import XML.Decode._ variant(List( { case (Nil, a) => Primrec(list(string)(a)) }, { case (Nil, Nil) => Recdef }, { case (Nil, a) => Primcorec(list(string)(a)) }, { case (Nil, Nil) => Corec }, { case (Nil, Nil) => Unknown_Recursion })) } sealed abstract class Rough_Classification { def is_equational: Boolean = this.isInstanceOf[Equational] def is_inductive: Boolean = this == Inductive def is_co_inductive: Boolean = this == Co_Inductive def is_relational: Boolean = is_inductive || is_co_inductive def is_unknown: Boolean = this == Unknown def cache(cache: Term.Cache): Rough_Classification = this match { case Equational(recursion) => Equational(recursion.cache(cache)) case _ => this } } case class Equational(recursion: Recursion) extends Rough_Classification case object Inductive extends Rough_Classification case object Co_Inductive extends Rough_Classification case object Unknown extends Rough_Classification val decode_rough_classification: XML.Decode.T[Rough_Classification] = { import XML.Decode._ variant(List( { case (Nil, a) => Equational(decode_recursion(a)) }, { case (Nil, Nil) => Inductive }, { case (Nil, Nil) => Co_Inductive }, { case (Nil, Nil) => Unknown })) } sealed case class Spec_Rule( pos: Position.T, name: String, rough_classification: Rough_Classification, typargs: List[(String, Term.Sort)], args: List[(String, Term.Typ)], terms: List[(Term.Term, Term.Typ)], rules: List[Term.Term]) { def id: Option[Long] = Position.Id.unapply(pos) def cache(cache: Term.Cache): Spec_Rule = Spec_Rule( cache.position(pos), cache.string(name), rough_classification.cache(cache), typargs.map({ case (name, sort) => (cache.string(name), cache.sort(sort)) }), args.map({ case (name, typ) => (cache.string(name), cache.typ(typ)) }), terms.map({ case (term, typ) => (cache.term(term), cache.typ(typ)) }), rules.map(cache.term(_))) } def read_spec_rules(provider: Export.Provider): List[Spec_Rule] = { val body = provider.uncompressed_yxml(export_prefix + "spec_rules") val spec_rules = { import XML.Decode._ import Term_XML.Decode._ list( pair(properties, pair(string, pair(decode_rough_classification, pair(list(pair(string, sort)), pair(list(pair(string, typ)), pair(list(pair(term, typ)), list(term))))))))(body) } for ((pos, (name, (rough_classification, (typargs, (args, (terms, rules)))))) <- spec_rules) yield Spec_Rule(pos, name, rough_classification, typargs, args, terms, rules) } } diff --git a/src/Pure/Thy/file_format.scala b/src/Pure/Thy/file_format.scala --- a/src/Pure/Thy/file_format.scala +++ b/src/Pure/Thy/file_format.scala @@ -1,97 +1,97 @@ /* Title: Pure/Thy/file_format.scala Author: Makarius Support for user-defined file formats. */ package isabelle object File_Format { sealed case class Theory_Context(name: Document.Node.Name, content: String) /* environment */ def environment(): Environment = new Environment(Isabelle_System.init_classes[File_Format]("ISABELLE_FILE_FORMATS")) final class Environment private [File_Format](file_formats: List[File_Format]) { override def toString: String = file_formats.mkString("File_Format.Environment(", ",", ")") def get(name: String): Option[File_Format] = file_formats.find(_.detect(name)) def get(name: Document.Node.Name): Option[File_Format] = get(name.node) def get_theory(name: Document.Node.Name): Option[File_Format] = get(name.theory) def is_theory(name: Document.Node.Name): Boolean = get_theory(name).isDefined def start_session(session: isabelle.Session): Session = new Session(file_formats.map(_.start(session)).filterNot(_ == No_Agent)) } /* session */ final class Session private[File_Format](agents: List[Agent]) { override def toString: String = agents.mkString("File_Format.Session(", ", ", ")") def prover_options(options: Options): Options = (options /: agents)({ case (opts, agent) => agent.prover_options(opts) }) def stop_session { agents.foreach(_.stop) } } trait Agent { def prover_options(options: Options): Options = options def stop {} } object No_Agent extends Agent } trait File_Format { def format_name: String - override def toString = format_name + override def toString: String = format_name def file_ext: String def detect(name: String): Boolean = name.endsWith("." + file_ext) /* implicit theory context: name and content */ def theory_suffix: String = "" def theory_content(name: String): String = "" def make_theory_name(resources: Resources, name: Document.Node.Name): Option[Document.Node.Name] = { for { (_, ext_name) <- Thy_Header.split_file_name(name.node) if detect(ext_name) && theory_suffix.nonEmpty } yield { val thy_node = resources.append(name.node, Path.explode(theory_suffix)) Document.Node.Name(thy_node, name.master_dir, ext_name) } } def make_theory_content(resources: Resources, thy_name: Document.Node.Name): Option[String] = { for { (prefix, suffix) <- Thy_Header.split_file_name(thy_name.node) if suffix == theory_suffix (_, ext_name) <- Thy_Header.split_file_name(prefix) if detect(ext_name) s <- proper_string(theory_content(ext_name)) } yield s } def make_preview(snapshot: Document.Snapshot): Option[Present.Preview] = None /* PIDE session agent */ def start(session: isabelle.Session): File_Format.Agent = File_Format.No_Agent } diff --git a/src/Pure/Thy/html.scala b/src/Pure/Thy/html.scala --- a/src/Pure/Thy/html.scala +++ b/src/Pure/Thy/html.scala @@ -1,401 +1,401 @@ /* Title: Pure/Thy/html.scala Author: Makarius HTML presentation elements. */ package isabelle object HTML { /* output text with control symbols */ private val control = Map( Symbol.sub -> "sub", Symbol.sub_decoded -> "sub", Symbol.sup -> "sup", Symbol.sup_decoded -> "sup", Symbol.bold -> "b", Symbol.bold_decoded -> "b") private val control_block = Map( Symbol.bsub -> "", Symbol.bsub_decoded -> "", Symbol.esub -> "", Symbol.esub_decoded -> "", Symbol.bsup -> "", Symbol.bsup_decoded -> "", Symbol.esup -> "", Symbol.esup_decoded -> "") def is_control(sym: Symbol.Symbol): Boolean = control.isDefinedAt(sym) def output_char_permissive(c: Char, s: StringBuilder) { c match { case '<' => s ++= "<" case '>' => s ++= ">" case '&' => s ++= "&" case _ => s += c } } def output(text: String, s: StringBuilder, hidden: Boolean, permissive: Boolean) { def output_char(c: Char): Unit = if (permissive) output_char_permissive(c, s) else XML.output_char(c, s) def output_string(str: String): Unit = str.iterator.foreach(output_char) def output_hidden(body: => Unit): Unit = if (hidden) { s ++= ""; body; s ++= "" } def output_symbol(sym: Symbol.Symbol): Unit = if (sym != "") { control_block.get(sym) match { case Some(html) if html.startsWith(" s ++= html; output_hidden(output_string(sym)) case Some(html) => output_hidden(output_string(sym)); s ++= html case None => if (hidden && Symbol.is_control_encoded(sym)) { output_hidden(output_string(Symbol.control_prefix)) s ++= "" output_string(Symbol.control_name(sym).get) s ++= "" output_hidden(output_string(Symbol.control_suffix)) } else output_string(sym) } } var ctrl = "" for (sym <- Symbol.iterator(text)) { if (is_control(sym)) { output_symbol(ctrl); ctrl = sym } else { control.get(ctrl) match { case Some(elem) if Symbol.is_controllable(sym) => output_hidden(output_symbol(ctrl)) s += '<'; s ++= elem; s += '>' output_symbol(sym) s ++= "' case _ => output_symbol(ctrl) output_symbol(sym) } ctrl = "" } } output_symbol(ctrl) } def output(text: String): String = Library.make_string(output(text, _, hidden = false, permissive = true)) /* output XML as HTML */ private val structural_elements = Set("head", "body", "meta", "div", "pre", "p", "title", "h1", "h2", "h3", "h4", "h5", "h6", "ul", "ol", "dl", "li", "dt", "dd") def output(body: XML.Body, s: StringBuilder, hidden: Boolean, structural: Boolean) { def elem(markup: Markup) { s ++= markup.name for ((a, b) <- markup.properties) { s += ' ' s ++= a s += '=' s += '"' output(b, s, hidden = hidden, permissive = false) s += '"' } } def tree(t: XML.Tree): Unit = t match { case XML.Elem(markup, Nil) => s += '<'; elem(markup); s ++= "/>" case XML.Elem(markup, ts) => if (structural && structural_elements(markup.name)) s += '\n' s += '<'; elem(markup); s += '>' ts.foreach(tree) s ++= "' if (structural && structural_elements(markup.name)) s += '\n' case XML.Text(txt) => output(txt, s, hidden = hidden, permissive = true) } body.foreach(tree) } def output(body: XML.Body, hidden: Boolean, structural: Boolean): String = Library.make_string(output(body, _, hidden, structural)) def output(tree: XML.Tree, hidden: Boolean, structural: Boolean): String = output(List(tree), hidden, structural) /* attributes */ class Attribute(val name: String, value: String) { def xml: XML.Attribute = name -> value def apply(elem: XML.Elem): XML.Elem = elem + xml } def id(s: String): Attribute = new Attribute("id", s) def class_(name: String): Attribute = new Attribute("class", name) def width(w: Int): Attribute = new Attribute("width", w.toString) def height(h: Int): Attribute = new Attribute("height", h.toString) def size(w: Int, h: Int)(elem: XML.Elem): XML.Elem = width(w)(height(h)(elem)) /* structured markup operators */ def text(txt: String): XML.Body = if (txt.isEmpty) Nil else List(XML.Text(txt)) val no_text: XML.Tree = XML.Text("") val break: XML.Body = List(XML.elem("br")) class Operator(val name: String) { def apply(body: XML.Body): XML.Elem = XML.elem(name, body) def apply(att: Attribute, body: XML.Body): XML.Elem = att(apply(body)) def apply(c: String, body: XML.Body): XML.Elem = apply(class_(c), body) } class Heading(name: String) extends Operator(name) { def apply(txt: String): XML.Elem = super.apply(text(txt)) def apply(att: Attribute, txt: String): XML.Elem = super.apply(att, text(txt)) def apply(c: String, txt: String): XML.Elem = super.apply(c, text(txt)) } val div = new Operator("div") val span = new Operator("span") val pre = new Operator("pre") val par = new Operator("p") val emph = new Operator("em") val bold = new Operator("b") val code = new Operator("code") val item = new Operator("li") val list = new Operator("ul") val enum = new Operator("ol") val descr = new Operator("dl") val dt = new Operator("dt") val dd = new Operator("dd") val title = new Heading("title") val chapter = new Heading("h1") val section = new Heading("h2") val subsection = new Heading("h3") val subsubsection = new Heading("h4") val paragraph = new Heading("h5") val subparagraph = new Heading("h6") def itemize(items: List[XML.Body]): XML.Elem = list(items.map(item(_))) def enumerate(items: List[XML.Body]): XML.Elem = enum(items.map(item(_))) def description(items: List[(XML.Body, XML.Body)]): XML.Elem = descr(items.flatMap({ case (x, y) => List(dt(x), dd(y)) })) def link(href: String, body: XML.Body = Nil): XML.Elem = XML.Elem(Markup("a", List("href" -> href)), if (body.isEmpty) text(href) else body) def image(src: String, alt: String = ""): XML.Elem = XML.Elem(Markup("img", List("src" -> src) ::: proper_string(alt).map("alt" -> _).toList), Nil) def source(body: XML.Body): XML.Elem = pre("source", body) def source(src: String): XML.Elem = source(text(src)) def style(s: String): XML.Elem = XML.elem("style", text(s)) def style_file(href: String): XML.Elem = XML.Elem(Markup("link", List("rel" -> "stylesheet", "type" -> "text/css", "href" -> href)), Nil) def style_file(path: Path): XML.Elem = style_file(Url.print_file(path.file)) def script(s: String): XML.Elem = XML.elem("script", text(s)) def script_file(href: String): XML.Elem = XML.Elem(Markup("script", List("src" -> href)), Nil) def script_file(path: Path): XML.Elem = script_file(Url.print_file(path.file)) /* messages */ // background val writeln_message: Attribute = class_("writeln_message") val warning_message: Attribute = class_("warning_message") val error_message: Attribute = class_("error_message") // underline val writeln: Attribute = class_("writeln") val warning: Attribute = class_("warning") val error: Attribute = class_("error") /* tooltips */ def tooltip(item: XML.Body, tip: XML.Body): XML.Elem = span(item ::: List(div("tooltip", tip))) def tooltip_errors(item: XML.Body, msgs: List[XML.Body]): XML.Elem = HTML.error(tooltip(item, msgs.map(msg => error_message(pre(msg))))) /* GUI elements */ object GUI { private def optional_value(text: String): XML.Attributes = proper_string(text).map(a => "value" -> a).toList private def optional_name(name: String): XML.Attributes = proper_string(name).map(a => "name" -> a).toList private def optional_title(tooltip: String): XML.Attributes = proper_string(tooltip).map(a => "title" -> a).toList private def optional_submit(submit: Boolean): XML.Attributes = if (submit) List("onChange" -> "this.form.submit()") else Nil private def optional_checked(selected: Boolean): XML.Attributes = if (selected) List("checked" -> "") else Nil private def optional_action(action: String): XML.Attributes = proper_string(action).map(a => "action" -> a).toList private def optional_onclick(script: String): XML.Attributes = proper_string(script).map(a => "onclick" -> a).toList private def optional_onchange(script: String): XML.Attributes = proper_string(script).map(a => "onchange" -> a).toList def button(body: XML.Body, name: String = "", tooltip: String = "", submit: Boolean = false, script: String = ""): XML.Elem = XML.Elem( Markup("button", List("type" -> (if (submit) "submit" else "button"), "value" -> "true") ::: optional_name(name) ::: optional_title(tooltip) ::: optional_onclick(script)), body) def checkbox(body: XML.Body, name: String = "", tooltip: String = "", submit: Boolean = false, selected: Boolean = false, script: String = ""): XML.Elem = XML.elem("label", XML.elem( Markup("input", List("type" -> "checkbox", "value" -> "true") ::: optional_name(name) ::: optional_title(tooltip) ::: optional_submit(submit) ::: optional_checked(selected) ::: optional_onchange(script))) :: body) def text_field(columns: Int = 0, text: String = "", name: String = "", tooltip: String = "", submit: Boolean = false, script: String = ""): XML.Elem = XML.elem(Markup("input", List("type" -> "text") ::: (if (columns > 0) List("size" -> columns.toString) else Nil) ::: optional_value(text) ::: optional_name(name) ::: optional_title(tooltip) ::: optional_submit(submit) ::: optional_onchange(script))) def parameter(text: String = "", name: String = ""): XML.Elem = XML.elem( Markup("input", List("type" -> "hidden") ::: optional_value(text) ::: optional_name(name))) def form(body: XML.Body, name: String = "", action: String = "", http_post: Boolean = false) : XML.Elem = XML.Elem( Markup("form", optional_name(name) ::: optional_action(action) ::: (if (http_post) List("method" -> "post") else Nil)), body) } /* GUI layout */ object Wrap_Panel { object Alignment extends Enumeration { val left, right, center = Value } def apply(contents: List[XML.Elem], name: String = "", action: String = "", alignment: Alignment.Value = Alignment.right): XML.Elem = { val body = Library.separate(XML.Text(" "), contents) GUI.form(List(div(body) + ("style" -> ("text-align: " + alignment))), name = name, action = action) } } /* document */ val header: String = XML.header + """ """ val head_meta: XML.Elem = XML.Elem(Markup("meta", List("http-equiv" -> "Content-Type", "content" -> "text/html; charset=utf-8")), Nil) def output_document(head: XML.Body, body: XML.Body, css: String = "isabelle.css", hidden: Boolean = true, structural: Boolean = true): String = { cat_lines( List(header, output( XML.elem("head", head_meta :: (if (css == "") Nil else List(style_file(css))) ::: head), hidden = hidden, structural = structural), output(XML.elem("body", body), hidden = hidden, structural = structural))) } /* fonts */ def fonts_url(): String => String = (for (entry <- Isabelle_Fonts.fonts(hidden = true)) yield (entry.path.file_name -> Url.print_file(entry.path.file))).toMap def fonts_dir(prefix: String)(ttf_name: String): String = prefix + "/" + ttf_name def fonts_css(make_url: String => String = fonts_url()): String = { def font_face(entry: Isabelle_Fonts.Entry): String = cat_lines( List( "@font-face {", " font-family: '" + entry.family + "';", " src: url('" + make_url(entry.path.file_name) + "') format('truetype');") ::: (if (entry.is_bold) List(" font-weight: bold;") else Nil) ::: (if (entry.is_italic) List(" font-style: italic;") else Nil) ::: List("}")) - ("/* Isabelle fonts */" :: Isabelle_Fonts.fonts(hidden = true).map(font_face(_))) + ("/* Isabelle fonts */" :: Isabelle_Fonts.fonts(hidden = true).map(font_face)) .mkString("", "\n\n", "\n") } /* document directory */ def isabelle_css: Path = Path.explode("~~/etc/isabelle.css") def write_isabelle_css(dir: Path, make_url: String => String = fonts_dir("fonts")) { File.write(dir + isabelle_css.base, fonts_css(make_url) + "\n\n" + File.read(isabelle_css)) } def init_dir(dir: Path) { Isabelle_System.mkdirs(dir) write_isabelle_css(dir) } def write_document(dir: Path, name: String, head: XML.Body, body: XML.Body, css: String = isabelle_css.file_name, hidden: Boolean = true, structural: Boolean = true) { init_dir(dir) File.write(dir + Path.basic(name), output_document(head, body, css = css, hidden = hidden, structural = structural)) } } diff --git a/src/Pure/Thy/latex.scala b/src/Pure/Thy/latex.scala --- a/src/Pure/Thy/latex.scala +++ b/src/Pure/Thy/latex.scala @@ -1,177 +1,178 @@ /* Title: Pure/Thy/latex.scala Author: Makarius Support for LaTeX. */ package isabelle import java.io.{File => JFile} import scala.annotation.tailrec +import scala.util.matching.Regex object Latex { /** latex errors **/ def latex_errors(dir: Path, root_name: String): List[String] = { val root_log_path = dir + Path.explode(root_name).ext("log") if (root_log_path.is_file) { for { (msg, pos) <- filter_errors(dir, File.read(root_log_path)) } yield "Latex error" + Position.here(pos) + ":\n" + Library.prefix_lines(" ", msg) } else Nil } /* generated .tex file */ private val File_Pattern = """^%:%file=(.+)%:%$""".r private val Line_Pattern = """^*%:%(\d+)=(\d+)%:%$""".r def read_tex_file(tex_file: Path): Tex_File = { val positions = Line.logical_lines(File.read(tex_file)).reverse. takeWhile(_ != "\\endinput").reverse val source_file = positions match { case File_Pattern(file) :: _ => Some(file) case _ => None } val source_lines = if (source_file.isEmpty) Nil else positions.flatMap(line => line match { case Line_Pattern(Value.Int(line), Value.Int(source_line)) => Some(line -> source_line) case _ => None }) new Tex_File(tex_file, source_file, source_lines) } final class Tex_File private[Latex]( tex_file: Path, source_file: Option[String], source_lines: List[(Int, Int)]) { override def toString: String = tex_file.toString def source_position(l: Int): Option[Position.T] = source_file match { case None => None case Some(file) => val source_line = (0 /: source_lines.iterator.takeWhile({ case (m, _) => m <= l }))( { case (_, (_, n)) => n }) if (source_line == 0) None else Some(Position.Line_File(source_line, file)) } def position(line: Int): Position.T = source_position(line) getOrElse Position.Line_File(line, tex_file.implode) } /* latex log */ def filter_errors(dir: Path, root_log: String): List[(String, Position.T)] = { val seen_files = Synchronized(Map.empty[JFile, Tex_File]) def check_tex_file(path: Path): Option[Tex_File] = seen_files.change_result(seen => seen.get(path.file) match { case None => if (path.is_file) { val tex_file = read_tex_file(path) (Some(tex_file), seen + (path.file -> tex_file)) } else (None, seen) case some => (some, seen) }) def tex_file_position(path: Path, line: Int): Position.T = check_tex_file(path) match { case Some(tex_file) => tex_file.position(line) case None => Position.Line_File(line, path.implode) } object File_Line_Error { - val Pattern = """^(.*?\.\w\w\w):(\d+): (.*)$""".r + val Pattern: Regex = """^(.*?\.\w\w\w):(\d+): (.*)$""".r def unapply(line: String): Option[(Path, Int, String)] = line match { case Pattern(file, Value.Int(line), message) => val path = File.standard_path(file) if (Path.is_wellformed(path)) { val file = (dir + Path.explode(path)).canonical val msg = Library.perhaps_unprefix("LaTeX Error: ", message) if (file.is_file) Some((file, line, msg)) else None } else None case _ => None } } val Line_Error = """^l\.\d+ (.*)$""".r val More_Error = List( "", "

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