| author | wenzelm |
| Thu, 02 Apr 2020 12:19:09 +0200 | |
| changeset 71655 | dad29591645a |
| parent 71652 | 721f143a679b |
| child 71667 | 4d2de35214c5 |
| permissions | -rw-r--r-- |
| 50686 | 1 |
/* Title: Pure/Tools/build.scala |
| 48276 | 2 |
Author: Makarius |
| 57923 | 3 |
Options: :folding=explicit: |
| 48276 | 4 |
|
5 |
Build and manage Isabelle sessions. |
|
6 |
*/ |
|
7 |
||
8 |
package isabelle |
|
9 |
||
10 |
||
| 71648 | 11 |
import scala.collection.{SortedSet, mutable}
|
|
48340
6f4fc030882a
allow explicit specification of additional session directories;
wenzelm
parents:
48339
diff
changeset
|
12 |
import scala.annotation.tailrec |
|
48337
9c7f8e5805b4
cumulate semantic Session_Info, based on syntactic Session_Entry;
wenzelm
parents:
48336
diff
changeset
|
13 |
|
| 48335 | 14 |
|
| 48276 | 15 |
object Build |
16 |
{
|
|
| 62631 | 17 |
/** auxiliary **/ |
| 48424 | 18 |
|
|
65291
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
19 |
/* persistent build info */ |
|
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
20 |
|
|
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
21 |
sealed case class Session_Info( |
|
66749
0445cfaf6132
more compact (second-order) digest for 10^2..10^3 source files, with slightly increased risk of collisions;
wenzelm
parents:
66747
diff
changeset
|
22 |
sources: String, |
|
65291
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
23 |
input_heaps: List[String], |
|
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
24 |
output_heap: Option[String], |
|
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
25 |
return_code: Int) |
| 66747 | 26 |
{
|
27 |
def ok: Boolean = return_code == 0 |
|
28 |
} |
|
|
65291
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
29 |
|
|
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
30 |
|
| 65289 | 31 |
/* queue with scheduling information */ |
|
48676
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
32 |
|
| 62631 | 33 |
private object Queue |
|
48676
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
34 |
{
|
|
68086
9e1c670301b8
cleanup session output before starting build job;
wenzelm
parents:
67852
diff
changeset
|
35 |
type Timings = (List[Properties.T], Double) |
|
9e1c670301b8
cleanup session output before starting build job;
wenzelm
parents:
67852
diff
changeset
|
36 |
|
|
9e1c670301b8
cleanup session output before starting build job;
wenzelm
parents:
67852
diff
changeset
|
37 |
def load_timings(progress: Progress, store: Sessions.Store, name: String): Timings = |
| 65289 | 38 |
{
|
|
68086
9e1c670301b8
cleanup session output before starting build job;
wenzelm
parents:
67852
diff
changeset
|
39 |
val no_timings: Timings = (Nil, 0.0) |
| 65289 | 40 |
|
| 68221 | 41 |
store.access_database(name) match {
|
|
65291
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
42 |
case None => no_timings |
| 68214 | 43 |
case Some(db) => |
|
65291
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
44 |
def ignore_error(msg: String) = |
|
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
45 |
{
|
| 68214 | 46 |
progress.echo_warning("Ignoring bad database " + db + (if (msg == "") "" else "\n" + msg))
|
|
65291
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
47 |
no_timings |
|
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
48 |
} |
|
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
49 |
try {
|
| 68214 | 50 |
val command_timings = store.read_command_timings(db, name) |
51 |
val session_timing = |
|
52 |
store.read_session_timing(db, name) match {
|
|
53 |
case Markup.Elapsed(t) => t |
|
54 |
case _ => 0.0 |
|
55 |
} |
|
56 |
(command_timings, session_timing) |
|
|
65291
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
57 |
} |
|
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
58 |
catch {
|
|
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
59 |
case ERROR(msg) => ignore_error(msg) |
|
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
60 |
case exn: java.lang.Error => ignore_error(Exn.message(exn)) |
|
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
61 |
case _: XML.Error => ignore_error("")
|
|
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
62 |
} |
| 68214 | 63 |
finally { db.close }
|
| 65289 | 64 |
} |
65 |
} |
|
66 |
||
| 68731 | 67 |
def make_session_timing(sessions_structure: Sessions.Structure, timing: Map[String, Double]) |
|
68486
6984a55f3cba
clarified queue ordering: take session descendants into account, notably for "slow" AFP sessions;
wenzelm
parents:
68331
diff
changeset
|
68 |
: Map[String, Double] = |
|
6984a55f3cba
clarified queue ordering: take session descendants into account, notably for "slow" AFP sessions;
wenzelm
parents:
68331
diff
changeset
|
69 |
{
|
| 68731 | 70 |
val maximals = sessions_structure.build_graph.maximals.toSet |
|
68486
6984a55f3cba
clarified queue ordering: take session descendants into account, notably for "slow" AFP sessions;
wenzelm
parents:
68331
diff
changeset
|
71 |
def desc_timing(name: String): Double = |
|
6984a55f3cba
clarified queue ordering: take session descendants into account, notably for "slow" AFP sessions;
wenzelm
parents:
68331
diff
changeset
|
72 |
{
|
| 68487 | 73 |
if (maximals.contains(name)) timing(name) |
74 |
else {
|
|
| 68731 | 75 |
val descendants = sessions_structure.build_descendants(List(name)).toSet |
76 |
val g = sessions_structure.build_graph.restrict(descendants) |
|
| 68487 | 77 |
(0.0 :: g.maximals.flatMap(desc => {
|
78 |
val ps = g.all_preds(List(desc)) |
|
79 |
if (ps.exists(p => timing.get(p).isEmpty)) None |
|
80 |
else Some(ps.map(timing(_)).sum) |
|
81 |
})).max |
|
82 |
} |
|
|
68486
6984a55f3cba
clarified queue ordering: take session descendants into account, notably for "slow" AFP sessions;
wenzelm
parents:
68331
diff
changeset
|
83 |
} |
|
6984a55f3cba
clarified queue ordering: take session descendants into account, notably for "slow" AFP sessions;
wenzelm
parents:
68331
diff
changeset
|
84 |
timing.keySet.iterator.map(name => (name -> desc_timing(name))).toMap.withDefaultValue(0.0) |
|
6984a55f3cba
clarified queue ordering: take session descendants into account, notably for "slow" AFP sessions;
wenzelm
parents:
68331
diff
changeset
|
85 |
} |
|
6984a55f3cba
clarified queue ordering: take session descendants into account, notably for "slow" AFP sessions;
wenzelm
parents:
68331
diff
changeset
|
86 |
|
| 68731 | 87 |
def apply(progress: Progress, sessions_structure: Sessions.Structure, store: Sessions.Store) |
88 |
: Queue = |
|
|
48676
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
89 |
{
|
| 68731 | 90 |
val graph = sessions_structure.build_graph |
|
65415
8cd54b18b68b
clarified signature: tree structure is not essential;
wenzelm
parents:
65406
diff
changeset
|
91 |
val names = graph.keys |
| 51220 | 92 |
|
| 65831 | 93 |
val timings = names.map(name => (name, load_timings(progress, store, name))) |
| 51220 | 94 |
val command_timings = |
|
68486
6984a55f3cba
clarified queue ordering: take session descendants into account, notably for "slow" AFP sessions;
wenzelm
parents:
68331
diff
changeset
|
95 |
timings.map({ case (name, (ts, _)) => (name, ts) }).toMap.withDefaultValue(Nil)
|
| 51220 | 96 |
val session_timing = |
| 68731 | 97 |
make_session_timing(sessions_structure, |
98 |
timings.map({ case (name, (_, t)) => (name, t) }).toMap)
|
|
|
48678
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
99 |
|
|
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
100 |
object Ordering extends scala.math.Ordering[String] |
|
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
101 |
{
|
|
51227
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
102 |
def compare_timing(name1: String, name2: String): Int = |
|
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
103 |
{
|
|
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
104 |
val t1 = session_timing(name1) |
|
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
105 |
val t2 = session_timing(name2) |
|
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
106 |
if (t1 == 0.0 || t2 == 0.0) 0 |
|
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
107 |
else t1 compare t2 |
|
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
108 |
} |
|
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
109 |
|
|
48678
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
110 |
def compare(name1: String, name2: String): Int = |
|
68486
6984a55f3cba
clarified queue ordering: take session descendants into account, notably for "slow" AFP sessions;
wenzelm
parents:
68331
diff
changeset
|
111 |
compare_timing(name2, name1) match {
|
|
48678
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
112 |
case 0 => |
| 68731 | 113 |
sessions_structure(name2).timeout compare sessions_structure(name1).timeout match {
|
|
68486
6984a55f3cba
clarified queue ordering: take session descendants into account, notably for "slow" AFP sessions;
wenzelm
parents:
68331
diff
changeset
|
114 |
case 0 => name1 compare name2 |
|
48678
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
115 |
case ord => ord |
|
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
116 |
} |
|
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
117 |
case ord => ord |
|
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
118 |
} |
|
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
119 |
} |
|
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
120 |
|
|
65415
8cd54b18b68b
clarified signature: tree structure is not essential;
wenzelm
parents:
65406
diff
changeset
|
121 |
new Queue(graph, SortedSet(names: _*)(Ordering), command_timings) |
|
48676
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
122 |
} |
|
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
123 |
} |
|
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
124 |
|
| 65314 | 125 |
private class Queue( |
| 62631 | 126 |
graph: Graph[String, Sessions.Info], |
| 51220 | 127 |
order: SortedSet[String], |
128 |
val command_timings: String => List[Properties.T]) |
|
|
48676
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
129 |
{
|
|
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
130 |
def is_inner(name: String): Boolean = !graph.is_maximal(name) |
|
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
131 |
|
|
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
132 |
def is_empty: Boolean = graph.is_empty |
|
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
133 |
|
|
51227
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
134 |
def - (name: String): Queue = |
|
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
135 |
new Queue(graph.del_node(name), |
| 67011 | 136 |
order - name, // FIXME scala-2.10.0 .. 2.12.4 TreeSet problem!? |
|
51227
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
137 |
command_timings) |
|
48676
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
138 |
|
| 62631 | 139 |
def dequeue(skip: String => Boolean): Option[(String, Sessions.Info)] = |
|
48676
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
140 |
{
|
|
51227
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
141 |
val it = order.iterator.dropWhile(name => |
|
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
142 |
skip(name) |
| 67011 | 143 |
|| !graph.defined(name) // FIXME scala-2.10.0 .. 2.12.4 TreeSet problem!? |
|
51227
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
144 |
|| !graph.is_minimal(name)) |
| 48680 | 145 |
if (it.hasNext) { val name = it.next; Some((name, graph.get_node(name))) }
|
|
48676
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
146 |
else None |
|
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
147 |
} |
|
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
148 |
} |
|
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
149 |
|
|
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
150 |
|
| 65313 | 151 |
/* PIDE protocol handler */ |
152 |
||
153 |
class Handler(progress: Progress, session: Session, session_name: String) |
|
154 |
extends Session.Protocol_Handler |
|
155 |
{
|
|
| 71631 | 156 |
val build_session_errors: Promise[List[String]] = Future.promise |
| 65313 | 157 |
|
| 71631 | 158 |
override def exit() { build_session_errors.cancel }
|
| 65313 | 159 |
|
160 |
private def build_session_finished(msg: Prover.Protocol_Output): Boolean = |
|
161 |
{
|
|
| 71631 | 162 |
val errors = |
163 |
try {
|
|
| 71649 | 164 |
for (err <- XML.Decode.list(x => x)(Symbol.decode_yxml(msg.text))) yield |
165 |
Pretty.string_of(Protocol_Message.expose_no_reports(err), metric = Symbol.Metric) |
|
| 71631 | 166 |
} |
167 |
catch { case ERROR(err) => List(err) }
|
|
168 |
build_session_errors.fulfill(errors) |
|
| 65313 | 169 |
session.send_stop() |
170 |
true |
|
171 |
} |
|
172 |
||
173 |
private def loading_theory(msg: Prover.Protocol_Output): Boolean = |
|
174 |
msg.properties match {
|
|
175 |
case Markup.Loading_Theory(name) => |
|
| 68957 | 176 |
progress.theory(Progress.Theory(name, session = session_name)) |
| 65313 | 177 |
true |
178 |
case _ => false |
|
179 |
} |
|
180 |
||
181 |
val functions = |
|
182 |
List( |
|
| 71610 | 183 |
Markup.BUILD_SESSION_FINISHED -> build_session_finished, |
184 |
Markup.LOADING_THEORY -> loading_theory) |
|
| 65313 | 185 |
} |
186 |
||
187 |
||
| 65308 | 188 |
/* job: running prover process */ |
| 48341 | 189 |
|
| 65308 | 190 |
private class Job(progress: Progress, |
191 |
name: String, |
|
192 |
val info: Sessions.Info, |
|
| 65313 | 193 |
deps: Sessions.Deps, |
| 65308 | 194 |
store: Sessions.Store, |
| 71614 | 195 |
do_store: Boolean, |
| 65308 | 196 |
verbose: Boolean, |
197 |
val numa_node: Option[Int], |
|
198 |
command_timings: List[Properties.T]) |
|
| 48418 | 199 |
{
|
| 71610 | 200 |
private val options = NUMA.policy_options(info.options, numa_node) |
| 65312 | 201 |
|
| 71610 | 202 |
private val sessions_structure = deps.sessions_structure |
|
70683
8c7706b053c7
find theory files via session structure: much faster Prover IDE startup;
wenzelm
parents:
70671
diff
changeset
|
203 |
|
| 59445 | 204 |
private val graph_file = Isabelle_System.tmp_file("session_graph", "pdf")
|
| 66822 | 205 |
isabelle.graphview.Graph_File.write(options, graph_file, deps(name).session_graph_display) |
| 59445 | 206 |
|
|
68198
6710167e17af
avoid race condition wrt. ISABELLE_TMP, which is removed in Bash.cleanup() before Bash.result(progress_stdout);
wenzelm
parents:
68148
diff
changeset
|
207 |
private val export_tmp_dir = Isabelle_System.tmp_dir("export")
|
| 68289 | 208 |
private val export_consumer = |
209 |
Export.consumer(store.open_database(name, output = true), cache = store.xz_cache) |
|
|
68092
888d35a19866
store exports in session database, with asynchronous / parallel compression;
wenzelm
parents:
68088
diff
changeset
|
210 |
|
| 62573 | 211 |
private val future_result: Future[Process_Result] = |
|
61559
313eca3fa847
more direct task future implementation, with proper cancel operation;
wenzelm
parents:
61556
diff
changeset
|
212 |
Future.thread("build") {
|
| 65308 | 213 |
val parent = info.parent.getOrElse("")
|
| 71645 | 214 |
val base = deps(parent) |
| 65308 | 215 |
val args_yxml = |
216 |
YXML.string_of_body( |
|
|
62944
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
217 |
{
|
|
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
218 |
import XML.Encode._ |
|
71611
fb6953e77000
eliminated pointless flag (see also 6533ceee4cd7);
wenzelm
parents:
71610
diff
changeset
|
219 |
pair(list(pair(string, int)), pair(list(properties), pair(bool, |
|
62944
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
220 |
pair(Path.encode, pair(list(pair(Path.encode, Path.encode)), pair(string, |
| 65307 | 221 |
pair(string, pair(string, pair(string, pair(Path.encode, |
| 65517 | 222 |
pair(list(pair(Options.encode, list(pair(string, properties)))), |
|
70683
8c7706b053c7
find theory files via session structure: much faster Prover IDE startup;
wenzelm
parents:
70671
diff
changeset
|
223 |
pair(list(pair(string, properties)), |
|
67493
c4e9e0c50487
treat sessions as entities with defining position;
wenzelm
parents:
67471
diff
changeset
|
224 |
pair(list(pair(string, string)), |
|
70683
8c7706b053c7
find theory files via session structure: much faster Prover IDE startup;
wenzelm
parents:
70671
diff
changeset
|
225 |
pair(list(string), pair(list(pair(string, string)), |
|
71611
fb6953e77000
eliminated pointless flag (see also 6533ceee4cd7);
wenzelm
parents:
71610
diff
changeset
|
226 |
pair(list(string), list(string)))))))))))))))))( |
|
fb6953e77000
eliminated pointless flag (see also 6533ceee4cd7);
wenzelm
parents:
71610
diff
changeset
|
227 |
(Symbol.codes, (command_timings, (verbose, |
|
62944
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
228 |
(store.browser_info, (info.document_files, (File.standard_path(graph_file), |
| 65307 | 229 |
(parent, (info.chapter, (name, (Path.current, |
|
70683
8c7706b053c7
find theory files via session structure: much faster Prover IDE startup;
wenzelm
parents:
70671
diff
changeset
|
230 |
(info.theories, |
|
8c7706b053c7
find theory files via session structure: much faster Prover IDE startup;
wenzelm
parents:
70671
diff
changeset
|
231 |
(sessions_structure.session_positions, |
|
8c7706b053c7
find theory files via session structure: much faster Prover IDE startup;
wenzelm
parents:
70671
diff
changeset
|
232 |
(sessions_structure.dest_session_directories, |
|
8c7706b053c7
find theory files via session structure: much faster Prover IDE startup;
wenzelm
parents:
70671
diff
changeset
|
233 |
(base.doc_names, (base.global_theories.toList, |
|
71611
fb6953e77000
eliminated pointless flag (see also 6533ceee4cd7);
wenzelm
parents:
71610
diff
changeset
|
234 |
(base.loaded_theories.keys, info.bibtex_entries.map(_.info)))))))))))))))))) |
| 65308 | 235 |
}) |
236 |
||
237 |
val env = |
|
238 |
Isabelle_System.settings() + |
|
|
68198
6710167e17af
avoid race condition wrt. ISABELLE_TMP, which is removed in Bash.cleanup() before Bash.result(progress_stdout);
wenzelm
parents:
68148
diff
changeset
|
239 |
("ISABELLE_EXPORT_TMP" -> File.standard_path(export_tmp_dir)) +
|
| 65312 | 240 |
("ISABELLE_ML_DEBUGGER" -> options.bool("ML_debugger").toString)
|
| 65308 | 241 |
|
| 71610 | 242 |
val is_pure = Sessions.is_pure(name) |
243 |
||
|
71639
ec84f542e411
clarified signature of ML_Process vs. Isabelle_Process: proper support for "isabelle build -P -b";
wenzelm
parents:
71632
diff
changeset
|
244 |
val use_prelude = if (is_pure) Thy_Header.ml_roots.map(_._1) else Nil |
|
ec84f542e411
clarified signature of ML_Process vs. Isabelle_Process: proper support for "isabelle build -P -b";
wenzelm
parents:
71632
diff
changeset
|
245 |
|
| 71613 | 246 |
val eval_store = |
|
71639
ec84f542e411
clarified signature of ML_Process vs. Isabelle_Process: proper support for "isabelle build -P -b";
wenzelm
parents:
71632
diff
changeset
|
247 |
if (do_store) {
|
| 71613 | 248 |
(if (info.theories.nonEmpty) List("ML_Heap.share_common_data ()") else Nil) :::
|
249 |
List("ML_Heap.save_child " +
|
|
250 |
ML_Syntax.print_string_bytes(File.platform_path(store.output_heap(name)))) |
|
251 |
} |
|
|
71639
ec84f542e411
clarified signature of ML_Process vs. Isabelle_Process: proper support for "isabelle build -P -b";
wenzelm
parents:
71632
diff
changeset
|
252 |
else Nil |
|
62944
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
253 |
|
| 71651 | 254 |
if (options.bool("pide_build")) {
|
| 71643 | 255 |
val resources = new Resources(sessions_structure, deps(parent)) |
| 65313 | 256 |
val session = new Session(options, resources) |
257 |
val handler = new Handler(progress, session, name) |
|
| 65315 | 258 |
session.init_protocol_handler(handler) |
| 65313 | 259 |
|
| 71648 | 260 |
val stdout = new StringBuilder(1000) |
261 |
val messages = new mutable.ListBuffer[String] |
|
| 71652 | 262 |
val command_timings = new mutable.ListBuffer[Properties.T] |
263 |
val theory_timings = new mutable.ListBuffer[Properties.T] |
|
264 |
val runtime_statistics = new mutable.ListBuffer[Properties.T] |
|
265 |
val task_statistics = new mutable.ListBuffer[Properties.T] |
|
| 71648 | 266 |
|
| 71652 | 267 |
val consumer = |
268 |
Session.Consumer[Any]("build_session_output") {
|
|
| 71648 | 269 |
case msg: Prover.Output => |
270 |
val message = msg.message |
|
271 |
if (msg.is_stdout) {
|
|
272 |
stdout ++= Symbol.encode(XML.content(message)) |
|
273 |
} |
|
274 |
else if (Protocol.is_exported(message)) {
|
|
| 71649 | 275 |
messages += |
276 |
Symbol.encode(Protocol.message_text(List(message), metric = Symbol.Metric)) |
|
| 71648 | 277 |
} |
| 71652 | 278 |
case Session.Command_Timing(props) => command_timings += props |
279 |
case Session.Theory_Timing(props) => theory_timings += props |
|
280 |
case Session.Runtime_Statistics(props) => runtime_statistics += props |
|
281 |
case Session.Task_Statistics(props) => task_statistics += props |
|
| 71648 | 282 |
case _ => |
283 |
} |
|
284 |
||
| 71652 | 285 |
session.all_messages += consumer |
286 |
session.command_timings += consumer |
|
287 |
session.theory_timings += consumer |
|
288 |
session.runtime_statistics += consumer |
|
289 |
session.task_statistics += consumer |
|
290 |
||
| 71644 | 291 |
val eval_main = Command_Line.ML_tool("Isabelle_Process.init_build ()" :: eval_store)
|
|
71639
ec84f542e411
clarified signature of ML_Process vs. Isabelle_Process: proper support for "isabelle build -P -b";
wenzelm
parents:
71632
diff
changeset
|
292 |
|
| 71607 | 293 |
val process = |
294 |
Isabelle_Process(session, options, sessions_structure, store, |
|
|
71639
ec84f542e411
clarified signature of ML_Process vs. Isabelle_Process: proper support for "isabelle build -P -b";
wenzelm
parents:
71632
diff
changeset
|
295 |
logic = parent, raw_ml_system = is_pure, |
|
ec84f542e411
clarified signature of ML_Process vs. Isabelle_Process: proper support for "isabelle build -P -b";
wenzelm
parents:
71632
diff
changeset
|
296 |
use_prelude = use_prelude, eval_main = eval_main, |
|
ec84f542e411
clarified signature of ML_Process vs. Isabelle_Process: proper support for "isabelle build -P -b";
wenzelm
parents:
71632
diff
changeset
|
297 |
cwd = info.dir.file, env = env).await_startup |
| 65313 | 298 |
|
| 71607 | 299 |
session.protocol_command("build_session", args_yxml)
|
300 |
||
| 71648 | 301 |
val process_result = process.join |
| 71652 | 302 |
val process_output = |
303 |
stdout.toString :: messages.toList ::: |
|
304 |
command_timings.toList.map(Protocol.Command_Timing_Marker.apply) ::: |
|
305 |
theory_timings.toList.map(Protocol.Theory_Timing_Marker.apply) ::: |
|
306 |
runtime_statistics.toList.map(Protocol.ML_Statistics_Marker.apply) ::: |
|
307 |
task_statistics.toList.map(Protocol.Task_Statistics_Marker.apply) |
|
| 71648 | 308 |
|
| 71652 | 309 |
val result = process_result.output(process_output) |
| 71631 | 310 |
handler.build_session_errors.join match {
|
311 |
case Nil => result |
|
312 |
case errors => |
|
313 |
result.error_rc.output( |
|
314 |
errors.flatMap(s => split_lines(Output.error_message_text(s))) ::: |
|
315 |
errors.map(Protocol.Error_Message_Marker.apply)) |
|
| 65313 | 316 |
} |
| 65308 | 317 |
} |
318 |
else {
|
|
319 |
val args_file = Isabelle_System.tmp_file("build")
|
|
320 |
File.write(args_file, args_yxml) |
|
321 |
||
|
71639
ec84f542e411
clarified signature of ML_Process vs. Isabelle_Process: proper support for "isabelle build -P -b";
wenzelm
parents:
71632
diff
changeset
|
322 |
val eval_build = |
|
ec84f542e411
clarified signature of ML_Process vs. Isabelle_Process: proper support for "isabelle build -P -b";
wenzelm
parents:
71632
diff
changeset
|
323 |
"Build.build " + ML_Syntax.print_string_bytes(File.standard_path(args_file)) |
|
ec84f542e411
clarified signature of ML_Process vs. Isabelle_Process: proper support for "isabelle build -P -b";
wenzelm
parents:
71632
diff
changeset
|
324 |
val eval_main = Command_Line.ML_tool(eval_build :: eval_store) |
| 64265 | 325 |
|
| 65308 | 326 |
val process = |
|
71639
ec84f542e411
clarified signature of ML_Process vs. Isabelle_Process: proper support for "isabelle build -P -b";
wenzelm
parents:
71632
diff
changeset
|
327 |
ML_Process(options, deps.sessions_structure, store, |
|
ec84f542e411
clarified signature of ML_Process vs. Isabelle_Process: proper support for "isabelle build -P -b";
wenzelm
parents:
71632
diff
changeset
|
328 |
logic = parent, raw_ml_system = is_pure, |
|
ec84f542e411
clarified signature of ML_Process vs. Isabelle_Process: proper support for "isabelle build -P -b";
wenzelm
parents:
71632
diff
changeset
|
329 |
use_prelude = use_prelude, eval_main = eval_main, |
|
ec84f542e411
clarified signature of ML_Process vs. Isabelle_Process: proper support for "isabelle build -P -b";
wenzelm
parents:
71632
diff
changeset
|
330 |
cwd = info.dir.file, env = env, cleanup = () => args_file.delete) |
| 64265 | 331 |
|
| 65308 | 332 |
process.result( |
| 71625 | 333 |
progress_stdout = |
334 |
{
|
|
|
71630
50425e4c3910
clarified modules: global quasi-scope for markers;
wenzelm
parents:
71625
diff
changeset
|
335 |
case Protocol.Loading_Theory_Marker(theory) => |
| 68957 | 336 |
progress.theory(Progress.Theory(theory, session = name)) |
| 71624 | 337 |
case Protocol.Export.Marker((args, path)) => |
|
71623
b3bddebe44ca
clarified signature: more explicit type Protocol_Message.Marker;
wenzelm
parents:
71614
diff
changeset
|
338 |
val body = |
|
b3bddebe44ca
clarified signature: more explicit type Protocol_Message.Marker;
wenzelm
parents:
71614
diff
changeset
|
339 |
try { Bytes.read(path) }
|
|
b3bddebe44ca
clarified signature: more explicit type Protocol_Message.Marker;
wenzelm
parents:
71614
diff
changeset
|
340 |
catch {
|
|
b3bddebe44ca
clarified signature: more explicit type Protocol_Message.Marker;
wenzelm
parents:
71614
diff
changeset
|
341 |
case ERROR(msg) => |
|
b3bddebe44ca
clarified signature: more explicit type Protocol_Message.Marker;
wenzelm
parents:
71614
diff
changeset
|
342 |
error("Failed to read export " + quote(args.compound_name) + "\n" + msg)
|
|
b3bddebe44ca
clarified signature: more explicit type Protocol_Message.Marker;
wenzelm
parents:
71614
diff
changeset
|
343 |
} |
|
b3bddebe44ca
clarified signature: more explicit type Protocol_Message.Marker;
wenzelm
parents:
71614
diff
changeset
|
344 |
path.file.delete |
|
b3bddebe44ca
clarified signature: more explicit type Protocol_Message.Marker;
wenzelm
parents:
71614
diff
changeset
|
345 |
export_consumer(name, args, body) |
|
b3bddebe44ca
clarified signature: more explicit type Protocol_Message.Marker;
wenzelm
parents:
71614
diff
changeset
|
346 |
case _ => |
| 65308 | 347 |
}, |
348 |
progress_limit = |
|
| 65312 | 349 |
options.int("process_output_limit") match {
|
| 65308 | 350 |
case 0 => None |
351 |
case m => Some(m * 1000000L) |
|
352 |
}, |
|
353 |
strict = false) |
|
354 |
} |
|
| 50845 | 355 |
} |
| 48674 | 356 |
|
| 62572 | 357 |
def terminate: Unit = future_result.cancel |
358 |
def is_finished: Boolean = future_result.is_finished |
|
| 48674 | 359 |
|
| 56779 | 360 |
private val timeout_request: Option[Event_Timer.Request] = |
361 |
{
|
|
| 62569 | 362 |
if (info.timeout > Time.zero) |
| 66943 | 363 |
Some(Event_Timer.request(Time.now() + info.timeout) { terminate })
|
| 48674 | 364 |
else None |
| 56779 | 365 |
} |
| 48674 | 366 |
|
| 68217 | 367 |
def join: (Process_Result, Option[String]) = |
| 50845 | 368 |
{
|
| 68217 | 369 |
val result0 = future_result.join |
370 |
val result1 = |
|
| 71601 | 371 |
export_consumer.shutdown(close = true).map(Output.error_message_text) match {
|
| 68217 | 372 |
case Nil => result0 |
| 68928 | 373 |
case errs => result0.errors(errs).error_rc |
|
68092
888d35a19866
store exports in session database, with asynchronous / parallel compression;
wenzelm
parents:
68088
diff
changeset
|
374 |
} |
| 50845 | 375 |
|
|
68198
6710167e17af
avoid race condition wrt. ISABELLE_TMP, which is removed in Bash.cleanup() before Bash.result(progress_stdout);
wenzelm
parents:
68148
diff
changeset
|
376 |
Isabelle_System.rm_tree(export_tmp_dir) |
|
6710167e17af
avoid race condition wrt. ISABELLE_TMP, which is removed in Bash.cleanup() before Bash.result(progress_stdout);
wenzelm
parents:
68148
diff
changeset
|
377 |
|
|
69811
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents:
69777
diff
changeset
|
378 |
if (result1.ok) |
| 62633 | 379 |
Present.finish(progress, store.browser_info, graph_file, info, name) |
| 61372 | 380 |
|
| 59445 | 381 |
graph_file.delete |
| 66943 | 382 |
|
383 |
val was_timeout = |
|
384 |
timeout_request match {
|
|
385 |
case None => false |
|
386 |
case Some(request) => !request.cancel |
|
387 |
} |
|
| 48674 | 388 |
|
| 68217 | 389 |
val result2 = |
390 |
if (result1.interrupted) {
|
|
391 |
if (was_timeout) result1.error(Output.error_message_text("Timeout")).was_timeout
|
|
392 |
else result1.error(Output.error_message_text("Interrupt"))
|
|
393 |
} |
|
394 |
else result1 |
|
395 |
||
396 |
val heap_digest = |
|
| 71614 | 397 |
if (result2.ok && do_store && store.output_heap(name).is_file) |
| 68217 | 398 |
Some(Sessions.write_heap_digest(store.output_heap(name))) |
399 |
else None |
|
400 |
||
401 |
(result2, heap_digest) |
|
| 48674 | 402 |
} |
| 48364 | 403 |
} |
404 |
||
| 48424 | 405 |
|
| 62631 | 406 |
|
| 62641 | 407 |
/** build with results **/ |
| 48424 | 408 |
|
| 63996 | 409 |
class Results private[Build](results: Map[String, (Option[Process_Result], Sessions.Info)]) |
| 62403 | 410 |
{
|
411 |
def sessions: Set[String] = results.keySet |
|
|
63082
6af03422535a
expose Sessions.Info in Build.Results
Lars Hupel <lars.hupel@mytum.de>
parents:
62946
diff
changeset
|
412 |
def cancelled(name: String): Boolean = results(name)._1.isEmpty |
|
6af03422535a
expose Sessions.Info in Build.Results
Lars Hupel <lars.hupel@mytum.de>
parents:
62946
diff
changeset
|
413 |
def apply(name: String): Process_Result = results(name)._1.getOrElse(Process_Result(1)) |
|
6af03422535a
expose Sessions.Info in Build.Results
Lars Hupel <lars.hupel@mytum.de>
parents:
62946
diff
changeset
|
414 |
def info(name: String): Sessions.Info = results(name)._2 |
| 71601 | 415 |
val rc: Int = |
| 65253 | 416 |
(0 /: results.iterator.map( |
417 |
{ case (_, (Some(r), _)) => r.rc case (_, (None, _)) => 1 }))(_ max _)
|
|
| 62641 | 418 |
def ok: Boolean = rc == 0 |
| 62406 | 419 |
|
420 |
override def toString: String = rc.toString |
|
| 62403 | 421 |
} |
422 |
||
| 62641 | 423 |
def build( |
|
50404
898cac1dad5e
avoid startup within GUI thread -- it is only required later for dialog;
wenzelm
parents:
50367
diff
changeset
|
424 |
options: Options, |
| 64909 | 425 |
progress: Progress = No_Progress, |
|
65832
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
wenzelm
parents:
65831
diff
changeset
|
426 |
check_unknown_files: Boolean = false, |
|
48511
37999ee01156
remove old output heaps, to ensure that result is valid wrt. check_stamps;
wenzelm
parents:
48509
diff
changeset
|
427 |
build_heap: Boolean = false, |
| 48595 | 428 |
clean_build: Boolean = false, |
| 56890 | 429 |
dirs: List[Path] = Nil, |
430 |
select_dirs: List[Path] = Nil, |
|
|
66968
9991671c98aa
allow to augment session context via explicit session infos;
wenzelm
parents:
66962
diff
changeset
|
431 |
infos: List[Sessions.Info] = Nil, |
| 64265 | 432 |
numa_shuffling: Boolean = false, |
| 48509 | 433 |
max_jobs: Int = 1, |
| 48903 | 434 |
list_files: Boolean = false, |
|
59891
9ce697050455
added isabelle build option -k, for fast off-line checking of theory sources;
wenzelm
parents:
59811
diff
changeset
|
435 |
check_keywords: Set[String] = Set.empty, |
| 66841 | 436 |
fresh_build: Boolean = false, |
| 48509 | 437 |
no_build: Boolean = false, |
| 66745 | 438 |
soft_build: Boolean = false, |
| 48509 | 439 |
verbose: Boolean = false, |
|
69811
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents:
69777
diff
changeset
|
440 |
export_files: Boolean = false, |
| 63224 | 441 |
requirements: Boolean = false, |
442 |
all_sessions: Boolean = false, |
|
|
66737
2edc0c42c883
option -B for "isabelle build" and "isabelle imports";
wenzelm
parents:
66736
diff
changeset
|
443 |
base_sessions: List[String] = Nil, |
| 63224 | 444 |
exclude_session_groups: List[String] = Nil, |
|
59892
2a616319c171
added isabelle build option -x, to exclude sessions;
wenzelm
parents:
59891
diff
changeset
|
445 |
exclude_sessions: List[String] = Nil, |
| 63224 | 446 |
session_groups: List[String] = Nil, |
| 65422 | 447 |
sessions: List[String] = Nil, |
448 |
selection: Sessions.Selection = Sessions.Selection.empty): Results = |
|
| 63224 | 449 |
{
|
| 66745 | 450 |
val build_options = options.int.update("completion_limit", 0).bool.update("ML_statistics", true)
|
| 51220 | 451 |
|
|
69854
cc0b3e177b49
system option "system_heaps" supersedes various command-line options for "system build mode";
wenzelm
parents:
69811
diff
changeset
|
452 |
val store = Sessions.store(build_options) |
| 66745 | 453 |
|
|
69369
6ecc85955e04
prefer "Isabelle DejaVu Sans", even for headless batch-build (session_graph.pdf);
wenzelm
parents:
68957
diff
changeset
|
454 |
Isabelle_Fonts.init() |
|
6ecc85955e04
prefer "Isabelle DejaVu Sans", even for headless batch-build (session_graph.pdf);
wenzelm
parents:
68957
diff
changeset
|
455 |
|
| 66745 | 456 |
|
457 |
/* session selection and dependencies */ |
|
| 65422 | 458 |
|
| 66961 | 459 |
val full_sessions = |
| 67026 | 460 |
Sessions.load_structure(build_options, dirs = dirs, select_dirs = select_dirs, infos = infos) |
| 65422 | 461 |
|
|
66749
0445cfaf6132
more compact (second-order) digest for 10^2..10^3 source files, with slightly increased risk of collisions;
wenzelm
parents:
66747
diff
changeset
|
462 |
def sources_stamp(deps: Sessions.Deps, name: String): String = |
|
0445cfaf6132
more compact (second-order) digest for 10^2..10^3 source files, with slightly increased risk of collisions;
wenzelm
parents:
66747
diff
changeset
|
463 |
{
|
|
0445cfaf6132
more compact (second-order) digest for 10^2..10^3 source files, with slightly increased risk of collisions;
wenzelm
parents:
66747
diff
changeset
|
464 |
val digests = |
|
0445cfaf6132
more compact (second-order) digest for 10^2..10^3 source files, with slightly increased risk of collisions;
wenzelm
parents:
66747
diff
changeset
|
465 |
full_sessions(name).meta_digest :: deps.sources(name) ::: deps.imported_sources(name) |
|
0445cfaf6132
more compact (second-order) digest for 10^2..10^3 source files, with slightly increased risk of collisions;
wenzelm
parents:
66747
diff
changeset
|
466 |
SHA1.digest(cat_lines(digests.map(_.toString).sorted)).toString |
|
0445cfaf6132
more compact (second-order) digest for 10^2..10^3 source files, with slightly increased risk of collisions;
wenzelm
parents:
66747
diff
changeset
|
467 |
} |
| 66745 | 468 |
|
|
67030
a9859e879f38
proper build_selection for clean_build (amending 961285f581e6): e.g. relevant for "isabelle build_doc";
wenzelm
parents:
67026
diff
changeset
|
469 |
val selection1 = |
|
a9859e879f38
proper build_selection for clean_build (amending 961285f581e6): e.g. relevant for "isabelle build_doc";
wenzelm
parents:
67026
diff
changeset
|
470 |
Sessions.Selection(requirements, all_sessions, base_sessions, exclude_session_groups, |
|
a9859e879f38
proper build_selection for clean_build (amending 961285f581e6): e.g. relevant for "isabelle build_doc";
wenzelm
parents:
67026
diff
changeset
|
471 |
exclude_sessions, session_groups, sessions) ++ selection |
|
a9859e879f38
proper build_selection for clean_build (amending 961285f581e6): e.g. relevant for "isabelle build_doc";
wenzelm
parents:
67026
diff
changeset
|
472 |
|
| 68204 | 473 |
val deps = |
| 66745 | 474 |
{
|
475 |
val deps0 = |
|
|
70671
cb1776c8e216
clarified signature: retain global session information, unaffected by later restriction;
wenzelm
parents:
70509
diff
changeset
|
476 |
Sessions.deps(full_sessions.selection(selection1), |
|
66962
e1bde71bace6
clarified signature: global_theories is always required;
wenzelm
parents:
66961
diff
changeset
|
477 |
progress = progress, inlined_files = true, verbose = verbose, |
|
e1bde71bace6
clarified signature: global_theories is always required;
wenzelm
parents:
66961
diff
changeset
|
478 |
list_files = list_files, check_keywords = check_keywords).check_errors |
| 66745 | 479 |
|
| 66841 | 480 |
if (soft_build && !fresh_build) {
|
| 66745 | 481 |
val outdated = |
| 68204 | 482 |
deps0.sessions_structure.build_topological_order.flatMap(name => |
| 68221 | 483 |
store.access_database(name) match {
|
| 68214 | 484 |
case Some(db) => |
| 68216 | 485 |
using(db)(store.read_build(_, name)) match {
|
| 66745 | 486 |
case Some(build) |
| 66747 | 487 |
if build.ok && build.sources == sources_stamp(deps0, name) => None |
| 66745 | 488 |
case _ => Some(name) |
489 |
} |
|
490 |
case None => Some(name) |
|
491 |
}) |
|
| 68204 | 492 |
|
493 |
Sessions.deps(full_sessions.selection(Sessions.Selection(sessions = outdated)), |
|
|
70671
cb1776c8e216
clarified signature: retain global session information, unaffected by later restriction;
wenzelm
parents:
70509
diff
changeset
|
494 |
progress = progress, inlined_files = true).check_errors |
| 66745 | 495 |
} |
| 68204 | 496 |
else deps0 |
| 66745 | 497 |
} |
498 |
||
499 |
||
500 |
/* check unknown files */ |
|
|
48504
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
wenzelm
parents:
48494
diff
changeset
|
501 |
|
|
65832
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
wenzelm
parents:
65831
diff
changeset
|
502 |
if (check_unknown_files) {
|
|
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
wenzelm
parents:
65831
diff
changeset
|
503 |
val source_files = |
|
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
wenzelm
parents:
65831
diff
changeset
|
504 |
(for {
|
|
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
wenzelm
parents:
65831
diff
changeset
|
505 |
(_, base) <- deps.session_bases.iterator |
|
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
wenzelm
parents:
65831
diff
changeset
|
506 |
(path, _) <- base.sources.iterator |
|
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
wenzelm
parents:
65831
diff
changeset
|
507 |
} yield path).toList |
| 67098 | 508 |
val exclude_files = List(Path.explode("$POLYML_EXE")).map(_.canonical_file)
|
509 |
val unknown_files = |
|
|
67782
7e223a05e6d8
clarified notion of unknown files: ignore files outside of a Mercurial repository;
wenzelm
parents:
67493
diff
changeset
|
510 |
Mercurial.check_files(source_files)._2. |
| 67098 | 511 |
filterNot(path => exclude_files.contains(path.canonical_file)) |
|
65832
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
wenzelm
parents:
65831
diff
changeset
|
512 |
if (unknown_files.nonEmpty) {
|
|
67782
7e223a05e6d8
clarified notion of unknown files: ignore files outside of a Mercurial repository;
wenzelm
parents:
67493
diff
changeset
|
513 |
progress.echo_warning("Unknown files (not part of the underlying Mercurial repository):" +
|
|
65832
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
wenzelm
parents:
65831
diff
changeset
|
514 |
unknown_files.map(path => path.expand.implode).sorted.mkString("\n ", "\n ", ""))
|
|
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
wenzelm
parents:
65831
diff
changeset
|
515 |
} |
|
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
wenzelm
parents:
65831
diff
changeset
|
516 |
} |
|
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
wenzelm
parents:
65831
diff
changeset
|
517 |
|
| 51220 | 518 |
|
519 |
/* main build process */ |
|
520 |
||
| 68204 | 521 |
val queue = Queue(progress, deps.sessions_structure, store) |
| 65289 | 522 |
|
| 68219 | 523 |
store.prepare_output_dir() |
| 48373 | 524 |
|
| 48595 | 525 |
if (clean_build) {
|
|
68734
c14a2cc9b5ef
isabelle build options -c -x -B refer to imports_graph;
wenzelm
parents:
68731
diff
changeset
|
526 |
for (name <- full_sessions.imports_descendants(full_sessions.imports_selection(selection1))) {
|
|
68220
8fc4e3d1df86
clarified store.clean_output: cleanup user_output_dir even in system_mode;
wenzelm
parents:
68219
diff
changeset
|
527 |
val (relevant, ok) = store.clean_output(name) |
|
8fc4e3d1df86
clarified store.clean_output: cleanup user_output_dir even in system_mode;
wenzelm
parents:
68219
diff
changeset
|
528 |
if (relevant) {
|
|
8fc4e3d1df86
clarified store.clean_output: cleanup user_output_dir even in system_mode;
wenzelm
parents:
68219
diff
changeset
|
529 |
if (ok) progress.echo("Cleaned " + name)
|
|
8fc4e3d1df86
clarified store.clean_output: cleanup user_output_dir even in system_mode;
wenzelm
parents:
68219
diff
changeset
|
530 |
else progress.echo(name + " FAILED to clean") |
|
8fc4e3d1df86
clarified store.clean_output: cleanup user_output_dir even in system_mode;
wenzelm
parents:
68219
diff
changeset
|
531 |
} |
| 48595 | 532 |
} |
533 |
} |
|
534 |
||
| 48425 | 535 |
// scheduler loop |
|
63082
6af03422535a
expose Sessions.Info in Build.Results
Lars Hupel <lars.hupel@mytum.de>
parents:
62946
diff
changeset
|
536 |
case class Result( |
| 66594 | 537 |
current: Boolean, |
| 68213 | 538 |
heap_digest: Option[String], |
| 66594 | 539 |
process: Option[Process_Result], |
540 |
info: Sessions.Info) |
|
| 62402 | 541 |
{
|
542 |
def ok: Boolean = |
|
543 |
process match {
|
|
544 |
case None => false |
|
545 |
case Some(res) => res.rc == 0 |
|
546 |
} |
|
547 |
} |
|
|
48639
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
wenzelm
parents:
48626
diff
changeset
|
548 |
|
|
56837
5a598f1eecfd
more robust interrupt handling for Scala_Console, which uses JVM Thread.interrupt instead of POSIX SIGINT;
wenzelm
parents:
56831
diff
changeset
|
549 |
def sleep() |
|
5a598f1eecfd
more robust interrupt handling for Scala_Console, which uses JVM Thread.interrupt instead of POSIX SIGINT;
wenzelm
parents:
56831
diff
changeset
|
550 |
{
|
|
5a598f1eecfd
more robust interrupt handling for Scala_Console, which uses JVM Thread.interrupt instead of POSIX SIGINT;
wenzelm
parents:
56831
diff
changeset
|
551 |
try { Thread.sleep(500) }
|
| 56861 | 552 |
catch { case Exn.Interrupt() => Exn.Interrupt.impose() }
|
|
56837
5a598f1eecfd
more robust interrupt handling for Scala_Console, which uses JVM Thread.interrupt instead of POSIX SIGINT;
wenzelm
parents:
56831
diff
changeset
|
553 |
} |
| 50366 | 554 |
|
| 64265 | 555 |
val numa_nodes = new NUMA.Nodes(numa_shuffling) |
556 |
||
| 48425 | 557 |
@tailrec def loop( |
|
48676
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
558 |
pending: Queue, |
| 62628 | 559 |
running: Map[String, (List[String], Job)], |
|
48639
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
wenzelm
parents:
48626
diff
changeset
|
560 |
results: Map[String, Result]): Map[String, Result] = |
| 48425 | 561 |
{
|
| 64265 | 562 |
def used_node(i: Int): Boolean = |
563 |
running.iterator.exists( |
|
564 |
{ case (_, (_, job)) => job.numa_node.isDefined && job.numa_node.get == i })
|
|
565 |
||
| 48425 | 566 |
if (pending.is_empty) results |
| 51253 | 567 |
else {
|
568 |
if (progress.stopped) |
|
569 |
for ((_, (_, job)) <- running) job.terminate |
|
570 |
||
| 48674 | 571 |
running.find({ case (_, (_, job)) => job.is_finished }) match {
|
| 62628 | 572 |
case Some((name, (input_heaps, job))) => |
| 50367 | 573 |
//{{{ finish job
|
| 48424 | 574 |
|
| 68217 | 575 |
val (process_result, heap_digest) = job.join |
| 48373 | 576 |
|
|
71623
b3bddebe44ca
clarified signature: more explicit type Protocol_Message.Marker;
wenzelm
parents:
71614
diff
changeset
|
577 |
val log_lines = process_result.out_lines.filterNot(Protocol_Message.Marker.test) |
|
62404
13a0f537e232
retain tail out_lines as printed, but not the whole log content;
wenzelm
parents:
62403
diff
changeset
|
578 |
val process_result_tail = |
|
13a0f537e232
retain tail out_lines as printed, but not the whole log content;
wenzelm
parents:
62403
diff
changeset
|
579 |
{
|
|
62409
e391528eff3b
proper option process_output_tail, more generous default;
wenzelm
parents:
62406
diff
changeset
|
580 |
val tail = job.info.options.int("process_output_tail")
|
| 62632 | 581 |
process_result.copy( |
582 |
out_lines = |
|
|
68086
9e1c670301b8
cleanup session output before starting build job;
wenzelm
parents:
67852
diff
changeset
|
583 |
"(see also " + store.output_log(name).file.toString + ")" :: |
| 65294 | 584 |
(if (tail == 0) log_lines else log_lines.drop(log_lines.length - tail max 0))) |
|
62404
13a0f537e232
retain tail out_lines as printed, but not the whole log content;
wenzelm
parents:
62403
diff
changeset
|
585 |
} |
|
13a0f537e232
retain tail out_lines as printed, but not the whole log content;
wenzelm
parents:
62403
diff
changeset
|
586 |
|
|
66666
1a620647285c
clarified messages: after writing all files (see also 27f90319a499 and 57c85c83c11b);
wenzelm
parents:
66595
diff
changeset
|
587 |
|
|
1a620647285c
clarified messages: after writing all files (see also 27f90319a499 and 57c85c83c11b);
wenzelm
parents:
66595
diff
changeset
|
588 |
// write log file |
| 68217 | 589 |
if (process_result.ok) {
|
590 |
File.write_gzip(store.output_log_gz(name), terminate_lines(log_lines)) |
|
591 |
} |
|
592 |
else File.write(store.output_log(name), terminate_lines(log_lines)) |
|
|
65291
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
593 |
|
|
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
594 |
// write database |
|
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
595 |
{
|
|
66944
05df740cb54b
more informative timeout message, notably for build_status;
wenzelm
parents:
66943
diff
changeset
|
596 |
val build_log = |
|
05df740cb54b
more informative timeout message, notably for build_status;
wenzelm
parents:
66943
diff
changeset
|
597 |
Build_Log.Log_File(name, process_result.out_lines). |
|
05df740cb54b
more informative timeout message, notably for build_status;
wenzelm
parents:
66943
diff
changeset
|
598 |
parse_session_info( |
|
05df740cb54b
more informative timeout message, notably for build_status;
wenzelm
parents:
66943
diff
changeset
|
599 |
command_timings = true, |
|
05df740cb54b
more informative timeout message, notably for build_status;
wenzelm
parents:
66943
diff
changeset
|
600 |
theory_timings = true, |
|
05df740cb54b
more informative timeout message, notably for build_status;
wenzelm
parents:
66943
diff
changeset
|
601 |
ml_statistics = true, |
|
05df740cb54b
more informative timeout message, notably for build_status;
wenzelm
parents:
66943
diff
changeset
|
602 |
task_statistics = true) |
|
05df740cb54b
more informative timeout message, notably for build_status;
wenzelm
parents:
66943
diff
changeset
|
603 |
|
| 68221 | 604 |
using(store.open_database(name, output = true))(db => |
|
65318
342efc382558
eliminated somewhat redundant inlined name (despite a7aa17a1f721);
wenzelm
parents:
65317
diff
changeset
|
605 |
store.write_session_info(db, name, |
|
65291
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
606 |
build_log = |
|
66944
05df740cb54b
more informative timeout message, notably for build_status;
wenzelm
parents:
66943
diff
changeset
|
607 |
if (process_result.timeout) build_log.error("Timeout") else build_log,
|
|
65291
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
608 |
build = |
| 68213 | 609 |
Session_Info(sources_stamp(deps, name), input_heaps, heap_digest, |
| 66744 | 610 |
process_result.rc))) |
|
65291
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
611 |
} |
|
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
wenzelm
parents:
65289
diff
changeset
|
612 |
|
|
66666
1a620647285c
clarified messages: after writing all files (see also 27f90319a499 and 57c85c83c11b);
wenzelm
parents:
66595
diff
changeset
|
613 |
// messages |
| 71601 | 614 |
process_result.err_lines.foreach(progress.echo) |
|
66666
1a620647285c
clarified messages: after writing all files (see also 27f90319a499 and 57c85c83c11b);
wenzelm
parents:
66595
diff
changeset
|
615 |
|
|
1a620647285c
clarified messages: after writing all files (see also 27f90319a499 and 57c85c83c11b);
wenzelm
parents:
66595
diff
changeset
|
616 |
if (process_result.ok) |
|
1a620647285c
clarified messages: after writing all files (see also 27f90319a499 and 57c85c83c11b);
wenzelm
parents:
66595
diff
changeset
|
617 |
progress.echo("Finished " + name + " (" + process_result.timing.message_resources + ")")
|
|
1a620647285c
clarified messages: after writing all files (see also 27f90319a499 and 57c85c83c11b);
wenzelm
parents:
66595
diff
changeset
|
618 |
else {
|
|
1a620647285c
clarified messages: after writing all files (see also 27f90319a499 and 57c85c83c11b);
wenzelm
parents:
66595
diff
changeset
|
619 |
progress.echo(name + " FAILED") |
|
1a620647285c
clarified messages: after writing all files (see also 27f90319a499 and 57c85c83c11b);
wenzelm
parents:
66595
diff
changeset
|
620 |
if (!process_result.interrupted) progress.echo(process_result_tail.out) |
|
1a620647285c
clarified messages: after writing all files (see also 27f90319a499 and 57c85c83c11b);
wenzelm
parents:
66595
diff
changeset
|
621 |
} |
|
1a620647285c
clarified messages: after writing all files (see also 27f90319a499 and 57c85c83c11b);
wenzelm
parents:
66595
diff
changeset
|
622 |
|
|
50707
5b2bf7611662
maintain session index on Scala side, for more determistic results;
wenzelm
parents:
50686
diff
changeset
|
623 |
loop(pending - name, running - name, |
| 68213 | 624 |
results + (name -> Result(false, heap_digest, Some(process_result_tail), job.info))) |
| 50367 | 625 |
//}}} |
| 60215 | 626 |
case None if running.size < (max_jobs max 1) => |
| 50367 | 627 |
//{{{ check/start next job
|
| 71601 | 628 |
pending.dequeue(running.isDefinedAt) match {
|
| 48547 | 629 |
case Some((name, info)) => |
| 66848 | 630 |
val ancestor_results = |
| 68204 | 631 |
deps.sessions_structure.build_requirements(List(name)).filterNot(_ == name). |
| 66848 | 632 |
map(results(_)) |
| 68213 | 633 |
val ancestor_heaps = ancestor_results.flatMap(_.heap_digest) |
| 62628 | 634 |
|
| 71614 | 635 |
val do_store = build_heap || Sessions.is_pure(name) || queue.is_inner(name) |
| 48547 | 636 |
|
| 68213 | 637 |
val (current, heap_digest) = |
| 48547 | 638 |
{
|
| 68221 | 639 |
store.access_database(name) match {
|
|
68212
5a59fded83c7
clarified heap vs. database operations: discontinued correlation of directory;
wenzelm
parents:
68209
diff
changeset
|
640 |
case Some(db) => |
| 68216 | 641 |
using(db)(store.read_build(_, name)) match {
|
642 |
case Some(build) => |
|
643 |
val heap_digest = store.find_heap_digest(name) |
|
644 |
val current = |
|
645 |
!fresh_build && |
|
646 |
build.ok && |
|
647 |
build.sources == sources_stamp(deps, name) && |
|
648 |
build.input_heaps == ancestor_heaps && |
|
649 |
build.output_heap == heap_digest && |
|
| 71614 | 650 |
!(do_store && heap_digest.isEmpty) |
| 68216 | 651 |
(current, heap_digest) |
652 |
case None => (false, None) |
|
653 |
} |
|
| 62636 | 654 |
case None => (false, None) |
|
48504
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
wenzelm
parents:
48494
diff
changeset
|
655 |
} |
| 48547 | 656 |
} |
| 62628 | 657 |
val all_current = current && ancestor_results.forall(_.current) |
|
48528
784c6f63d79c
proper all_current, which regards parent status as well;
wenzelm
parents:
48511
diff
changeset
|
658 |
|
| 48547 | 659 |
if (all_current) |
| 62402 | 660 |
loop(pending - name, running, |
| 68213 | 661 |
results + (name -> Result(true, heap_digest, Some(Process_Result(0)), info))) |
|
48678
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
662 |
else if (no_build) {
|
| 50366 | 663 |
if (verbose) progress.echo("Skipping " + name + " ...")
|
| 62402 | 664 |
loop(pending - name, running, |
| 68213 | 665 |
results + (name -> Result(false, heap_digest, Some(Process_Result(1)), info))) |
|
48678
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
666 |
} |
| 62628 | 667 |
else if (ancestor_results.forall(_.ok) && !progress.stopped) {
|
| 71614 | 668 |
progress.echo((if (do_store) "Building " else "Running ") + name + " ...") |
|
68086
9e1c670301b8
cleanup session output before starting build job;
wenzelm
parents:
67852
diff
changeset
|
669 |
|
|
68220
8fc4e3d1df86
clarified store.clean_output: cleanup user_output_dir even in system_mode;
wenzelm
parents:
68219
diff
changeset
|
670 |
store.clean_output(name) |
| 68221 | 671 |
using(store.open_database(name, output = true))(store.init_session_info(_, name)) |
|
68086
9e1c670301b8
cleanup session output before starting build job;
wenzelm
parents:
67852
diff
changeset
|
672 |
|
| 71601 | 673 |
val numa_node = numa_nodes.next(used_node) |
| 51220 | 674 |
val job = |
| 71651 | 675 |
new Job(progress, name, info, deps, store, do_store, verbose, |
|
69811
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents:
69777
diff
changeset
|
676 |
numa_node, queue.command_timings(name)) |
| 62628 | 677 |
loop(pending, running + (name -> (ancestor_heaps, job)), results) |
| 48547 | 678 |
} |
679 |
else {
|
|
| 50366 | 680 |
progress.echo(name + " CANCELLED") |
| 65253 | 681 |
loop(pending - name, running, |
| 68213 | 682 |
results + (name -> Result(false, heap_digest, None, info))) |
| 48547 | 683 |
} |
684 |
case None => sleep(); loop(pending, running, results) |
|
| 48425 | 685 |
} |
| 50367 | 686 |
///}}} |
| 48425 | 687 |
case None => sleep(); loop(pending, running, results) |
| 48373 | 688 |
} |
| 51253 | 689 |
} |
| 48425 | 690 |
} |
691 |
||
| 51220 | 692 |
|
693 |
/* build results */ |
|
694 |
||
| 62641 | 695 |
val results0 = |
| 48583 | 696 |
if (deps.is_empty) {
|
| 65827 | 697 |
progress.echo_warning("Nothing to build")
|
|
50707
5b2bf7611662
maintain session index on Scala side, for more determistic results;
wenzelm
parents:
50686
diff
changeset
|
698 |
Map.empty[String, Result] |
| 48583 | 699 |
} |
700 |
else loop(queue, Map.empty, Map.empty) |
|
701 |
||
| 62641 | 702 |
val results = |
| 64265 | 703 |
new Results( |
704 |
(for ((name, result) <- results0.iterator) |
|
705 |
yield (name, (result.process, result.info))).toMap) |
|
| 62641 | 706 |
|
|
69811
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents:
69777
diff
changeset
|
707 |
if (export_files) {
|
|
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents:
69777
diff
changeset
|
708 |
for (name <- full_sessions.imports_selection(selection1).iterator if results(name).ok) {
|
|
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents:
69777
diff
changeset
|
709 |
val info = results.info(name) |
|
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents:
69777
diff
changeset
|
710 |
if (info.export_files.nonEmpty) {
|
|
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents:
69777
diff
changeset
|
711 |
progress.echo("Exporting " + info.name + " ...")
|
|
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents:
69777
diff
changeset
|
712 |
for ((dir, prune, pats) <- info.export_files) {
|
|
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents:
69777
diff
changeset
|
713 |
Export.export_files(store, name, info.dir + dir, |
|
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents:
69777
diff
changeset
|
714 |
progress = if (verbose) progress else No_Progress, |
|
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents:
69777
diff
changeset
|
715 |
export_prune = prune, |
|
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents:
69777
diff
changeset
|
716 |
export_patterns = pats) |
|
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents:
69777
diff
changeset
|
717 |
} |
|
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents:
69777
diff
changeset
|
718 |
} |
|
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents:
69777
diff
changeset
|
719 |
} |
|
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents:
69777
diff
changeset
|
720 |
} |
|
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents:
69777
diff
changeset
|
721 |
|
| 62641 | 722 |
if (results.rc != 0 && (verbose || !no_build)) {
|
723 |
val unfinished = |
|
724 |
(for {
|
|
725 |
name <- results.sessions.iterator |
|
726 |
if !results(name).ok |
|
727 |
} yield name).toList.sorted |
|
728 |
progress.echo("Unfinished session(s): " + commas(unfinished))
|
|
729 |
} |
|
730 |
||
|
51418
7b8ce8403340
more accurate handling of global browser info at the very end (without races), subject to no_build and info.browser_info;
wenzelm
parents:
51402
diff
changeset
|
731 |
|
|
7b8ce8403340
more accurate handling of global browser info at the very end (without races), subject to no_build and info.browser_info;
wenzelm
parents:
51402
diff
changeset
|
732 |
/* global browser info */ |
|
7b8ce8403340
more accurate handling of global browser info at the very end (without races), subject to no_build and info.browser_info;
wenzelm
parents:
51402
diff
changeset
|
733 |
|
|
7b8ce8403340
more accurate handling of global browser info at the very end (without races), subject to no_build and info.browser_info;
wenzelm
parents:
51402
diff
changeset
|
734 |
if (!no_build) {
|
|
7b8ce8403340
more accurate handling of global browser info at the very end (without races), subject to no_build and info.browser_info;
wenzelm
parents:
51402
diff
changeset
|
735 |
val browser_chapters = |
|
7b8ce8403340
more accurate handling of global browser info at the very end (without races), subject to no_build and info.browser_info;
wenzelm
parents:
51402
diff
changeset
|
736 |
(for {
|
| 62641 | 737 |
(name, result) <- results0.iterator |
|
62944
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
738 |
if result.ok |
|
65415
8cd54b18b68b
clarified signature: tree structure is not essential;
wenzelm
parents:
65406
diff
changeset
|
739 |
info = full_sessions(name) |
|
51418
7b8ce8403340
more accurate handling of global browser info at the very end (without races), subject to no_build and info.browser_info;
wenzelm
parents:
51402
diff
changeset
|
740 |
if info.options.bool("browser_info")
|
|
7b8ce8403340
more accurate handling of global browser info at the very end (without races), subject to no_build and info.browser_info;
wenzelm
parents:
51402
diff
changeset
|
741 |
} yield (info.chapter, (name, info.description))).toList.groupBy(_._1). |
|
7b8ce8403340
more accurate handling of global browser info at the very end (without races), subject to no_build and info.browser_info;
wenzelm
parents:
51402
diff
changeset
|
742 |
map({ case (chapter, es) => (chapter, es.map(_._2)) }).filterNot(_._2.isEmpty)
|
|
7b8ce8403340
more accurate handling of global browser info at the very end (without races), subject to no_build and info.browser_info;
wenzelm
parents:
51402
diff
changeset
|
743 |
|
|
7b8ce8403340
more accurate handling of global browser info at the very end (without races), subject to no_build and info.browser_info;
wenzelm
parents:
51402
diff
changeset
|
744 |
for ((chapter, entries) <- browser_chapters) |
| 62632 | 745 |
Present.update_chapter_index(store.browser_info, chapter, entries) |
|
51418
7b8ce8403340
more accurate handling of global browser info at the very end (without races), subject to no_build and info.browser_info;
wenzelm
parents:
51402
diff
changeset
|
746 |
|
| 62632 | 747 |
if (browser_chapters.nonEmpty) Present.make_global_index(store.browser_info) |
|
51418
7b8ce8403340
more accurate handling of global browser info at the very end (without races), subject to no_build and info.browser_info;
wenzelm
parents:
51402
diff
changeset
|
748 |
} |
|
7b8ce8403340
more accurate handling of global browser info at the very end (without races), subject to no_build and info.browser_info;
wenzelm
parents:
51402
diff
changeset
|
749 |
|
| 62641 | 750 |
results |
| 48341 | 751 |
} |
752 |
||
753 |
||
| 62833 | 754 |
/* Isabelle tool wrapper */ |
| 48341 | 755 |
|
| 62833 | 756 |
val isabelle_tool = Isabelle_Tool("build", "build and manage Isabelle sessions", args =>
|
| 48341 | 757 |
{
|
| 62833 | 758 |
val build_options = Word.explode(Isabelle_System.getenv("ISABELLE_BUILD_OPTIONS"))
|
| 62590 | 759 |
|
|
66737
2edc0c42c883
option -B for "isabelle build" and "isabelle imports";
wenzelm
parents:
66736
diff
changeset
|
760 |
var base_sessions: List[String] = Nil |
| 62833 | 761 |
var select_dirs: List[Path] = Nil |
| 64265 | 762 |
var numa_shuffling = false |
| 62833 | 763 |
var requirements = false |
| 66745 | 764 |
var soft_build = false |
| 62833 | 765 |
var exclude_session_groups: List[String] = Nil |
766 |
var all_sessions = false |
|
767 |
var build_heap = false |
|
768 |
var clean_build = false |
|
769 |
var dirs: List[Path] = Nil |
|
|
69811
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents:
69777
diff
changeset
|
770 |
var export_files = false |
| 66841 | 771 |
var fresh_build = false |
| 62833 | 772 |
var session_groups: List[String] = Nil |
773 |
var max_jobs = 1 |
|
774 |
var check_keywords: Set[String] = Set.empty |
|
775 |
var list_files = false |
|
776 |
var no_build = false |
|
| 67847 | 777 |
var options = Options.init(opts = build_options) |
| 62833 | 778 |
var verbose = false |
779 |
var exclude_sessions: List[String] = Nil |
|
| 62590 | 780 |
|
| 62833 | 781 |
val getopts = Getopts("""
|
| 62590 | 782 |
Usage: isabelle build [OPTIONS] [SESSIONS ...] |
783 |
||
784 |
Options are: |
|
|
66737
2edc0c42c883
option -B for "isabelle build" and "isabelle imports";
wenzelm
parents:
66736
diff
changeset
|
785 |
-B NAME include session NAME and all descendants |
| 62590 | 786 |
-D DIR include session directory and select its sessions |
| 64265 | 787 |
-N cyclic shuffling of NUMA CPU nodes (performance tuning) |
| 62590 | 788 |
-R operate on requirements of selected sessions |
| 66745 | 789 |
-S soft build: only observe changes of sources, not heap images |
| 62590 | 790 |
-X NAME exclude sessions from group NAME and all descendants |
791 |
-a select all sessions |
|
792 |
-b build heap images |
|
793 |
-c clean build |
|
794 |
-d DIR include session directory |
|
|
69811
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents:
69777
diff
changeset
|
795 |
-e export files from session specification into file-system |
| 66841 | 796 |
-f fresh build |
| 62590 | 797 |
-g NAME select session group NAME |
798 |
-j INT maximum number of parallel jobs (default 1) |
|
799 |
-k KEYWORD check theory sources for conflicts with proposed keywords |
|
800 |
-l list session source files |
|
801 |
-n no build -- test dependencies only |
|
802 |
-o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) |
|
803 |
-v verbose |
|
804 |
-x NAME exclude session NAME and all descendants |
|
805 |
||
| 62596 | 806 |
Build and manage Isabelle sessions, depending on implicit settings: |
807 |
||
| 64455 | 808 |
""" + Library.prefix_lines(" ", Build_Log.Settings.show()) + "\n",
|
|
66737
2edc0c42c883
option -B for "isabelle build" and "isabelle imports";
wenzelm
parents:
66736
diff
changeset
|
809 |
"B:" -> (arg => base_sessions = base_sessions ::: List(arg)), |
| 62833 | 810 |
"D:" -> (arg => select_dirs = select_dirs ::: List(Path.explode(arg))), |
| 64265 | 811 |
"N" -> (_ => numa_shuffling = true), |
| 62833 | 812 |
"R" -> (_ => requirements = true), |
| 66745 | 813 |
"S" -> (_ => soft_build = true), |
| 62833 | 814 |
"X:" -> (arg => exclude_session_groups = exclude_session_groups ::: List(arg)), |
815 |
"a" -> (_ => all_sessions = true), |
|
816 |
"b" -> (_ => build_heap = true), |
|
817 |
"c" -> (_ => clean_build = true), |
|
818 |
"d:" -> (arg => dirs = dirs ::: List(Path.explode(arg))), |
|
|
69811
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents:
69777
diff
changeset
|
819 |
"e" -> (_ => export_files = true), |
| 66841 | 820 |
"f" -> (_ => fresh_build = true), |
| 62833 | 821 |
"g:" -> (arg => session_groups = session_groups ::: List(arg)), |
| 63805 | 822 |
"j:" -> (arg => max_jobs = Value.Int.parse(arg)), |
| 62833 | 823 |
"k:" -> (arg => check_keywords = check_keywords + arg), |
824 |
"l" -> (_ => list_files = true), |
|
825 |
"n" -> (_ => no_build = true), |
|
826 |
"o:" -> (arg => options = options + arg), |
|
827 |
"v" -> (_ => verbose = true), |
|
828 |
"x:" -> (arg => exclude_sessions = exclude_sessions ::: List(arg))) |
|
| 62590 | 829 |
|
| 62833 | 830 |
val sessions = getopts(args) |
| 62590 | 831 |
|
| 64115 | 832 |
val progress = new Console_Progress(verbose = verbose) |
| 62590 | 833 |
|
| 64140 | 834 |
val start_date = Date.now() |
835 |
||
| 62833 | 836 |
if (verbose) {
|
837 |
progress.echo( |
|
| 64155 | 838 |
"Started at " + Build_Log.print_date(start_date) + |
| 64140 | 839 |
" (" + Isabelle_System.getenv("ML_IDENTIFIER") + " on " + Isabelle_System.hostname() +")")
|
| 64081 | 840 |
progress.echo(Build_Log.Settings.show() + "\n") |
| 62833 | 841 |
} |
| 62590 | 842 |
|
| 62833 | 843 |
val results = |
844 |
progress.interrupt_handler {
|
|
| 67846 | 845 |
build(options, |
846 |
progress = progress, |
|
|
65832
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
wenzelm
parents:
65831
diff
changeset
|
847 |
check_unknown_files = Mercurial.is_repository(Path.explode("~~")),
|
| 63224 | 848 |
build_heap = build_heap, |
849 |
clean_build = clean_build, |
|
850 |
dirs = dirs, |
|
851 |
select_dirs = select_dirs, |
|
| 65831 | 852 |
numa_shuffling = NUMA.enabled_warning(progress, numa_shuffling), |
| 63224 | 853 |
max_jobs = max_jobs, |
854 |
list_files = list_files, |
|
855 |
check_keywords = check_keywords, |
|
| 66841 | 856 |
fresh_build = fresh_build, |
| 63224 | 857 |
no_build = no_build, |
| 66745 | 858 |
soft_build = soft_build, |
| 63224 | 859 |
verbose = verbose, |
|
69811
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents:
69777
diff
changeset
|
860 |
export_files = export_files, |
| 63224 | 861 |
requirements = requirements, |
862 |
all_sessions = all_sessions, |
|
|
66737
2edc0c42c883
option -B for "isabelle build" and "isabelle imports";
wenzelm
parents:
66736
diff
changeset
|
863 |
base_sessions = base_sessions, |
| 63224 | 864 |
exclude_session_groups = exclude_session_groups, |
865 |
exclude_sessions = exclude_sessions, |
|
866 |
session_groups = session_groups, |
|
867 |
sessions = sessions) |
|
| 62833 | 868 |
} |
| 64140 | 869 |
val end_date = Date.now() |
870 |
val elapsed_time = end_date.time - start_date.time |
|
| 62590 | 871 |
|
| 62833 | 872 |
if (verbose) {
|
| 64155 | 873 |
progress.echo("\nFinished at " + Build_Log.print_date(end_date))
|
| 62833 | 874 |
} |
| 62590 | 875 |
|
| 62833 | 876 |
val total_timing = |
877 |
(Timing.zero /: results.sessions.iterator.map(a => results(a).timing))(_ + _). |
|
878 |
copy(elapsed = elapsed_time) |
|
879 |
progress.echo(total_timing.message_resources) |
|
| 62590 | 880 |
|
| 62833 | 881 |
sys.exit(results.rc) |
882 |
}) |
|
| 68305 | 883 |
|
884 |
||
885 |
/* build logic image */ |
|
886 |
||
887 |
def build_logic(options: Options, logic: String, |
|
888 |
progress: Progress = No_Progress, |
|
889 |
build_heap: Boolean = false, |
|
890 |
dirs: List[Path] = Nil, |
|
| 70791 | 891 |
fresh: Boolean = false, |
| 69540 | 892 |
strict: Boolean = false): Int = |
| 68305 | 893 |
{
|
| 69540 | 894 |
val rc = |
| 70791 | 895 |
if (!fresh && build(options, build_heap = build_heap, no_build = true, dirs = dirs, |
|
69854
cc0b3e177b49
system option "system_heaps" supersedes various command-line options for "system build mode";
wenzelm
parents:
69811
diff
changeset
|
896 |
sessions = List(logic)).ok) 0 |
| 69540 | 897 |
else {
|
898 |
progress.echo("Build started for Isabelle/" + logic + " ...")
|
|
| 70791 | 899 |
Build.build(options, progress = progress, build_heap = build_heap, fresh_build = fresh, |
900 |
dirs = dirs, sessions = List(logic)).rc |
|
| 69540 | 901 |
} |
902 |
||
903 |
if (strict && rc != 0) error("Failed to build Isabelle/" + logic) else rc
|
|
| 68305 | 904 |
} |
| 48276 | 905 |
} |