1 /* Title: Pure/Tools/build.scala |
|
2 Author: Makarius |
|
3 Options: :folding=explicit: |
|
4 |
|
5 Command-line tools to build and manage Isabelle sessions. |
|
6 */ |
|
7 |
|
8 package isabelle |
|
9 |
|
10 import scala.collection.mutable |
|
11 import scala.util.matching.Regex |
|
12 |
|
13 |
|
14 object Build { |
|
15 /** "isabelle build" **/ |
|
16 |
|
17 /* options */ |
|
18 |
|
19 def hostname(options: Options): String = |
|
20 Isabelle_System.hostname(options.string("build_hostname")) |
|
21 |
|
22 def engine_name(options: Options): String = options.string("build_engine") |
|
23 |
|
24 |
|
25 |
|
26 /* context */ |
|
27 |
|
28 sealed case class Context( |
|
29 store: Store, |
|
30 engine: Engine, |
|
31 build_deps: isabelle.Sessions.Deps, |
|
32 afp_root: Option[Path] = None, |
|
33 build_hosts: List[Build_Cluster.Host] = Nil, |
|
34 ml_platform: String = Isabelle_System.getenv("ML_PLATFORM"), |
|
35 hostname: String = Isabelle_System.hostname(), |
|
36 numa_shuffling: Boolean = false, |
|
37 build_heap: Boolean = false, |
|
38 max_jobs: Int = 1, |
|
39 fresh_build: Boolean = false, |
|
40 no_build: Boolean = false, |
|
41 session_setup: (String, Session) => Unit = (_, _) => (), |
|
42 build_uuid: String = UUID.random().toString, |
|
43 master: Boolean = false |
|
44 ) { |
|
45 override def toString: String = |
|
46 "Build.Context(build_uuid = " + quote(build_uuid) + if_proper(master, ", master = true") + ")" |
|
47 |
|
48 def build_options: Options = store.options |
|
49 |
|
50 def sessions_structure: isabelle.Sessions.Structure = build_deps.sessions_structure |
|
51 |
|
52 def worker_active: Boolean = max_jobs > 0 |
|
53 } |
|
54 |
|
55 |
|
56 /* results */ |
|
57 |
|
58 object Results { |
|
59 def apply( |
|
60 context: Context, |
|
61 results: Map[String, Process_Result] = Map.empty, |
|
62 other_rc: Int = Process_Result.RC.ok |
|
63 ): Results = { |
|
64 new Results(context.store, context.build_deps, results, other_rc) |
|
65 } |
|
66 } |
|
67 |
|
68 class Results private( |
|
69 val store: Store, |
|
70 val deps: Sessions.Deps, |
|
71 results: Map[String, Process_Result], |
|
72 other_rc: Int |
|
73 ) { |
|
74 def cache: Term.Cache = store.cache |
|
75 |
|
76 def sessions_ok: List[String] = |
|
77 (for { |
|
78 name <- deps.sessions_structure.build_topological_order.iterator |
|
79 result <- results.get(name) if result.ok |
|
80 } yield name).toList |
|
81 |
|
82 def info(name: String): Sessions.Info = deps.sessions_structure(name) |
|
83 def sessions: Set[String] = results.keySet |
|
84 def cancelled(name: String): Boolean = !results(name).defined |
|
85 def apply(name: String): Process_Result = results(name).strict |
|
86 |
|
87 val rc: Int = |
|
88 Process_Result.RC.merge(other_rc, |
|
89 Process_Result.RC.merge(results.valuesIterator.map(_.strict.rc))) |
|
90 def ok: Boolean = rc == Process_Result.RC.ok |
|
91 |
|
92 lazy val unfinished: List[String] = sessions.iterator.filterNot(apply(_).ok).toList.sorted |
|
93 |
|
94 override def toString: String = rc.toString |
|
95 } |
|
96 |
|
97 |
|
98 /* engine */ |
|
99 |
|
100 object Engine { |
|
101 lazy val services: List[Engine] = |
|
102 Isabelle_System.make_services(classOf[Engine]) |
|
103 |
|
104 def apply(name: String): Engine = |
|
105 services.find(_.name == name).getOrElse(error("Bad build engine " + quote(name))) |
|
106 } |
|
107 |
|
108 class Engine(val name: String) extends Isabelle_System.Service { |
|
109 override def toString: String = name |
|
110 |
|
111 def build_options(options: Options, build_hosts: List[Build_Cluster.Host] = Nil): Options = { |
|
112 val options1 = options + "completion_limit=0" + "editor_tracing_messages=0" |
|
113 if (build_hosts.isEmpty) options1 |
|
114 else options1 + "build_database_server" + "build_database" |
|
115 } |
|
116 |
|
117 final def build_store(options: Options, |
|
118 build_hosts: List[Build_Cluster.Host] = Nil, |
|
119 cache: Term.Cache = Term.Cache.make() |
|
120 ): Store = { |
|
121 val store = Store(build_options(options, build_hosts = build_hosts), cache = cache) |
|
122 Isabelle_System.make_directory(store.output_dir + Path.basic("log")) |
|
123 Isabelle_Fonts.init() |
|
124 store |
|
125 } |
|
126 |
|
127 def open_build_process( |
|
128 build_context: Context, |
|
129 build_progress: Progress, |
|
130 server: SSH.Server |
|
131 ): Build_Process = new Build_Process(build_context, build_progress, server) |
|
132 |
|
133 final def run_build_process( |
|
134 context: Context, |
|
135 progress: Progress, |
|
136 server: SSH.Server |
|
137 ): Results = { |
|
138 Isabelle_Thread.uninterruptible { |
|
139 using(open_build_process(context, progress, server))(_.run()) |
|
140 } |
|
141 } |
|
142 } |
|
143 |
|
144 class Default_Engine extends Engine("") { override def toString: String = "<default>" } |
|
145 |
|
146 |
|
147 /* build */ |
|
148 |
|
149 def build( |
|
150 options: Options, |
|
151 build_hosts: List[Build_Cluster.Host] = Nil, |
|
152 selection: Sessions.Selection = Sessions.Selection.empty, |
|
153 browser_info: Browser_Info.Config = Browser_Info.Config.none, |
|
154 progress: Progress = new Progress, |
|
155 check_unknown_files: Boolean = false, |
|
156 build_heap: Boolean = false, |
|
157 clean_build: Boolean = false, |
|
158 afp_root: Option[Path] = None, |
|
159 dirs: List[Path] = Nil, |
|
160 select_dirs: List[Path] = Nil, |
|
161 infos: List[Sessions.Info] = Nil, |
|
162 numa_shuffling: Boolean = false, |
|
163 max_jobs: Int = 1, |
|
164 list_files: Boolean = false, |
|
165 check_keywords: Set[String] = Set.empty, |
|
166 fresh_build: Boolean = false, |
|
167 no_build: Boolean = false, |
|
168 soft_build: Boolean = false, |
|
169 export_files: Boolean = false, |
|
170 augment_options: String => List[Options.Spec] = _ => Nil, |
|
171 session_setup: (String, Session) => Unit = (_, _) => (), |
|
172 cache: Term.Cache = Term.Cache.make() |
|
173 ): Results = { |
|
174 val build_engine = Engine(engine_name(options)) |
|
175 |
|
176 val store = build_engine.build_store(options, build_hosts = build_hosts, cache = cache) |
|
177 val build_options = store.options |
|
178 |
|
179 using(store.open_server()) { server => |
|
180 using_optional(store.maybe_open_database_server(server = server)) { database_server => |
|
181 |
|
182 |
|
183 /* session selection and dependencies */ |
|
184 |
|
185 val full_sessions = |
|
186 Sessions.load_structure(build_options, dirs = AFP.make_dirs(afp_root) ::: dirs, |
|
187 select_dirs = select_dirs, infos = infos, augment_options = augment_options) |
|
188 val full_sessions_selection = full_sessions.imports_selection(selection) |
|
189 |
|
190 val build_deps = { |
|
191 val deps0 = |
|
192 Sessions.deps(full_sessions.selection(selection), progress = progress, inlined_files = true, |
|
193 list_files = list_files, check_keywords = check_keywords).check_errors |
|
194 |
|
195 if (soft_build && !fresh_build) { |
|
196 val outdated = |
|
197 deps0.sessions_structure.build_topological_order.flatMap(name => |
|
198 store.try_open_database(name, server = server) match { |
|
199 case Some(db) => |
|
200 using(db)(store.read_build(_, name)) match { |
|
201 case Some(build) if build.ok => |
|
202 val session_options = deps0.sessions_structure(name).options |
|
203 val session_sources = deps0.sources_shasum(name) |
|
204 if (Sessions.eq_sources(session_options, build.sources, session_sources)) None |
|
205 else Some(name) |
|
206 case _ => Some(name) |
|
207 } |
|
208 case None => Some(name) |
|
209 }) |
|
210 |
|
211 Sessions.deps(full_sessions.selection(Sessions.Selection(sessions = outdated)), |
|
212 progress = progress, inlined_files = true).check_errors |
|
213 } |
|
214 else deps0 |
|
215 } |
|
216 |
|
217 |
|
218 /* check unknown files */ |
|
219 |
|
220 if (check_unknown_files) { |
|
221 val source_files = |
|
222 (for { |
|
223 (_, base) <- build_deps.session_bases.iterator |
|
224 (path, _) <- base.session_sources.iterator |
|
225 } yield path).toList |
|
226 Mercurial.check_files(source_files)._2 match { |
|
227 case Nil => |
|
228 case unknown_files => |
|
229 progress.echo_warning("Unknown files (not part of the underlying Mercurial repository):" + |
|
230 unknown_files.map(File.standard_path).sorted.mkString("\n ", "\n ", "")) |
|
231 } |
|
232 } |
|
233 |
|
234 |
|
235 /* build process and results */ |
|
236 |
|
237 val build_context = |
|
238 Context(store, build_engine, build_deps, afp_root = afp_root, build_hosts = build_hosts, |
|
239 hostname = hostname(build_options), build_heap = build_heap, |
|
240 numa_shuffling = numa_shuffling, max_jobs = max_jobs, fresh_build = fresh_build, |
|
241 no_build = no_build, session_setup = session_setup, master = true) |
|
242 |
|
243 if (clean_build) { |
|
244 for (name <- full_sessions.imports_descendants(full_sessions_selection)) { |
|
245 store.clean_output(database_server, name) match { |
|
246 case None => |
|
247 case Some(true) => progress.echo("Cleaned " + name) |
|
248 case Some(false) => progress.echo(name + " FAILED to clean") |
|
249 } |
|
250 } |
|
251 } |
|
252 |
|
253 val results = build_engine.run_build_process(build_context, progress, server) |
|
254 |
|
255 if (export_files) { |
|
256 for (name <- full_sessions_selection.iterator if results(name).ok) { |
|
257 val info = results.info(name) |
|
258 if (info.export_files.nonEmpty) { |
|
259 progress.echo("Exporting " + info.name + " ...") |
|
260 for ((dir, prune, pats) <- info.export_files) { |
|
261 Export.export_files(store, name, info.dir + dir, |
|
262 progress = if (progress.verbose) progress else new Progress, |
|
263 export_prune = prune, |
|
264 export_patterns = pats) |
|
265 } |
|
266 } |
|
267 } |
|
268 } |
|
269 |
|
270 val presentation_sessions = |
|
271 results.sessions_ok.filter(name => browser_info.enabled(results.info(name))) |
|
272 if (presentation_sessions.nonEmpty && !progress.stopped) { |
|
273 Browser_Info.build(browser_info, results.store, results.deps, presentation_sessions, |
|
274 progress = progress, server = server) |
|
275 } |
|
276 |
|
277 if (results.unfinished.nonEmpty && (progress.verbose || !no_build)) { |
|
278 progress.echo("Unfinished session(s): " + commas(results.unfinished)) |
|
279 } |
|
280 |
|
281 results |
|
282 } |
|
283 } |
|
284 } |
|
285 |
|
286 |
|
287 /* build logic image */ |
|
288 |
|
289 def build_logic(options: Options, logic: String, |
|
290 progress: Progress = new Progress, |
|
291 build_heap: Boolean = false, |
|
292 dirs: List[Path] = Nil, |
|
293 fresh: Boolean = false, |
|
294 strict: Boolean = false |
|
295 ): Int = { |
|
296 val selection = Sessions.Selection.session(logic) |
|
297 val rc = |
|
298 if (!fresh && build(options, selection = selection, |
|
299 build_heap = build_heap, no_build = true, dirs = dirs).ok) Process_Result.RC.ok |
|
300 else { |
|
301 progress.echo("Build started for Isabelle/" + logic + " ...") |
|
302 build(options, selection = selection, progress = progress, |
|
303 build_heap = build_heap, fresh_build = fresh, dirs = dirs).rc |
|
304 } |
|
305 if (strict && rc != Process_Result.RC.ok) error("Failed to build Isabelle/" + logic) else rc |
|
306 } |
|
307 |
|
308 |
|
309 /* command-line wrapper */ |
|
310 |
|
311 val isabelle_tool1 = Isabelle_Tool("build", "build and manage Isabelle sessions", |
|
312 Scala_Project.here, |
|
313 { args => |
|
314 var afp_root: Option[Path] = None |
|
315 val base_sessions = new mutable.ListBuffer[String] |
|
316 val select_dirs = new mutable.ListBuffer[Path] |
|
317 val build_hosts = new mutable.ListBuffer[Build_Cluster.Host] |
|
318 var numa_shuffling = false |
|
319 var browser_info = Browser_Info.Config.none |
|
320 var requirements = false |
|
321 var soft_build = false |
|
322 val exclude_session_groups = new mutable.ListBuffer[String] |
|
323 var all_sessions = false |
|
324 var build_heap = false |
|
325 var clean_build = false |
|
326 val dirs = new mutable.ListBuffer[Path] |
|
327 var export_files = false |
|
328 var fresh_build = false |
|
329 val session_groups = new mutable.ListBuffer[String] |
|
330 var max_jobs = 1 |
|
331 var check_keywords: Set[String] = Set.empty |
|
332 var list_files = false |
|
333 var no_build = false |
|
334 var options = Options.init(specs = Options.Spec.ISABELLE_BUILD_OPTIONS) |
|
335 var verbose = false |
|
336 val exclude_sessions = new mutable.ListBuffer[String] |
|
337 |
|
338 val getopts = Getopts(""" |
|
339 Usage: isabelle build [OPTIONS] [SESSIONS ...] |
|
340 |
|
341 Options are: |
|
342 -A ROOT include AFP with given root directory (":" for """ + AFP.BASE.implode + """) |
|
343 -B NAME include session NAME and all descendants |
|
344 -D DIR include session directory and select its sessions |
|
345 -H HOSTS additional build cluster host specifications, of the form |
|
346 "NAMES:PARAMETERS" (separated by commas) |
|
347 -N cyclic shuffling of NUMA CPU nodes (performance tuning) |
|
348 -P DIR enable HTML/PDF presentation in directory (":" for default) |
|
349 -R refer to requirements of selected sessions |
|
350 -S soft build: only observe changes of sources, not heap images |
|
351 -X NAME exclude sessions from group NAME and all descendants |
|
352 -a select all sessions |
|
353 -b build heap images |
|
354 -c clean build |
|
355 -d DIR include session directory |
|
356 -e export files from session specification into file-system |
|
357 -f fresh build |
|
358 -g NAME select session group NAME |
|
359 -j INT maximum number of parallel jobs (default 1) |
|
360 -k KEYWORD check theory sources for conflicts with proposed keywords |
|
361 -l list session source files |
|
362 -n no build -- take existing session build databases |
|
363 -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) |
|
364 -v verbose |
|
365 -x NAME exclude session NAME and all descendants |
|
366 |
|
367 Build and manage Isabelle sessions: ML heaps, session databases, documents. |
|
368 |
|
369 Parameters for host specifications (option -H), apart from system options: |
|
370 """ + Library.indent_lines(4, Build_Cluster.Host.parameters.print()) + |
|
371 """ |
|
372 |
|
373 Notable system options: see "isabelle options -l -t build" |
|
374 |
|
375 Notable system settings: |
|
376 """ + Library.indent_lines(4, Build_Log.Settings.show()) + "\n", |
|
377 "A:" -> (arg => afp_root = Some(if (arg == ":") AFP.BASE else Path.explode(arg))), |
|
378 "B:" -> (arg => base_sessions += arg), |
|
379 "D:" -> (arg => select_dirs += Path.explode(arg)), |
|
380 "H:" -> (arg => build_hosts ++= Build_Cluster.Host.parse(Registry.global, arg)), |
|
381 "N" -> (_ => numa_shuffling = true), |
|
382 "P:" -> (arg => browser_info = Browser_Info.Config.make(arg)), |
|
383 "R" -> (_ => requirements = true), |
|
384 "S" -> (_ => soft_build = true), |
|
385 "X:" -> (arg => exclude_session_groups += arg), |
|
386 "a" -> (_ => all_sessions = true), |
|
387 "b" -> (_ => build_heap = true), |
|
388 "c" -> (_ => clean_build = true), |
|
389 "d:" -> (arg => dirs += Path.explode(arg)), |
|
390 "e" -> (_ => export_files = true), |
|
391 "f" -> (_ => fresh_build = true), |
|
392 "g:" -> (arg => session_groups += arg), |
|
393 "j:" -> (arg => max_jobs = Value.Int.parse(arg)), |
|
394 "k:" -> (arg => check_keywords = check_keywords + arg), |
|
395 "l" -> (_ => list_files = true), |
|
396 "n" -> (_ => no_build = true), |
|
397 "o:" -> (arg => options = options + arg), |
|
398 "v" -> (_ => verbose = true), |
|
399 "x:" -> (arg => exclude_sessions += arg)) |
|
400 |
|
401 val sessions = getopts(args) |
|
402 |
|
403 val progress = new Console_Progress(verbose = verbose) |
|
404 |
|
405 progress.echo( |
|
406 "Started at " + Build_Log.print_date(progress.start) + |
|
407 " (" + Isabelle_System.ml_identifier() + " on " + hostname(options) +")", |
|
408 verbose = true) |
|
409 progress.echo(Build_Log.Settings.show() + "\n", verbose = true) |
|
410 |
|
411 val results = |
|
412 progress.interrupt_handler { |
|
413 build(options, |
|
414 selection = Sessions.Selection( |
|
415 requirements = requirements, |
|
416 all_sessions = all_sessions, |
|
417 base_sessions = base_sessions.toList, |
|
418 exclude_session_groups = exclude_session_groups.toList, |
|
419 exclude_sessions = exclude_sessions.toList, |
|
420 session_groups = session_groups.toList, |
|
421 sessions = sessions), |
|
422 browser_info = browser_info, |
|
423 progress = progress, |
|
424 check_unknown_files = Mercurial.is_repository(Path.ISABELLE_HOME), |
|
425 build_heap = build_heap, |
|
426 clean_build = clean_build, |
|
427 afp_root = afp_root, |
|
428 dirs = dirs.toList, |
|
429 select_dirs = select_dirs.toList, |
|
430 numa_shuffling = Host.numa_check(progress, numa_shuffling), |
|
431 max_jobs = max_jobs, |
|
432 list_files = list_files, |
|
433 check_keywords = check_keywords, |
|
434 fresh_build = fresh_build, |
|
435 no_build = no_build, |
|
436 soft_build = soft_build, |
|
437 export_files = export_files, |
|
438 build_hosts = build_hosts.toList) |
|
439 } |
|
440 val stop_date = Date.now() |
|
441 val elapsed_time = stop_date.time - progress.start.time |
|
442 |
|
443 progress.echo("\nFinished at " + Build_Log.print_date(stop_date), verbose = true) |
|
444 |
|
445 val total_timing = |
|
446 results.sessions.iterator.map(a => results(a).timing).foldLeft(Timing.zero)(_ + _). |
|
447 copy(elapsed = elapsed_time) |
|
448 progress.echo(total_timing.message_resources) |
|
449 |
|
450 sys.exit(results.rc) |
|
451 }) |
|
452 |
|
453 |
|
454 |
|
455 /** build cluster management **/ |
|
456 |
|
457 /* identified builds */ |
|
458 |
|
459 def read_builds(build_database: Option[SQL.Database]): List[Build_Process.Build] = |
|
460 build_database match { |
|
461 case None => Nil |
|
462 case Some(db) => Build_Process.read_builds(db) |
|
463 } |
|
464 |
|
465 def print_builds(build_database: Option[SQL.Database], builds: List[Build_Process.Build]): String = |
|
466 { |
|
467 val print_database = |
|
468 build_database match { |
|
469 case None => "" |
|
470 case Some(db) => " (database " + db + ")" |
|
471 } |
|
472 if (builds.isEmpty) "No build processes" + print_database |
|
473 else "Build processes" + print_database + builds.map(build => "\n " + build.print).mkString |
|
474 } |
|
475 |
|
476 def find_builds( |
|
477 build_database: Option[SQL.Database], |
|
478 build_id: String, |
|
479 builds: List[Build_Process.Build] |
|
480 ): Build_Process.Build = { |
|
481 (build_id, builds.length) match { |
|
482 case (UUID(_), _) if builds.exists(_.build_uuid == build_id) => |
|
483 builds.find(_.build_uuid == build_id).get |
|
484 case ("", 1) => builds.head |
|
485 case ("", 0) => error(print_builds(build_database, builds)) |
|
486 case _ => |
|
487 cat_error("Cannot identify build process " + quote(build_id), |
|
488 print_builds(build_database, builds)) |
|
489 } |
|
490 } |
|
491 |
|
492 |
|
493 /* "isabelle build_process" */ |
|
494 |
|
495 def build_process( |
|
496 options: Options, |
|
497 list_builds: Boolean = false, |
|
498 remove_builds: Boolean = false, |
|
499 force: Boolean = false, |
|
500 progress: Progress = new Progress |
|
501 ): Unit = { |
|
502 val build_engine = Engine(engine_name(options)) |
|
503 val store = build_engine.build_store(options) |
|
504 |
|
505 using(store.open_server()) { server => |
|
506 using_optional(store.maybe_open_build_database(server = server)) { build_database => |
|
507 def print(builds: List[Build_Process.Build]): Unit = |
|
508 if (list_builds) progress.echo(print_builds(build_database, builds)) |
|
509 |
|
510 build_database match { |
|
511 case None => print(Nil) |
|
512 case Some(db) => |
|
513 Build_Process.private_data.transaction_lock(db, create = true, label = "build_process") { |
|
514 val builds = Build_Process.private_data.read_builds(db) |
|
515 val remove = |
|
516 if (!remove_builds) Nil |
|
517 else if (force) builds.map(_.build_uuid) |
|
518 else builds.flatMap(build => if (build.active) None else Some(build.build_uuid)) |
|
519 |
|
520 print(builds) |
|
521 if (remove.nonEmpty) { |
|
522 if (remove_builds) { |
|
523 progress.echo("Removing " + commas(remove) + " ...") |
|
524 Build_Process.private_data.remove_builds(db, remove) |
|
525 print(Build_Process.private_data.read_builds(db)) |
|
526 } |
|
527 } |
|
528 } |
|
529 } |
|
530 } |
|
531 } |
|
532 } |
|
533 |
|
534 val isabelle_tool2 = Isabelle_Tool("build_process", "manage session build process", |
|
535 Scala_Project.here, |
|
536 { args => |
|
537 var force = false |
|
538 var list_builds = false |
|
539 var options = |
|
540 Options.init(specs = Options.Spec.ISABELLE_BUILD_OPTIONS ::: |
|
541 List( |
|
542 Options.Spec.make("build_database_server"), |
|
543 Options.Spec.make("build_database"))) |
|
544 var remove_builds = false |
|
545 |
|
546 val getopts = Getopts(""" |
|
547 Usage: isabelle build_process [OPTIONS] |
|
548 |
|
549 Options are: |
|
550 -f extra force for option -r |
|
551 -l list build processes |
|
552 -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) |
|
553 -r remove data from build processes: inactive processes (default) |
|
554 or all processes (option -f) |
|
555 """, |
|
556 "f" -> (_ => force = true), |
|
557 "l" -> (_ => list_builds = true), |
|
558 "o:" -> (arg => options = options + arg), |
|
559 "r" -> (_ => remove_builds = true)) |
|
560 |
|
561 val more_args = getopts(args) |
|
562 if (more_args.nonEmpty) getopts.usage() |
|
563 |
|
564 val progress = new Console_Progress() |
|
565 |
|
566 build_process(options, list_builds = list_builds, remove_builds = remove_builds, |
|
567 force = force, progress = progress) |
|
568 }) |
|
569 |
|
570 |
|
571 |
|
572 /* "isabelle build_worker" */ |
|
573 |
|
574 def build_worker_command( |
|
575 host: Build_Cluster.Host, |
|
576 ssh: SSH.System = SSH.Local, |
|
577 build_options: List[Options.Spec] = Nil, |
|
578 build_id: String = "", |
|
579 isabelle_home: Path = Path.current, |
|
580 afp_root: Option[Path] = None, |
|
581 dirs: List[Path] = Nil, |
|
582 quiet: Boolean = false, |
|
583 verbose: Boolean = false |
|
584 ): String = { |
|
585 val options = build_options ::: Options.Spec.eq("build_hostname", host.name) :: host.options |
|
586 ssh.bash_path(isabelle_home + Path.explode("bin/isabelle")) + " build_worker" + |
|
587 if_proper(build_id, " -B " + Bash.string(build_id)) + |
|
588 if_proper(afp_root, " -A " + ssh.bash_path(afp_root.get)) + |
|
589 dirs.map(dir => " -d " + ssh.bash_path(dir)).mkString + |
|
590 if_proper(host.numa, " -N") + " -j" + host.jobs + |
|
591 Options.Spec.bash_strings(options, bg = true) + |
|
592 if_proper(quiet, " -q") + |
|
593 if_proper(verbose, " -v") |
|
594 } |
|
595 |
|
596 def build_worker( |
|
597 options: Options, |
|
598 build_id: String = "", |
|
599 progress: Progress = new Progress, |
|
600 afp_root: Option[Path] = None, |
|
601 dirs: List[Path] = Nil, |
|
602 numa_shuffling: Boolean = false, |
|
603 max_jobs: Int = 1 |
|
604 ): Results = { |
|
605 val build_engine = Engine(engine_name(options)) |
|
606 val store = build_engine.build_store(options) |
|
607 val build_options = store.options |
|
608 |
|
609 using(store.open_server()) { server => |
|
610 using_optional(store.maybe_open_build_database(server = server)) { build_database => |
|
611 val builds = read_builds(build_database) |
|
612 |
|
613 val build_master = find_builds(build_database, build_id, builds.filter(_.active)) |
|
614 |
|
615 val sessions_structure = |
|
616 Sessions.load_structure(build_options, dirs = AFP.make_dirs(afp_root) ::: dirs). |
|
617 selection(Sessions.Selection(sessions = build_master.sessions)) |
|
618 |
|
619 val build_deps = |
|
620 Sessions.deps(sessions_structure, progress = progress, inlined_files = true).check_errors |
|
621 |
|
622 val build_context = |
|
623 Context(store, build_engine, build_deps, afp_root = afp_root, |
|
624 hostname = hostname(build_options), numa_shuffling = numa_shuffling, |
|
625 max_jobs = max_jobs, build_uuid = build_master.build_uuid) |
|
626 |
|
627 build_engine.run_build_process(build_context, progress, server) |
|
628 } |
|
629 } |
|
630 } |
|
631 |
|
632 val isabelle_tool3 = Isabelle_Tool("build_worker", "start worker for session build process", |
|
633 Scala_Project.here, |
|
634 { args => |
|
635 var afp_root: Option[Path] = None |
|
636 var build_id = "" |
|
637 var numa_shuffling = false |
|
638 val dirs = new mutable.ListBuffer[Path] |
|
639 var max_jobs = 1 |
|
640 var options = |
|
641 Options.init(specs = Options.Spec.ISABELLE_BUILD_OPTIONS ::: |
|
642 List( |
|
643 Options.Spec.make("build_database_server"), |
|
644 Options.Spec.make("build_database"))) |
|
645 var quiet = false |
|
646 var verbose = false |
|
647 |
|
648 val getopts = Getopts(""" |
|
649 Usage: isabelle build_worker [OPTIONS] |
|
650 |
|
651 Options are: |
|
652 -A ROOT include AFP with given root directory (":" for """ + AFP.BASE.implode + """) |
|
653 -B UUID existing UUID for build process (default: from database) |
|
654 -N cyclic shuffling of NUMA CPU nodes (performance tuning) |
|
655 -d DIR include session directory |
|
656 -j INT maximum number of parallel jobs (default 1) |
|
657 -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) |
|
658 -q quiet mode: no progress |
|
659 -v verbose |
|
660 """, |
|
661 "A:" -> (arg => afp_root = Some(if (arg == ":") AFP.BASE else Path.explode(arg))), |
|
662 "B:" -> (arg => build_id = arg), |
|
663 "N" -> (_ => numa_shuffling = true), |
|
664 "d:" -> (arg => dirs += Path.explode(arg)), |
|
665 "j:" -> (arg => max_jobs = Value.Int.parse(arg)), |
|
666 "o:" -> (arg => options = options + arg), |
|
667 "q" -> (_ => quiet = true), |
|
668 "v" -> (_ => verbose = true)) |
|
669 |
|
670 val more_args = getopts(args) |
|
671 if (more_args.nonEmpty) getopts.usage() |
|
672 |
|
673 val progress = |
|
674 if (quiet && verbose) new Progress { override def verbose: Boolean = true } |
|
675 else if (quiet) new Progress |
|
676 else new Console_Progress(verbose = verbose) |
|
677 |
|
678 val results = |
|
679 progress.interrupt_handler { |
|
680 build_worker(options, |
|
681 build_id = build_id, |
|
682 progress = progress, |
|
683 afp_root = afp_root, |
|
684 dirs = dirs.toList, |
|
685 numa_shuffling = Host.numa_check(progress, numa_shuffling), |
|
686 max_jobs = max_jobs) |
|
687 } |
|
688 |
|
689 if (!results.ok) sys.exit(results.rc) |
|
690 }) |
|
691 |
|
692 |
|
693 |
|
694 /** "isabelle build_log" **/ |
|
695 |
|
696 /* theory markup/messages from session database */ |
|
697 |
|
698 def read_theory( |
|
699 theory_context: Export.Theory_Context, |
|
700 unicode_symbols: Boolean = false |
|
701 ): Option[Document.Snapshot] = { |
|
702 def decode_bytes(bytes: Bytes): String = |
|
703 Symbol.output(unicode_symbols, UTF8.decode_permissive(bytes)) |
|
704 |
|
705 def read(name: String): Export.Entry = theory_context(name, permissive = true) |
|
706 |
|
707 def read_xml(name: String): XML.Body = |
|
708 YXML.parse_body(decode_bytes(read(name).bytes), cache = theory_context.cache) |
|
709 |
|
710 def read_source_file(name: String): Store.Source_File = |
|
711 theory_context.session_context.source_file(name) |
|
712 |
|
713 for { |
|
714 id <- theory_context.document_id() |
|
715 (thy_file, blobs_files) <- theory_context.files(permissive = true) |
|
716 } |
|
717 yield { |
|
718 val master_dir = |
|
719 Path.explode(Url.strip_base_name(thy_file).getOrElse( |
|
720 error("Cannot determine theory master directory: " + quote(thy_file)))) |
|
721 |
|
722 val blobs = |
|
723 blobs_files.map { name => |
|
724 val path = Path.explode(name) |
|
725 val src_path = File.relative_path(master_dir, path).getOrElse(path) |
|
726 |
|
727 val file = read_source_file(name) |
|
728 val bytes = file.bytes |
|
729 val text = decode_bytes(bytes) |
|
730 val chunk = Symbol.Text_Chunk(text) |
|
731 |
|
732 Command.Blob(Document.Node.Name(name), src_path, Some((file.digest, chunk))) -> |
|
733 Document.Blobs.Item(bytes, text, chunk, changed = false) |
|
734 } |
|
735 |
|
736 val thy_source = decode_bytes(read_source_file(thy_file).bytes) |
|
737 val thy_xml = read_xml(Export.MARKUP) |
|
738 val blobs_xml = |
|
739 for (i <- (1 to blobs.length).toList) |
|
740 yield read_xml(Export.MARKUP + i) |
|
741 |
|
742 val markups_index = Command.Markup_Index.make(blobs.map(_._1)) |
|
743 val markups = |
|
744 Command.Markups.make( |
|
745 for ((index, xml) <- markups_index.zip(thy_xml :: blobs_xml)) |
|
746 yield index -> Markup_Tree.from_XML(xml)) |
|
747 |
|
748 val results = |
|
749 Command.Results.make( |
|
750 for (case elem@XML.Elem(Markup(_, Markup.Serial(i)), _) <- read_xml(Export.MESSAGES)) |
|
751 yield i -> elem) |
|
752 |
|
753 val command = |
|
754 Command.unparsed(thy_source, theory = true, id = id, |
|
755 node_name = Document.Node.Name(thy_file, theory = theory_context.theory), |
|
756 blobs_info = Command.Blobs_Info.make(blobs), |
|
757 markups = markups, results = results) |
|
758 |
|
759 val doc_blobs = Document.Blobs.make(blobs) |
|
760 |
|
761 Document.State.init.snippet(command, doc_blobs) |
|
762 } |
|
763 } |
|
764 |
|
765 |
|
766 /* print messages */ |
|
767 |
|
768 def print_log( |
|
769 options: Options, |
|
770 sessions: List[String], |
|
771 theories: List[String] = Nil, |
|
772 message_head: List[Regex] = Nil, |
|
773 message_body: List[Regex] = Nil, |
|
774 progress: Progress = new Progress, |
|
775 margin: Double = Pretty.default_margin, |
|
776 breakgain: Double = Pretty.default_breakgain, |
|
777 metric: Pretty.Metric = Symbol.Metric, |
|
778 unicode_symbols: Boolean = false |
|
779 ): Unit = { |
|
780 val store = Store(options) |
|
781 val session = new Session(options, Resources.bootstrap) |
|
782 |
|
783 def check(filter: List[Regex], make_string: => String): Boolean = |
|
784 filter.isEmpty || { |
|
785 val s = Output.clean_yxml(make_string) |
|
786 filter.forall(r => r.findFirstIn(Output.clean_yxml(s)).nonEmpty) |
|
787 } |
|
788 |
|
789 def print(session_name: String): Unit = { |
|
790 using(Export.open_session_context0(store, session_name)) { session_context => |
|
791 val result = |
|
792 for { |
|
793 db <- session_context.session_db() |
|
794 theories = store.read_theories(db, session_name) |
|
795 errors = store.read_errors(db, session_name) |
|
796 info <- store.read_build(db, session_name) |
|
797 } yield (theories, errors, info.return_code) |
|
798 result match { |
|
799 case None => store.error_database(session_name) |
|
800 case Some((used_theories, errors, rc)) => |
|
801 theories.filterNot(used_theories.toSet) match { |
|
802 case Nil => |
|
803 case bad => error("Unknown theories " + commas_quote(bad)) |
|
804 } |
|
805 val print_theories = |
|
806 if (theories.isEmpty) used_theories else used_theories.filter(theories.toSet) |
|
807 |
|
808 for (thy <- print_theories) { |
|
809 val thy_heading = "\nTheory " + quote(thy) + " (in " + session_name + ")" + ":" |
|
810 |
|
811 Build.read_theory(session_context.theory(thy), unicode_symbols = unicode_symbols) match { |
|
812 case None => progress.echo(thy_heading + " MISSING") |
|
813 case Some(snapshot) => |
|
814 val rendering = new Rendering(snapshot, options, session) |
|
815 val messages = |
|
816 rendering.text_messages(Text.Range.full) |
|
817 .filter(message => progress.verbose || Protocol.is_exported(message.info)) |
|
818 if (messages.nonEmpty) { |
|
819 val line_document = Line.Document(snapshot.node.source) |
|
820 val buffer = new mutable.ListBuffer[String] |
|
821 for (Text.Info(range, elem) <- messages) { |
|
822 val line = line_document.position(range.start).line1 |
|
823 val pos = Position.Line_File(line, snapshot.node_name.node) |
|
824 def message_text: String = |
|
825 Protocol.message_text(elem, heading = true, pos = pos, |
|
826 margin = margin, breakgain = breakgain, metric = metric) |
|
827 val ok = |
|
828 check(message_head, Protocol.message_heading(elem, pos)) && |
|
829 check(message_body, XML.content(Pretty.unformatted(List(elem)))) |
|
830 if (ok) buffer += message_text |
|
831 } |
|
832 if (buffer.nonEmpty) { |
|
833 progress.echo(thy_heading) |
|
834 buffer.foreach(progress.echo(_)) |
|
835 } |
|
836 } |
|
837 } |
|
838 } |
|
839 |
|
840 if (errors.nonEmpty) { |
|
841 val msg = Symbol.output(unicode_symbols, cat_lines(errors)) |
|
842 progress.echo("\nBuild errors:\n" + Output.error_message_text(msg)) |
|
843 } |
|
844 if (rc != Process_Result.RC.ok) { |
|
845 progress.echo("\n" + Process_Result.RC.print_long(rc)) |
|
846 } |
|
847 } |
|
848 } |
|
849 } |
|
850 |
|
851 val errors = new mutable.ListBuffer[String] |
|
852 for (session_name <- sessions) { |
|
853 Exn.result(print(session_name)) match { |
|
854 case Exn.Res(_) => |
|
855 case Exn.Exn(exn) => errors += Exn.message(exn) |
|
856 } |
|
857 } |
|
858 if (errors.nonEmpty) error(cat_lines(errors.toList)) |
|
859 } |
|
860 |
|
861 |
|
862 /* command-line wrapper */ |
|
863 |
|
864 val isabelle_tool4 = Isabelle_Tool("build_log", "print messages from session build database", |
|
865 Scala_Project.here, |
|
866 { args => |
|
867 /* arguments */ |
|
868 |
|
869 val message_head = new mutable.ListBuffer[Regex] |
|
870 val message_body = new mutable.ListBuffer[Regex] |
|
871 var unicode_symbols = false |
|
872 val theories = new mutable.ListBuffer[String] |
|
873 var margin = Pretty.default_margin |
|
874 var options = Options.init() |
|
875 var verbose = false |
|
876 |
|
877 val getopts = Getopts(""" |
|
878 Usage: isabelle build_log [OPTIONS] [SESSIONS ...] |
|
879 |
|
880 Options are: |
|
881 -H REGEX filter messages by matching against head |
|
882 -M REGEX filter messages by matching against body |
|
883 -T NAME restrict to given theories (multiple options possible) |
|
884 -U output Unicode symbols |
|
885 -m MARGIN margin for pretty printing (default: """ + margin + """) |
|
886 -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) |
|
887 -v print all messages, including information etc. |
|
888 |
|
889 Print messages from the session build database of the given sessions, |
|
890 without any checks against current sources nor session structure: results |
|
891 from old sessions or failed builds can be printed as well. |
|
892 |
|
893 Multiple options -H and -M are conjunctive: all given patterns need to |
|
894 match. Patterns match any substring, but ^ or $ may be used to match the |
|
895 start or end explicitly. |
|
896 """, |
|
897 "H:" -> (arg => message_head += arg.r), |
|
898 "M:" -> (arg => message_body += arg.r), |
|
899 "T:" -> (arg => theories += arg), |
|
900 "U" -> (_ => unicode_symbols = true), |
|
901 "m:" -> (arg => margin = Value.Double.parse(arg)), |
|
902 "o:" -> (arg => options = options + arg), |
|
903 "v" -> (_ => verbose = true)) |
|
904 |
|
905 val sessions = getopts(args) |
|
906 |
|
907 val progress = new Console_Progress(verbose = verbose) |
|
908 |
|
909 if (sessions.isEmpty) progress.echo_warning("No sessions to print") |
|
910 else { |
|
911 print_log(options, sessions, theories = theories.toList, message_head = message_head.toList, |
|
912 message_body = message_body.toList, margin = margin, progress = progress, |
|
913 unicode_symbols = unicode_symbols) |
|
914 } |
|
915 }) |
|
916 } |
|