| author | wenzelm | 
| Sat, 20 May 2023 16:12:37 +0200 | |
| changeset 78084 | f0aca0506531 | 
| parent 77719 | cbfbf48b0281 | 
| child 78178 | a177f71dc79f | 
| permissions | -rw-r--r-- | 
| 50686 | 1 | /* Title: Pure/Tools/build.scala | 
| 48276 | 2 | Author: Makarius | 
| 57923 | 3 | Options: :folding=explicit: | 
| 48276 | 4 | |
| 77553 | 5 | Command-line tools to build and manage Isabelle sessions. | 
| 48276 | 6 | */ | 
| 7 | ||
| 8 | package isabelle | |
| 9 | ||
| 77553 | 10 | import scala.collection.mutable | 
| 11 | import scala.util.matching.Regex | |
| 12 | ||
| 48276 | 13 | |
| 75393 | 14 | object Build {
 | 
| 77553 | 15 | /** "isabelle build" **/ | 
| 16 | ||
| 77330 
47eb96592aa2
support alternative build engines, via system option "build_engine";
 wenzelm parents: 
77326diff
changeset | 17 | /* results */ | 
| 48424 | 18 | |
| 77311 | 19 |   object Results {
 | 
| 20 | def apply(context: Build_Process.Context, results: Map[String, Process_Result]): Results = | |
| 77453 | 21 | new Results(context.store, context.build_deps, results) | 
| 77311 | 22 | } | 
| 23 | ||
| 24 | class Results private( | |
| 76206 
769a7cd5a16a
clarified signature: re-use store/cache from build results;
 wenzelm parents: 
76202diff
changeset | 25 | val store: Sessions.Store, | 
| 76209 | 26 | val deps: Sessions.Deps, | 
| 77253 
792dad9cb04f
clarified data structure: absorb Option[Process_Result] into Process_Result, e.g. to simplify database storage;
 wenzelm parents: 
77252diff
changeset | 27 | results: Map[String, Process_Result] | 
| 76198 
fb4215da4919
clarified presentation_sessions: work with partial results;
 wenzelm parents: 
76197diff
changeset | 28 |   ) {
 | 
| 76206 
769a7cd5a16a
clarified signature: re-use store/cache from build results;
 wenzelm parents: 
76202diff
changeset | 29 | def cache: Term.Cache = store.cache | 
| 
769a7cd5a16a
clarified signature: re-use store/cache from build results;
 wenzelm parents: 
76202diff
changeset | 30 | |
| 77311 | 31 | def sessions_ok: List[String] = | 
| 32 |       (for {
 | |
| 33 | name <- deps.sessions_structure.build_topological_order.iterator | |
| 34 | result <- results.get(name) if result.ok | |
| 35 | } yield name).toList | |
| 36 | ||
| 77250 
22016642d6af
clarified data structure: use static info from deps, not dynamic results;
 wenzelm parents: 
77246diff
changeset | 37 | def info(name: String): Sessions.Info = deps.sessions_structure(name) | 
| 62403 | 38 | def sessions: Set[String] = results.keySet | 
| 77253 
792dad9cb04f
clarified data structure: absorb Option[Process_Result] into Process_Result, e.g. to simplify database storage;
 wenzelm parents: 
77252diff
changeset | 39 | def cancelled(name: String): Boolean = !results(name).defined | 
| 
792dad9cb04f
clarified data structure: absorb Option[Process_Result] into Process_Result, e.g. to simplify database storage;
 wenzelm parents: 
77252diff
changeset | 40 | def apply(name: String): Process_Result = results(name).strict | 
| 
792dad9cb04f
clarified data structure: absorb Option[Process_Result] into Process_Result, e.g. to simplify database storage;
 wenzelm parents: 
77252diff
changeset | 41 | val rc: Int = results.valuesIterator.map(_.strict.rc).foldLeft(Process_Result.RC.ok)(_ max _) | 
| 74306 | 42 | def ok: Boolean = rc == Process_Result.RC.ok | 
| 62406 | 43 | |
| 76202 | 44 | def unfinished: List[String] = sessions.iterator.filterNot(apply(_).ok).toList.sorted | 
| 76199 | 45 | |
| 62406 | 46 | override def toString: String = rc.toString | 
| 62403 | 47 | } | 
| 48 | ||
| 77330 
47eb96592aa2
support alternative build engines, via system option "build_engine";
 wenzelm parents: 
77326diff
changeset | 49 | |
| 
47eb96592aa2
support alternative build engines, via system option "build_engine";
 wenzelm parents: 
77326diff
changeset | 50 | /* engine */ | 
| 
47eb96592aa2
support alternative build engines, via system option "build_engine";
 wenzelm parents: 
77326diff
changeset | 51 | |
| 77411 
149cc77f7348
clafified signature: simplify object-oriented reuse;
 wenzelm parents: 
77384diff
changeset | 52 |   class Engine(val name: String) extends Isabelle_System.Service {
 | 
| 77330 
47eb96592aa2
support alternative build engines, via system option "build_engine";
 wenzelm parents: 
77326diff
changeset | 53 | override def toString: String = name | 
| 77505 | 54 | def init(build_context: Build_Process.Context, build_progress: Progress): Build_Process = | 
| 55 | new Build_Process(build_context, build_progress) | |
| 77330 
47eb96592aa2
support alternative build engines, via system option "build_engine";
 wenzelm parents: 
77326diff
changeset | 56 | } | 
| 
47eb96592aa2
support alternative build engines, via system option "build_engine";
 wenzelm parents: 
77326diff
changeset | 57 | |
| 77411 
149cc77f7348
clafified signature: simplify object-oriented reuse;
 wenzelm parents: 
77384diff
changeset | 58 |   class Default_Engine extends Engine("") { override def toString: String = "<default>" }
 | 
| 77330 
47eb96592aa2
support alternative build engines, via system option "build_engine";
 wenzelm parents: 
77326diff
changeset | 59 | |
| 
47eb96592aa2
support alternative build engines, via system option "build_engine";
 wenzelm parents: 
77326diff
changeset | 60 | lazy val engines: List[Engine] = | 
| 
47eb96592aa2
support alternative build engines, via system option "build_engine";
 wenzelm parents: 
77326diff
changeset | 61 | Isabelle_System.make_services(classOf[Engine]) | 
| 
47eb96592aa2
support alternative build engines, via system option "build_engine";
 wenzelm parents: 
77326diff
changeset | 62 | |
| 
47eb96592aa2
support alternative build engines, via system option "build_engine";
 wenzelm parents: 
77326diff
changeset | 63 | def get_engine(name: String): Engine = | 
| 
47eb96592aa2
support alternative build engines, via system option "build_engine";
 wenzelm parents: 
77326diff
changeset | 64 |     engines.find(_.name == name).getOrElse(error("Bad build engine " + quote(name)))
 | 
| 
47eb96592aa2
support alternative build engines, via system option "build_engine";
 wenzelm parents: 
77326diff
changeset | 65 | |
| 
47eb96592aa2
support alternative build engines, via system option "build_engine";
 wenzelm parents: 
77326diff
changeset | 66 | |
| 77558 | 67 | /* options */ | 
| 68 | ||
| 69 | def hostname(options: Options): String = | |
| 70 |     Isabelle_System.hostname(options.string("build_hostname"))
 | |
| 77330 
47eb96592aa2
support alternative build engines, via system option "build_engine";
 wenzelm parents: 
77326diff
changeset | 71 | |
| 77559 
4ad322ee6025
clarified signature: support all arguments of Sessions.store();
 wenzelm parents: 
77558diff
changeset | 72 |   def build_init(options: Options, cache: Term.Cache = Term.Cache.make()): Sessions.Store = {
 | 
| 77674 | 73 | val build_options = options + "completion_limit=0" + "editor_tracing_messages=0" | 
| 77559 
4ad322ee6025
clarified signature: support all arguments of Sessions.store();
 wenzelm parents: 
77558diff
changeset | 74 | val store = Sessions.store(build_options, cache = cache) | 
| 77557 | 75 | |
| 76 | Isabelle_Fonts.init() | |
| 77 | ||
| 78 | store | |
| 79 | } | |
| 80 | ||
| 77558 | 81 | |
| 82 | /* build */ | |
| 83 | ||
| 62641 | 84 | def build( | 
| 50404 
898cac1dad5e
avoid startup within GUI thread -- it is only required later for dialog;
 wenzelm parents: 
50367diff
changeset | 85 | options: Options, | 
| 71981 | 86 | selection: Sessions.Selection = Sessions.Selection.empty, | 
| 75942 
603852abed8f
clarified names: Browser_Info.Config vs. Browser_Info.Context;
 wenzelm parents: 
75941diff
changeset | 87 | browser_info: Browser_Info.Config = Browser_Info.Config.none, | 
| 71726 
a5fda30edae2
clarified signature: more uniform treatment of stopped/interrupted state;
 wenzelm parents: 
71718diff
changeset | 88 | progress: Progress = new Progress, | 
| 65832 
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
 wenzelm parents: 
65831diff
changeset | 89 | check_unknown_files: Boolean = false, | 
| 48511 
37999ee01156
remove old output heaps, to ensure that result is valid wrt. check_stamps;
 wenzelm parents: 
48509diff
changeset | 90 | build_heap: Boolean = false, | 
| 48595 | 91 | clean_build: Boolean = false, | 
| 56890 | 92 | dirs: List[Path] = Nil, | 
| 93 | select_dirs: List[Path] = Nil, | |
| 66968 
9991671c98aa
allow to augment session context via explicit session infos;
 wenzelm parents: 
66962diff
changeset | 94 | infos: List[Sessions.Info] = Nil, | 
| 64265 | 95 | numa_shuffling: Boolean = false, | 
| 48509 | 96 | max_jobs: Int = 1, | 
| 48903 | 97 | list_files: Boolean = false, | 
| 59891 
9ce697050455
added isabelle build option -k, for fast off-line checking of theory sources;
 wenzelm parents: 
59811diff
changeset | 98 | check_keywords: Set[String] = Set.empty, | 
| 66841 | 99 | fresh_build: Boolean = false, | 
| 48509 | 100 | no_build: Boolean = false, | 
| 66745 | 101 | soft_build: Boolean = false, | 
| 73805 
b73777a0c076
allow build session setup, e.g. for protocol handlers;
 wenzelm parents: 
73804diff
changeset | 102 | export_files: Boolean = false, | 
| 76927 
da13da82f6f9
treat update_options as part of Sessions.Info meta_digest, for proper re-build of updated sessions;
 wenzelm parents: 
76919diff
changeset | 103 | augment_options: String => List[Options.Spec] = _ => Nil, | 
| 77559 
4ad322ee6025
clarified signature: support all arguments of Sessions.store();
 wenzelm parents: 
77558diff
changeset | 104 | session_setup: (String, Session) => Unit = (_, _) => (), | 
| 
4ad322ee6025
clarified signature: support all arguments of Sessions.store();
 wenzelm parents: 
77558diff
changeset | 105 | cache: Term.Cache = Term.Cache.make() | 
| 75393 | 106 |   ): Results = {
 | 
| 77559 
4ad322ee6025
clarified signature: support all arguments of Sessions.store();
 wenzelm parents: 
77558diff
changeset | 107 | val store = build_init(options, cache = cache) | 
| 77557 | 108 | val build_options = store.options | 
| 69369 
6ecc85955e04
prefer "Isabelle DejaVu Sans", even for headless batch-build (session_graph.pdf);
 wenzelm parents: 
68957diff
changeset | 109 | |
| 66745 | 110 | |
| 111 | /* session selection and dependencies */ | |
| 65422 | 112 | |
| 66961 | 113 | val full_sessions = | 
| 76927 
da13da82f6f9
treat update_options as part of Sessions.Info meta_digest, for proper re-build of updated sessions;
 wenzelm parents: 
76919diff
changeset | 114 | Sessions.load_structure(build_options, dirs = dirs, select_dirs = select_dirs, infos = infos, | 
| 
da13da82f6f9
treat update_options as part of Sessions.Info meta_digest, for proper re-build of updated sessions;
 wenzelm parents: 
76919diff
changeset | 115 | augment_options = augment_options) | 
| 73012 
238ddf525da4
clarified HTML presentation, e.g. avoid bulky jobs like HOL or HOL-Analysis in applications;
 wenzelm parents: 
72993diff
changeset | 116 | val full_sessions_selection = full_sessions.imports_selection(selection) | 
| 
238ddf525da4
clarified HTML presentation, e.g. avoid bulky jobs like HOL or HOL-Analysis in applications;
 wenzelm parents: 
72993diff
changeset | 117 | |
| 75948 | 118 |     val build_deps = {
 | 
| 66745 | 119 | val deps0 = | 
| 77521 
5642de4d225d
clarified signature: manage "verbose" flag via "progress";
 wenzelm parents: 
77505diff
changeset | 120 | Sessions.deps(full_sessions.selection(selection), progress = progress, inlined_files = true, | 
| 66962 
e1bde71bace6
clarified signature: global_theories is always required;
 wenzelm parents: 
66961diff
changeset | 121 | list_files = list_files, check_keywords = check_keywords).check_errors | 
| 66745 | 122 | |
| 66841 | 123 |       if (soft_build && !fresh_build) {
 | 
| 66745 | 124 | val outdated = | 
| 68204 | 125 | deps0.sessions_structure.build_topological_order.flatMap(name => | 
| 72634 
5cea0993ee4f
clarified access to single database server vs. collection of database files;
 wenzelm parents: 
72624diff
changeset | 126 |             store.try_open_database(name) match {
 | 
| 68214 | 127 | case Some(db) => | 
| 68216 | 128 |                 using(db)(store.read_build(_, name)) match {
 | 
| 77675 
9e5f8f6e58a0
more thorough treatment of build prefs, guarded by system option "build_through": avoid accidental rebuild of HOL etc.;
 wenzelm parents: 
77674diff
changeset | 129 | case Some(build) if build.ok => | 
| 
9e5f8f6e58a0
more thorough treatment of build prefs, guarded by system option "build_through": avoid accidental rebuild of HOL etc.;
 wenzelm parents: 
77674diff
changeset | 130 | val session_options = deps0.sessions_structure(name).options | 
| 
9e5f8f6e58a0
more thorough treatment of build prefs, guarded by system option "build_through": avoid accidental rebuild of HOL etc.;
 wenzelm parents: 
77674diff
changeset | 131 | val session_sources = deps0.sources_shasum(name) | 
| 
9e5f8f6e58a0
more thorough treatment of build prefs, guarded by system option "build_through": avoid accidental rebuild of HOL etc.;
 wenzelm parents: 
77674diff
changeset | 132 | if (Sessions.eq_sources(session_options, build.sources, session_sources)) None | 
| 
9e5f8f6e58a0
more thorough treatment of build prefs, guarded by system option "build_through": avoid accidental rebuild of HOL etc.;
 wenzelm parents: 
77674diff
changeset | 133 | else Some(name) | 
| 66745 | 134 | case _ => Some(name) | 
| 135 | } | |
| 136 | case None => Some(name) | |
| 137 | }) | |
| 68204 | 138 | |
| 139 | Sessions.deps(full_sessions.selection(Sessions.Selection(sessions = outdated)), | |
| 70671 
cb1776c8e216
clarified signature: retain global session information, unaffected by later restriction;
 wenzelm parents: 
70509diff
changeset | 140 | progress = progress, inlined_files = true).check_errors | 
| 66745 | 141 | } | 
| 68204 | 142 | else deps0 | 
| 66745 | 143 | } | 
| 144 | ||
| 145 | ||
| 146 | /* check unknown files */ | |
| 48504 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 wenzelm parents: 
48494diff
changeset | 147 | |
| 65832 
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
 wenzelm parents: 
65831diff
changeset | 148 |     if (check_unknown_files) {
 | 
| 
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
 wenzelm parents: 
65831diff
changeset | 149 | val source_files = | 
| 
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
 wenzelm parents: 
65831diff
changeset | 150 |         (for {
 | 
| 75948 | 151 | (_, base) <- build_deps.session_bases.iterator | 
| 75741 | 152 | (path, _) <- base.session_sources.iterator | 
| 65832 
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
 wenzelm parents: 
65831diff
changeset | 153 | } yield path).toList | 
| 76670 
b04d45bebbc5
discontinued somewhat pointless dependency: avoid illusion of extra accuracy (see also 09fb749d1a1e and 0f750a6dc754);
 wenzelm parents: 
76669diff
changeset | 154 |       Mercurial.check_files(source_files)._2 match {
 | 
| 
b04d45bebbc5
discontinued somewhat pointless dependency: avoid illusion of extra accuracy (see also 09fb749d1a1e and 0f750a6dc754);
 wenzelm parents: 
76669diff
changeset | 155 | case Nil => | 
| 
b04d45bebbc5
discontinued somewhat pointless dependency: avoid illusion of extra accuracy (see also 09fb749d1a1e and 0f750a6dc754);
 wenzelm parents: 
76669diff
changeset | 156 | case unknown_files => | 
| 
b04d45bebbc5
discontinued somewhat pointless dependency: avoid illusion of extra accuracy (see also 09fb749d1a1e and 0f750a6dc754);
 wenzelm parents: 
76669diff
changeset | 157 |           progress.echo_warning("Unknown files (not part of the underlying Mercurial repository):" +
 | 
| 76883 | 158 |             unknown_files.map(File.standard_path).sorted.mkString("\n  ", "\n  ", ""))
 | 
| 65832 
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
 wenzelm parents: 
65831diff
changeset | 159 | } | 
| 
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
 wenzelm parents: 
65831diff
changeset | 160 | } | 
| 
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
 wenzelm parents: 
65831diff
changeset | 161 | |
| 51220 | 162 | |
| 77257 | 163 | /* build process and results */ | 
| 51220 | 164 | |
| 77315 
f34559b24277
clarified signature: move all parameters into Build_Process.Context;
 wenzelm parents: 
77311diff
changeset | 165 | val build_context = | 
| 
f34559b24277
clarified signature: move all parameters into Build_Process.Context;
 wenzelm parents: 
77311diff
changeset | 166 | Build_Process.Context(store, build_deps, progress = progress, | 
| 77558 | 167 | hostname = hostname(build_options), build_heap = build_heap, | 
| 168 | numa_shuffling = numa_shuffling, max_jobs = max_jobs, fresh_build = fresh_build, | |
| 77579 
69d3547206db
clarified signature: prefer Build_Process.Context for parameters;
 wenzelm parents: 
77563diff
changeset | 169 | no_build = no_build, session_setup = session_setup, master = true) | 
| 77254 
8d34f53871b4
clarified signature: make dynamic Queue from static Context;
 wenzelm parents: 
77253diff
changeset | 170 | |
| 77535 | 171 | store.prepare_output() | 
| 77536 
7c7f1473e51a
clarified database content and prepare/init stages;
 wenzelm parents: 
77535diff
changeset | 172 | build_context.prepare_database() | 
| 48373 | 173 | |
| 48595 | 174 |     if (clean_build) {
 | 
| 73012 
238ddf525da4
clarified HTML presentation, e.g. avoid bulky jobs like HOL or HOL-Analysis in applications;
 wenzelm parents: 
72993diff
changeset | 175 |       for (name <- full_sessions.imports_descendants(full_sessions_selection)) {
 | 
| 77677 | 176 |         store.clean_output(name) match {
 | 
| 177 | case None => | |
| 178 |           case Some(true) => progress.echo("Cleaned " + name)
 | |
| 179 | case Some(false) => progress.echo(name + " FAILED to clean") | |
| 68220 
8fc4e3d1df86
clarified store.clean_output: cleanup user_output_dir even in system_mode;
 wenzelm parents: 
68219diff
changeset | 180 | } | 
| 48595 | 181 | } | 
| 182 | } | |
| 183 | ||
| 77311 | 184 | val results = | 
| 185 |       Isabelle_Thread.uninterruptible {
 | |
| 77330 
47eb96592aa2
support alternative build engines, via system option "build_engine";
 wenzelm parents: 
77326diff
changeset | 186 |         val engine = get_engine(build_options.string("build_engine"))
 | 
| 77505 | 187 |         using(engine.init(build_context, progress)) { build_process =>
 | 
| 77579 
69d3547206db
clarified signature: prefer Build_Process.Context for parameters;
 wenzelm parents: 
77563diff
changeset | 188 | val res = build_process.run() | 
| 77372 | 189 | Results(build_context, res) | 
| 190 | } | |
| 77311 | 191 | } | 
| 62641 | 192 | |
| 69811 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 wenzelm parents: 
69777diff
changeset | 193 |     if (export_files) {
 | 
| 76196 | 194 |       for (name <- full_sessions_selection.iterator if results(name).ok) {
 | 
| 69811 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 wenzelm parents: 
69777diff
changeset | 195 | val info = results.info(name) | 
| 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 wenzelm parents: 
69777diff
changeset | 196 |         if (info.export_files.nonEmpty) {
 | 
| 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 wenzelm parents: 
69777diff
changeset | 197 |           progress.echo("Exporting " + info.name + " ...")
 | 
| 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 wenzelm parents: 
69777diff
changeset | 198 |           for ((dir, prune, pats) <- info.export_files) {
 | 
| 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 wenzelm parents: 
69777diff
changeset | 199 | Export.export_files(store, name, info.dir + dir, | 
| 77521 
5642de4d225d
clarified signature: manage "verbose" flag via "progress";
 wenzelm parents: 
77505diff
changeset | 200 | progress = if (progress.verbose) progress else new Progress, | 
| 69811 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 wenzelm parents: 
69777diff
changeset | 201 | export_prune = prune, | 
| 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 wenzelm parents: 
69777diff
changeset | 202 | export_patterns = pats) | 
| 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 wenzelm parents: 
69777diff
changeset | 203 | } | 
| 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 wenzelm parents: 
69777diff
changeset | 204 | } | 
| 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 wenzelm parents: 
69777diff
changeset | 205 | } | 
| 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 wenzelm parents: 
69777diff
changeset | 206 | } | 
| 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 wenzelm parents: 
69777diff
changeset | 207 | |
| 76901 | 208 | val presentation_sessions = | 
| 209 | results.sessions_ok.filter(name => browser_info.enabled(results.info(name))) | |
| 76230 | 210 |     if (presentation_sessions.nonEmpty && !progress.stopped) {
 | 
| 211 | Browser_Info.build(browser_info, results.store, results.deps, presentation_sessions, | |
| 77521 
5642de4d225d
clarified signature: manage "verbose" flag via "progress";
 wenzelm parents: 
77505diff
changeset | 212 | progress = progress) | 
| 76198 
fb4215da4919
clarified presentation_sessions: work with partial results;
 wenzelm parents: 
76197diff
changeset | 213 | } | 
| 
fb4215da4919
clarified presentation_sessions: work with partial results;
 wenzelm parents: 
76197diff
changeset | 214 | |
| 77521 
5642de4d225d
clarified signature: manage "verbose" flag via "progress";
 wenzelm parents: 
77505diff
changeset | 215 |     if (!results.ok && (progress.verbose || !no_build)) {
 | 
| 76199 | 216 |       progress.echo("Unfinished session(s): " + commas(results.unfinished))
 | 
| 62641 | 217 | } | 
| 218 | ||
| 219 | results | |
| 48341 | 220 | } | 
| 221 | ||
| 222 | ||
| 77555 | 223 | /* build logic image */ | 
| 224 | ||
| 225 | def build_logic(options: Options, logic: String, | |
| 226 | progress: Progress = new Progress, | |
| 227 | build_heap: Boolean = false, | |
| 228 | dirs: List[Path] = Nil, | |
| 229 | fresh: Boolean = false, | |
| 230 | strict: Boolean = false | |
| 231 |   ): Int = {
 | |
| 232 | val selection = Sessions.Selection.session(logic) | |
| 233 | val rc = | |
| 234 | if (!fresh && build(options, selection = selection, | |
| 235 | build_heap = build_heap, no_build = true, dirs = dirs).ok) Process_Result.RC.ok | |
| 236 |       else {
 | |
| 237 |         progress.echo("Build started for Isabelle/" + logic + " ...")
 | |
| 238 | build(options, selection = selection, progress = progress, | |
| 239 | build_heap = build_heap, fresh_build = fresh, dirs = dirs).rc | |
| 240 | } | |
| 241 |     if (strict && rc != Process_Result.RC.ok) error("Failed to build Isabelle/" + logic) else rc
 | |
| 242 | } | |
| 243 | ||
| 244 | ||
| 77553 | 245 | /* command-line wrapper */ | 
| 48341 | 246 | |
| 77553 | 247 |   val isabelle_tool1 = Isabelle_Tool("build", "build and manage Isabelle sessions",
 | 
| 75394 | 248 | Scala_Project.here, | 
| 249 |     { args =>
 | |
| 250 | var base_sessions: List[String] = Nil | |
| 251 | var select_dirs: List[Path] = Nil | |
| 252 | var numa_shuffling = false | |
| 75942 
603852abed8f
clarified names: Browser_Info.Config vs. Browser_Info.Context;
 wenzelm parents: 
75941diff
changeset | 253 | var browser_info = Browser_Info.Config.none | 
| 75394 | 254 | var requirements = false | 
| 255 | var soft_build = false | |
| 256 | var exclude_session_groups: List[String] = Nil | |
| 257 | var all_sessions = false | |
| 258 | var build_heap = false | |
| 259 | var clean_build = false | |
| 260 | var dirs: List[Path] = Nil | |
| 261 | var export_files = false | |
| 262 | var fresh_build = false | |
| 263 | var session_groups: List[String] = Nil | |
| 264 | var max_jobs = 1 | |
| 265 | var check_keywords: Set[String] = Set.empty | |
| 266 | var list_files = false | |
| 267 | var no_build = false | |
| 77628 | 268 | var options = Options.init(specs = Options.Spec.ISABELLE_BUILD_OPTIONS) | 
| 75394 | 269 | var verbose = false | 
| 270 | var exclude_sessions: List[String] = Nil | |
| 62590 | 271 | |
| 75394 | 272 |       val getopts = Getopts("""
 | 
| 62590 | 273 | Usage: isabelle build [OPTIONS] [SESSIONS ...] | 
| 274 | ||
| 275 | Options are: | |
| 66737 
2edc0c42c883
option -B for "isabelle build" and "isabelle imports";
 wenzelm parents: 
66736diff
changeset | 276 | -B NAME include session NAME and all descendants | 
| 62590 | 277 | -D DIR include session directory and select its sessions | 
| 64265 | 278 | -N cyclic shuffling of NUMA CPU nodes (performance tuning) | 
| 72648 | 279 |     -P DIR       enable HTML/PDF presentation in directory (":" for default)
 | 
| 71807 | 280 | -R refer to requirements of selected sessions | 
| 66745 | 281 | -S soft build: only observe changes of sources, not heap images | 
| 62590 | 282 | -X NAME exclude sessions from group NAME and all descendants | 
| 283 | -a select all sessions | |
| 284 | -b build heap images | |
| 285 | -c clean build | |
| 286 | -d DIR include session directory | |
| 69811 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 wenzelm parents: 
69777diff
changeset | 287 | -e export files from session specification into file-system | 
| 66841 | 288 | -f fresh build | 
| 62590 | 289 | -g NAME select session group NAME | 
| 290 | -j INT maximum number of parallel jobs (default 1) | |
| 291 | -k KEYWORD check theory sources for conflicts with proposed keywords | |
| 292 | -l list session source files | |
| 77554 
4465d9dff448
clarified terminology of "session build database", while "build database" is the one underlying Build_Process;
 wenzelm parents: 
77553diff
changeset | 293 | -n no build -- take existing session build databases | 
| 62590 | 294 | -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) | 
| 295 | -v verbose | |
| 296 | -x NAME exclude session NAME and all descendants | |
| 297 | ||
| 77647 | 298 | Build and manage Isabelle sessions: ML heaps, session databases, documents. | 
| 62596 | 299 | |
| 77647 | 300 | Notable system options: see "isabelle options -l -t build" | 
| 301 | ||
| 302 | Notable system settings: | |
| 303 | """ + Library.indent_lines(4, Build_Log.Settings.show()) + "\n", | |
| 75394 | 304 | "B:" -> (arg => base_sessions = base_sessions ::: List(arg)), | 
| 305 | "D:" -> (arg => select_dirs = select_dirs ::: List(Path.explode(arg))), | |
| 306 | "N" -> (_ => numa_shuffling = true), | |
| 75942 
603852abed8f
clarified names: Browser_Info.Config vs. Browser_Info.Context;
 wenzelm parents: 
75941diff
changeset | 307 | "P:" -> (arg => browser_info = Browser_Info.Config.make(arg)), | 
| 75394 | 308 | "R" -> (_ => requirements = true), | 
| 309 | "S" -> (_ => soft_build = true), | |
| 310 | "X:" -> (arg => exclude_session_groups = exclude_session_groups ::: List(arg)), | |
| 311 | "a" -> (_ => all_sessions = true), | |
| 312 | "b" -> (_ => build_heap = true), | |
| 313 | "c" -> (_ => clean_build = true), | |
| 314 | "d:" -> (arg => dirs = dirs ::: List(Path.explode(arg))), | |
| 315 | "e" -> (_ => export_files = true), | |
| 316 | "f" -> (_ => fresh_build = true), | |
| 317 | "g:" -> (arg => session_groups = session_groups ::: List(arg)), | |
| 318 | "j:" -> (arg => max_jobs = Value.Int.parse(arg)), | |
| 319 | "k:" -> (arg => check_keywords = check_keywords + arg), | |
| 320 | "l" -> (_ => list_files = true), | |
| 321 | "n" -> (_ => no_build = true), | |
| 322 | "o:" -> (arg => options = options + arg), | |
| 323 | "v" -> (_ => verbose = true), | |
| 324 | "x:" -> (arg => exclude_sessions = exclude_sessions ::: List(arg))) | |
| 62590 | 325 | |
| 75394 | 326 | val sessions = getopts(args) | 
| 62590 | 327 | |
| 75394 | 328 | val progress = new Console_Progress(verbose = verbose) | 
| 62590 | 329 | |
| 75394 | 330 | val start_date = Date.now() | 
| 64140 | 331 | |
| 77521 
5642de4d225d
clarified signature: manage "verbose" flag via "progress";
 wenzelm parents: 
77505diff
changeset | 332 | progress.echo( | 
| 
5642de4d225d
clarified signature: manage "verbose" flag via "progress";
 wenzelm parents: 
77505diff
changeset | 333 | "Started at " + Build_Log.print_date(start_date) + | 
| 77719 | 334 |           " (" + Isabelle_System.ml_identifier() + " on " + hostname(options) +")",
 | 
| 77521 
5642de4d225d
clarified signature: manage "verbose" flag via "progress";
 wenzelm parents: 
77505diff
changeset | 335 | verbose = true) | 
| 
5642de4d225d
clarified signature: manage "verbose" flag via "progress";
 wenzelm parents: 
77505diff
changeset | 336 | progress.echo(Build_Log.Settings.show() + "\n", verbose = true) | 
| 62590 | 337 | |
| 75394 | 338 | val results = | 
| 339 |         progress.interrupt_handler {
 | |
| 340 | build(options, | |
| 341 | selection = Sessions.Selection( | |
| 342 | requirements = requirements, | |
| 343 | all_sessions = all_sessions, | |
| 344 | base_sessions = base_sessions, | |
| 345 | exclude_session_groups = exclude_session_groups, | |
| 346 | exclude_sessions = exclude_sessions, | |
| 347 | session_groups = session_groups, | |
| 348 | sessions = sessions), | |
| 75942 
603852abed8f
clarified names: Browser_Info.Config vs. Browser_Info.Context;
 wenzelm parents: 
75941diff
changeset | 349 | browser_info = browser_info, | 
| 75394 | 350 | progress = progress, | 
| 351 | check_unknown_files = Mercurial.is_repository(Path.ISABELLE_HOME), | |
| 352 | build_heap = build_heap, | |
| 353 | clean_build = clean_build, | |
| 354 | dirs = dirs, | |
| 355 | select_dirs = select_dirs, | |
| 77477 | 356 | numa_shuffling = Host.numa_check(progress, numa_shuffling), | 
| 75394 | 357 | max_jobs = max_jobs, | 
| 358 | list_files = list_files, | |
| 359 | check_keywords = check_keywords, | |
| 360 | fresh_build = fresh_build, | |
| 361 | no_build = no_build, | |
| 362 | soft_build = soft_build, | |
| 363 | export_files = export_files) | |
| 364 | } | |
| 77545 | 365 | val stop_date = Date.now() | 
| 366 | val elapsed_time = stop_date.time - start_date.time | |
| 75394 | 367 | |
| 77545 | 368 |       progress.echo("\nFinished at " + Build_Log.print_date(stop_date), verbose = true)
 | 
| 62590 | 369 | |
| 75394 | 370 | val total_timing = | 
| 371 | results.sessions.iterator.map(a => results(a).timing).foldLeft(Timing.zero)(_ + _). | |
| 372 | copy(elapsed = elapsed_time) | |
| 373 | progress.echo(total_timing.message_resources) | |
| 62590 | 374 | |
| 75394 | 375 | sys.exit(results.rc) | 
| 376 | }) | |
| 68305 | 377 | |
| 378 | ||
| 77553 | 379 | |
| 77557 | 380 | /** "isabelle build_worker" **/ | 
| 381 | ||
| 382 | /* build_worker */ | |
| 383 | ||
| 384 | def build_worker( | |
| 385 | options: Options, | |
| 386 | build_uuid: String, | |
| 387 | progress: Progress = new Progress, | |
| 388 | dirs: List[Path] = Nil, | |
| 389 | infos: List[Sessions.Info] = Nil, | |
| 390 | numa_shuffling: Boolean = false, | |
| 391 | max_jobs: Int = 1, | |
| 77559 
4ad322ee6025
clarified signature: support all arguments of Sessions.store();
 wenzelm parents: 
77558diff
changeset | 392 | session_setup: (String, Session) => Unit = (_, _) => (), | 
| 
4ad322ee6025
clarified signature: support all arguments of Sessions.store();
 wenzelm parents: 
77558diff
changeset | 393 | cache: Term.Cache = Term.Cache.make() | 
| 77557 | 394 |   ): Results = {
 | 
| 77559 
4ad322ee6025
clarified signature: support all arguments of Sessions.store();
 wenzelm parents: 
77558diff
changeset | 395 | val store = build_init(options, cache = cache) | 
| 77557 | 396 | val build_options = store.options | 
| 397 | ||
| 398 |     progress.echo("build worker for " + build_uuid)
 | |
| 399 |     progress.echo_warning("FIXME")
 | |
| 400 | ??? | |
| 401 | } | |
| 402 | ||
| 403 | ||
| 404 | /* command-line wrapper */ | |
| 405 | ||
| 406 |   val isabelle_tool2 = Isabelle_Tool("build_worker", "external worker for session build process",
 | |
| 407 | Scala_Project.here, | |
| 408 |     { args =>
 | |
| 409 | var numa_shuffling = false | |
| 410 | var dirs: List[Path] = Nil | |
| 411 | var max_jobs = 1 | |
| 77628 | 412 | var options = Options.init(specs = Options.Spec.ISABELLE_BUILD_OPTIONS) | 
| 77557 | 413 | var verbose = false | 
| 414 | var build_uuid = "" | |
| 415 | ||
| 416 |       val getopts = Getopts("""
 | |
| 417 | Usage: isabelle build_worker [OPTIONS] ...] | |
| 418 | ||
| 419 | Options are: | |
| 420 | -N cyclic shuffling of NUMA CPU nodes (performance tuning) | |
| 421 | -U UUID Universally Unique Identifier of the build process | |
| 422 | -d DIR include session directory | |
| 423 | -j INT maximum number of parallel jobs (default 1) | |
| 424 | -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) | |
| 425 | -v verbose | |
| 426 | ||
| 427 | Run as external worker for session build process, as identified via | |
| 428 | option -U UUID. | |
| 77647 | 429 | """, | 
| 77557 | 430 | "N" -> (_ => numa_shuffling = true), | 
| 431 | "U:" -> (arg => build_uuid = arg), | |
| 432 | "d:" -> (arg => dirs = dirs ::: List(Path.explode(arg))), | |
| 433 | "j:" -> (arg => max_jobs = Value.Int.parse(arg)), | |
| 434 | "o:" -> (arg => options = options + arg), | |
| 435 | "v" -> (_ => verbose = true)) | |
| 436 | ||
| 437 | val more_args = getopts(args) | |
| 438 | if (more_args.nonEmpty) getopts.usage() | |
| 439 | ||
| 440 |       if (build_uuid.isEmpty) error("Missing UUID for build process (option -U)")
 | |
| 441 | ||
| 442 | val progress = new Console_Progress(verbose = verbose) | |
| 443 | ||
| 444 | val results = | |
| 445 |         progress.interrupt_handler {
 | |
| 446 | build_worker(options, build_uuid, | |
| 447 | progress = progress, | |
| 448 | dirs = dirs, | |
| 449 | numa_shuffling = Host.numa_check(progress, numa_shuffling), | |
| 450 | max_jobs = max_jobs) | |
| 451 | } | |
| 452 | ||
| 453 | sys.exit(results.rc) | |
| 454 | }) | |
| 455 | ||
| 456 | ||
| 457 | ||
| 77563 | 458 | /** "isabelle build_log" **/ | 
| 77553 | 459 | |
| 460 | /* theory markup/messages from session database */ | |
| 461 | ||
| 462 | def read_theory( | |
| 463 | theory_context: Export.Theory_Context, | |
| 464 | unicode_symbols: Boolean = false | |
| 465 |   ): Option[Document.Snapshot] = {
 | |
| 466 | def decode_bytes(bytes: Bytes): String = | |
| 467 | Symbol.output(unicode_symbols, UTF8.decode_permissive(bytes)) | |
| 468 | ||
| 469 | def read(name: String): Export.Entry = theory_context(name, permissive = true) | |
| 470 | ||
| 471 | def read_xml(name: String): XML.Body = | |
| 472 | YXML.parse_body(decode_bytes(read(name).bytes), cache = theory_context.cache) | |
| 473 | ||
| 474 | def read_source_file(name: String): Sessions.Source_File = | |
| 475 | theory_context.session_context.source_file(name) | |
| 476 | ||
| 477 |     for {
 | |
| 478 | id <- theory_context.document_id() | |
| 479 | (thy_file, blobs_files) <- theory_context.files(permissive = true) | |
| 480 | } | |
| 481 |     yield {
 | |
| 482 | val master_dir = | |
| 483 | Path.explode(Url.strip_base_name(thy_file).getOrElse( | |
| 484 |           error("Cannot determine theory master directory: " + quote(thy_file))))
 | |
| 485 | ||
| 486 | val blobs = | |
| 487 |         blobs_files.map { name =>
 | |
| 488 | val path = Path.explode(name) | |
| 489 | val src_path = File.relative_path(master_dir, path).getOrElse(path) | |
| 490 | ||
| 491 | val file = read_source_file(name) | |
| 492 | val bytes = file.bytes | |
| 493 | val text = decode_bytes(bytes) | |
| 494 | val chunk = Symbol.Text_Chunk(text) | |
| 495 | ||
| 496 | Command.Blob(Document.Node.Name(name), src_path, Some((file.digest, chunk))) -> | |
| 497 | Document.Blobs.Item(bytes, text, chunk, changed = false) | |
| 498 | } | |
| 499 | ||
| 500 | val thy_source = decode_bytes(read_source_file(thy_file).bytes) | |
| 501 | val thy_xml = read_xml(Export.MARKUP) | |
| 502 | val blobs_xml = | |
| 503 | for (i <- (1 to blobs.length).toList) | |
| 504 | yield read_xml(Export.MARKUP + i) | |
| 505 | ||
| 506 | val markups_index = Command.Markup_Index.make(blobs.map(_._1)) | |
| 507 | val markups = | |
| 508 | Command.Markups.make( | |
| 509 | for ((index, xml) <- markups_index.zip(thy_xml :: blobs_xml)) | |
| 510 | yield index -> Markup_Tree.from_XML(xml)) | |
| 511 | ||
| 512 | val results = | |
| 513 | Command.Results.make( | |
| 514 | for (elem @ XML.Elem(Markup(_, Markup.Serial(i)), _) <- read_xml(Export.MESSAGES)) | |
| 515 | yield i -> elem) | |
| 516 | ||
| 517 | val command = | |
| 518 | Command.unparsed(thy_source, theory = true, id = id, | |
| 519 | node_name = Document.Node.Name(thy_file, theory = theory_context.theory), | |
| 520 | blobs_info = Command.Blobs_Info.make(blobs), | |
| 521 | markups = markups, results = results) | |
| 522 | ||
| 523 | val doc_blobs = Document.Blobs.make(blobs) | |
| 524 | ||
| 525 | Document.State.init.snippet(command, doc_blobs) | |
| 526 | } | |
| 527 | } | |
| 528 | ||
| 529 | ||
| 530 | /* print messages */ | |
| 531 | ||
| 532 | def print_log( | |
| 533 | options: Options, | |
| 534 | sessions: List[String], | |
| 535 | theories: List[String] = Nil, | |
| 536 | message_head: List[Regex] = Nil, | |
| 537 | message_body: List[Regex] = Nil, | |
| 538 | progress: Progress = new Progress, | |
| 539 | margin: Double = Pretty.default_margin, | |
| 540 | breakgain: Double = Pretty.default_breakgain, | |
| 541 | metric: Pretty.Metric = Symbol.Metric, | |
| 542 | unicode_symbols: Boolean = false | |
| 543 |   ): Unit = {
 | |
| 544 | val store = Sessions.store(options) | |
| 545 | val session = new Session(options, Resources.bootstrap) | |
| 546 | ||
| 547 | def check(filter: List[Regex], make_string: => String): Boolean = | |
| 548 |       filter.isEmpty || {
 | |
| 549 | val s = Output.clean_yxml(make_string) | |
| 550 | filter.forall(r => r.findFirstIn(Output.clean_yxml(s)).nonEmpty) | |
| 551 | } | |
| 552 | ||
| 553 |     def print(session_name: String): Unit = {
 | |
| 554 |       using(Export.open_session_context0(store, session_name)) { session_context =>
 | |
| 555 | val result = | |
| 556 |           for {
 | |
| 557 | db <- session_context.session_db() | |
| 558 | theories = store.read_theories(db, session_name) | |
| 559 | errors = store.read_errors(db, session_name) | |
| 560 | info <- store.read_build(db, session_name) | |
| 561 | } yield (theories, errors, info.return_code) | |
| 562 |         result match {
 | |
| 563 | case None => store.error_database(session_name) | |
| 564 | case Some((used_theories, errors, rc)) => | |
| 565 |             theories.filterNot(used_theories.toSet) match {
 | |
| 566 | case Nil => | |
| 567 |               case bad => error("Unknown theories " + commas_quote(bad))
 | |
| 568 | } | |
| 569 | val print_theories = | |
| 570 | if (theories.isEmpty) used_theories else used_theories.filter(theories.toSet) | |
| 571 | ||
| 572 |             for (thy <- print_theories) {
 | |
| 573 | val thy_heading = "\nTheory " + quote(thy) + " (in " + session_name + ")" + ":" | |
| 574 | ||
| 575 |               Build_Job.read_theory(session_context.theory(thy), unicode_symbols = unicode_symbols) match {
 | |
| 576 | case None => progress.echo(thy_heading + " MISSING") | |
| 577 | case Some(snapshot) => | |
| 578 | val rendering = new Rendering(snapshot, options, session) | |
| 579 | val messages = | |
| 580 | rendering.text_messages(Text.Range.full) | |
| 581 | .filter(message => progress.verbose || Protocol.is_exported(message.info)) | |
| 582 |                   if (messages.nonEmpty) {
 | |
| 583 | val line_document = Line.Document(snapshot.node.source) | |
| 584 | val buffer = new mutable.ListBuffer[String] | |
| 585 |                     for (Text.Info(range, elem) <- messages) {
 | |
| 586 | val line = line_document.position(range.start).line1 | |
| 587 | val pos = Position.Line_File(line, snapshot.node_name.node) | |
| 588 | def message_text: String = | |
| 589 | Protocol.message_text(elem, heading = true, pos = pos, | |
| 590 | margin = margin, breakgain = breakgain, metric = metric) | |
| 591 | val ok = | |
| 592 | check(message_head, Protocol.message_heading(elem, pos)) && | |
| 593 | check(message_body, XML.content(Pretty.unformatted(List(elem)))) | |
| 594 | if (ok) buffer += message_text | |
| 595 | } | |
| 596 |                     if (buffer.nonEmpty) {
 | |
| 597 | progress.echo(thy_heading) | |
| 598 | buffer.foreach(progress.echo(_)) | |
| 599 | } | |
| 600 | } | |
| 601 | } | |
| 602 | } | |
| 603 | ||
| 604 |             if (errors.nonEmpty) {
 | |
| 605 | val msg = Symbol.output(unicode_symbols, cat_lines(errors)) | |
| 606 |               progress.echo("\nBuild errors:\n" + Output.error_message_text(msg))
 | |
| 607 | } | |
| 608 |             if (rc != Process_Result.RC.ok) {
 | |
| 609 |               progress.echo("\n" + Process_Result.RC.print_long(rc))
 | |
| 610 | } | |
| 611 | } | |
| 612 | } | |
| 613 | } | |
| 614 | ||
| 615 | val errors = new mutable.ListBuffer[String] | |
| 616 |     for (session_name <- sessions) {
 | |
| 617 |       Exn.interruptible_capture(print(session_name)) match {
 | |
| 618 | case Exn.Res(_) => | |
| 619 | case Exn.Exn(exn) => errors += Exn.message(exn) | |
| 620 | } | |
| 621 | } | |
| 622 | if (errors.nonEmpty) error(cat_lines(errors.toList)) | |
| 623 | } | |
| 624 | ||
| 625 | ||
| 626 | /* command-line wrapper */ | |
| 627 | ||
| 77563 | 628 |   val isabelle_tool3 = Isabelle_Tool("build_log", "print messages from session build database",
 | 
| 77553 | 629 | Scala_Project.here, | 
| 630 |     { args =>
 | |
| 631 | /* arguments */ | |
| 632 | ||
| 633 | var message_head = List.empty[Regex] | |
| 634 | var message_body = List.empty[Regex] | |
| 635 | var unicode_symbols = false | |
| 636 | var theories: List[String] = Nil | |
| 637 | var margin = Pretty.default_margin | |
| 638 | var options = Options.init() | |
| 639 | var verbose = false | |
| 640 | ||
| 641 |       val getopts = Getopts("""
 | |
| 77563 | 642 | Usage: isabelle build_log [OPTIONS] [SESSIONS ...] | 
| 77553 | 643 | |
| 644 | Options are: | |
| 645 | -H REGEX filter messages by matching against head | |
| 646 | -M REGEX filter messages by matching against body | |
| 647 | -T NAME restrict to given theories (multiple options possible) | |
| 648 | -U output Unicode symbols | |
| 649 | -m MARGIN margin for pretty printing (default: """ + margin + """) | |
| 650 | -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) | |
| 651 | -v print all messages, including information etc. | |
| 652 | ||
| 77554 
4465d9dff448
clarified terminology of "session build database", while "build database" is the one underlying Build_Process;
 wenzelm parents: 
77553diff
changeset | 653 | Print messages from the session build database of the given sessions, | 
| 
4465d9dff448
clarified terminology of "session build database", while "build database" is the one underlying Build_Process;
 wenzelm parents: 
77553diff
changeset | 654 | without any checks against current sources nor session structure: results | 
| 
4465d9dff448
clarified terminology of "session build database", while "build database" is the one underlying Build_Process;
 wenzelm parents: 
77553diff
changeset | 655 | from old sessions or failed builds can be printed as well. | 
| 77553 | 656 | |
| 657 | Multiple options -H and -M are conjunctive: all given patterns need to | |
| 658 | match. Patterns match any substring, but ^ or $ may be used to match the | |
| 659 | start or end explicitly. | |
| 660 | """, | |
| 661 | "H:" -> (arg => message_head = message_head ::: List(arg.r)), | |
| 662 | "M:" -> (arg => message_body = message_body ::: List(arg.r)), | |
| 663 | "T:" -> (arg => theories = theories ::: List(arg)), | |
| 664 | "U" -> (_ => unicode_symbols = true), | |
| 665 | "m:" -> (arg => margin = Value.Double.parse(arg)), | |
| 666 | "o:" -> (arg => options = options + arg), | |
| 667 | "v" -> (_ => verbose = true)) | |
| 668 | ||
| 669 | val sessions = getopts(args) | |
| 670 | ||
| 671 | val progress = new Console_Progress(verbose = verbose) | |
| 672 | ||
| 673 |       if (sessions.isEmpty) progress.echo_warning("No sessions to print")
 | |
| 674 |       else {
 | |
| 675 | print_log(options, sessions, theories = theories, message_head = message_head, | |
| 676 | message_body = message_body, margin = margin, progress = progress, | |
| 677 | unicode_symbols = unicode_symbols) | |
| 678 | } | |
| 679 | }) | |
| 48276 | 680 | } |