| author | wenzelm |
| Tue, 14 Mar 2017 11:16:23 +0100 | |
| changeset 65223 | 844c067bc3d4 |
| parent 65220 | 420f55912b3e |
| child 65250 | 13a6c81534a8 |
| 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 |
||
| 48548 | 11 |
import java.io.{BufferedInputStream, FileInputStream,
|
| 48494 | 12 |
BufferedReader, InputStreamReader, IOException} |
13 |
import java.util.zip.GZIPInputStream |
|
| 48335 | 14 |
|
|
48676
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
15 |
import scala.collection.SortedSet |
|
51397
03b586ee5930
support for 'chapter' specifications within session ROOT;
wenzelm
parents:
51300
diff
changeset
|
16 |
import scala.collection.mutable |
|
48340
6f4fc030882a
allow explicit specification of additional session directories;
wenzelm
parents:
48339
diff
changeset
|
17 |
import scala.annotation.tailrec |
|
48337
9c7f8e5805b4
cumulate semantic Session_Info, based on syntactic Session_Entry;
wenzelm
parents:
48336
diff
changeset
|
18 |
|
| 48335 | 19 |
|
| 48276 | 20 |
object Build |
21 |
{
|
|
| 62631 | 22 |
/** auxiliary **/ |
| 48424 | 23 |
|
|
48676
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
24 |
/* queue */ |
|
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
25 |
|
| 62631 | 26 |
private object Queue |
|
48676
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
27 |
{
|
| 62631 | 28 |
def apply(tree: Sessions.Tree, load_timings: String => (List[Properties.T], Double)): Queue = |
|
48676
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
29 |
{
|
|
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
30 |
val graph = tree.graph |
|
56372
fadb0fef09d7
more explicit iterator terminology, in accordance to Scala 2.8 library;
wenzelm
parents:
56208
diff
changeset
|
31 |
val sessions = graph.keys |
| 51220 | 32 |
|
|
59136
c2b23cb8a677
added Par_List in Scala, in accordance to ML version;
wenzelm
parents:
59083
diff
changeset
|
33 |
val timings = Par_List.map((name: String) => (name, load_timings(name)), sessions) |
| 51220 | 34 |
val command_timings = |
35 |
Map(timings.map({ case (name, (ts, _)) => (name, ts) }): _*).withDefaultValue(Nil)
|
|
36 |
val session_timing = |
|
37 |
Map(timings.map({ case (name, (_, t)) => (name, t) }): _*).withDefaultValue(0.0)
|
|
|
48678
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
38 |
|
|
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
39 |
def outdegree(name: String): Int = graph.imm_succs(name).size |
|
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
40 |
|
|
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
41 |
object Ordering extends scala.math.Ordering[String] |
|
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
42 |
{
|
|
51227
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
43 |
def compare_timing(name1: String, name2: String): Int = |
|
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
44 |
{
|
|
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
45 |
val t1 = session_timing(name1) |
|
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
46 |
val t2 = session_timing(name2) |
|
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
47 |
if (t1 == 0.0 || t2 == 0.0) 0 |
|
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
48 |
else t1 compare t2 |
|
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
49 |
} |
|
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
50 |
|
|
48678
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
51 |
def compare(name1: String, name2: String): Int = |
|
51229
6e40d0bb89e3
prefer outdegree in comparison again (cf. 88c96e836ed6) -- NB: big jobs might hide behind small ones in this naive queuing scheme;
wenzelm
parents:
51227
diff
changeset
|
52 |
outdegree(name2) compare outdegree(name1) match {
|
|
48678
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
53 |
case 0 => |
|
51229
6e40d0bb89e3
prefer outdegree in comparison again (cf. 88c96e836ed6) -- NB: big jobs might hide behind small ones in this naive queuing scheme;
wenzelm
parents:
51227
diff
changeset
|
54 |
compare_timing(name2, name1) match {
|
| 51220 | 55 |
case 0 => |
| 61602 | 56 |
tree(name2).timeout compare tree(name1).timeout match {
|
| 51220 | 57 |
case 0 => name1 compare name2 |
58 |
case ord => ord |
|
59 |
} |
|
|
48678
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
60 |
case ord => ord |
|
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
61 |
} |
|
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
62 |
case ord => ord |
|
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
63 |
} |
|
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
64 |
} |
|
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
65 |
|
| 51220 | 66 |
new Queue(graph, SortedSet(sessions: _*)(Ordering), command_timings) |
|
48676
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
67 |
} |
|
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
68 |
} |
|
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
69 |
|
| 62631 | 70 |
private final class Queue private( |
71 |
graph: Graph[String, Sessions.Info], |
|
| 51220 | 72 |
order: SortedSet[String], |
73 |
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
|
74 |
{
|
|
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
75 |
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
|
76 |
|
|
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
77 |
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
|
78 |
|
|
51227
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
79 |
def - (name: String): Queue = |
|
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
80 |
new Queue(graph.del_node(name), |
|
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
81 |
order - name, // FIXME scala-2.10.0 TreeSet problem!? |
|
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
82 |
command_timings) |
|
48676
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
83 |
|
| 62631 | 84 |
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
|
85 |
{
|
|
51227
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
86 |
val it = order.iterator.dropWhile(name => |
|
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
87 |
skip(name) |
|
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
88 |
|| !graph.defined(name) // FIXME scala-2.10.0 TreeSet problem!? |
|
88c96e836ed6
prefer comparison of session timing, if this is known already;
wenzelm
parents:
51223
diff
changeset
|
89 |
|| !graph.is_minimal(name)) |
| 48680 | 90 |
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
|
91 |
else None |
|
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
92 |
} |
|
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
93 |
} |
|
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
94 |
|
|
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
95 |
|
|
48710
5b51ccdc8623
prefer static Build.session_content for loaded theories -- discontinued incremental protocol;
wenzelm
parents:
48708
diff
changeset
|
96 |
/* source dependencies and static content */ |
|
48422
9613780a805b
determine source dependencies, relatively to preloaded theories;
wenzelm
parents:
48421
diff
changeset
|
97 |
|
| 64856 | 98 |
sealed case class Deps(deps: Map[String, Sessions.Base]) |
|
48423
0ccf143a2a69
maintain set of source digests, including relevant parts of session entry;
wenzelm
parents:
48422
diff
changeset
|
99 |
{
|
| 48583 | 100 |
def is_empty: Boolean = deps.isEmpty |
| 64856 | 101 |
def apply(name: String): Sessions.Base = deps(name) |
|
48504
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
wenzelm
parents:
48494
diff
changeset
|
102 |
def sources(name: String): List[SHA1.Digest] = deps(name).sources.map(_._2) |
|
48423
0ccf143a2a69
maintain set of source digests, including relevant parts of session entry;
wenzelm
parents:
48422
diff
changeset
|
103 |
} |
|
0ccf143a2a69
maintain set of source digests, including relevant parts of session entry;
wenzelm
parents:
48422
diff
changeset
|
104 |
|
|
59891
9ce697050455
added isabelle build option -k, for fast off-line checking of theory sources;
wenzelm
parents:
59811
diff
changeset
|
105 |
def dependencies( |
| 64909 | 106 |
progress: Progress = No_Progress, |
|
59891
9ce697050455
added isabelle build option -k, for fast off-line checking of theory sources;
wenzelm
parents:
59811
diff
changeset
|
107 |
inlined_files: Boolean = false, |
|
9ce697050455
added isabelle build option -k, for fast off-line checking of theory sources;
wenzelm
parents:
59811
diff
changeset
|
108 |
verbose: Boolean = false, |
|
9ce697050455
added isabelle build option -k, for fast off-line checking of theory sources;
wenzelm
parents:
59811
diff
changeset
|
109 |
list_files: Boolean = false, |
|
9ce697050455
added isabelle build option -k, for fast off-line checking of theory sources;
wenzelm
parents:
59811
diff
changeset
|
110 |
check_keywords: Set[String] = Set.empty, |
| 62631 | 111 |
tree: Sessions.Tree): Deps = |
| 64856 | 112 |
Deps((Map.empty[String, Sessions.Base] /: tree.topological_order)( |
|
48422
9613780a805b
determine source dependencies, relatively to preloaded theories;
wenzelm
parents:
48421
diff
changeset
|
113 |
{ case (deps, (name, info)) =>
|
|
56668
56335a8e2e8b
interruptible dependencies, which can take a few seconds;
wenzelm
parents:
56667
diff
changeset
|
114 |
if (progress.stopped) throw Exn.Interrupt() |
|
56335a8e2e8b
interruptible dependencies, which can take a few seconds;
wenzelm
parents:
56667
diff
changeset
|
115 |
|
|
54549
2a3053472ec3
actually expose errors of cumulative theory dependencies;
wenzelm
parents:
54514
diff
changeset
|
116 |
try {
|
| 64856 | 117 |
val resources = |
118 |
new Resources( |
|
119 |
info.parent match {
|
|
120 |
case None => Sessions.Base.bootstrap |
|
121 |
case Some(parent) => deps(parent) |
|
122 |
}) |
|
|
48422
9613780a805b
determine source dependencies, relatively to preloaded theories;
wenzelm
parents:
48421
diff
changeset
|
123 |
|
|
54549
2a3053472ec3
actually expose errors of cumulative theory dependencies;
wenzelm
parents:
54514
diff
changeset
|
124 |
if (verbose || list_files) {
|
|
2a3053472ec3
actually expose errors of cumulative theory dependencies;
wenzelm
parents:
54514
diff
changeset
|
125 |
val groups = |
|
2a3053472ec3
actually expose errors of cumulative theory dependencies;
wenzelm
parents:
54514
diff
changeset
|
126 |
if (info.groups.isEmpty) "" |
|
2a3053472ec3
actually expose errors of cumulative theory dependencies;
wenzelm
parents:
54514
diff
changeset
|
127 |
else info.groups.mkString(" (", " ", ")")
|
|
2a3053472ec3
actually expose errors of cumulative theory dependencies;
wenzelm
parents:
54514
diff
changeset
|
128 |
progress.echo("Session " + info.chapter + "/" + name + groups)
|
|
2a3053472ec3
actually expose errors of cumulative theory dependencies;
wenzelm
parents:
54514
diff
changeset
|
129 |
} |
| 48478 | 130 |
|
|
54549
2a3053472ec3
actually expose errors of cumulative theory dependencies;
wenzelm
parents:
54514
diff
changeset
|
131 |
val thy_deps = |
|
56801
8dd9df88f647
some support for session-qualified theories: allow to refer to resources via qualified name instead of odd file-system path;
wenzelm
parents:
56782
diff
changeset
|
132 |
{
|
|
8dd9df88f647
some support for session-qualified theories: allow to refer to resources via qualified name instead of odd file-system path;
wenzelm
parents:
56782
diff
changeset
|
133 |
val root_theories = |
|
8dd9df88f647
some support for session-qualified theories: allow to refer to resources via qualified name instead of odd file-system path;
wenzelm
parents:
56782
diff
changeset
|
134 |
info.theories.flatMap({
|
|
8dd9df88f647
some support for session-qualified theories: allow to refer to resources via qualified name instead of odd file-system path;
wenzelm
parents:
56782
diff
changeset
|
135 |
case (global, _, thys) => |
|
8dd9df88f647
some support for session-qualified theories: allow to refer to resources via qualified name instead of odd file-system path;
wenzelm
parents:
56782
diff
changeset
|
136 |
thys.map(thy => |
|
8dd9df88f647
some support for session-qualified theories: allow to refer to resources via qualified name instead of odd file-system path;
wenzelm
parents:
56782
diff
changeset
|
137 |
(resources.node_name( |
|
8dd9df88f647
some support for session-qualified theories: allow to refer to resources via qualified name instead of odd file-system path;
wenzelm
parents:
56782
diff
changeset
|
138 |
if (global) "" else name, info.dir + Resources.thy_path(thy)), info.pos)) |
|
8dd9df88f647
some support for session-qualified theories: allow to refer to resources via qualified name instead of odd file-system path;
wenzelm
parents:
56782
diff
changeset
|
139 |
}) |
| 64839 | 140 |
val thy_deps = resources.thy_info.dependencies(name, root_theories) |
|
48422
9613780a805b
determine source dependencies, relatively to preloaded theories;
wenzelm
parents:
48421
diff
changeset
|
141 |
|
|
56801
8dd9df88f647
some support for session-qualified theories: allow to refer to resources via qualified name instead of odd file-system path;
wenzelm
parents:
56782
diff
changeset
|
142 |
thy_deps.errors match {
|
|
8dd9df88f647
some support for session-qualified theories: allow to refer to resources via qualified name instead of odd file-system path;
wenzelm
parents:
56782
diff
changeset
|
143 |
case Nil => thy_deps |
|
8dd9df88f647
some support for session-qualified theories: allow to refer to resources via qualified name instead of odd file-system path;
wenzelm
parents:
56782
diff
changeset
|
144 |
case errs => error(cat_lines(errs)) |
|
8dd9df88f647
some support for session-qualified theories: allow to refer to resources via qualified name instead of odd file-system path;
wenzelm
parents:
56782
diff
changeset
|
145 |
} |
|
54549
2a3053472ec3
actually expose errors of cumulative theory dependencies;
wenzelm
parents:
54514
diff
changeset
|
146 |
} |
|
51294
0850d43cb355
discontinued obsolete header "files" -- these are loaded explicitly after exploring dependencies;
wenzelm
parents:
51293
diff
changeset
|
147 |
|
|
56801
8dd9df88f647
some support for session-qualified theories: allow to refer to resources via qualified name instead of odd file-system path;
wenzelm
parents:
56782
diff
changeset
|
148 |
val known_theories = |
| 64856 | 149 |
(resources.base.known_theories /: thy_deps.deps)({ case (known, dep) =>
|
|
56801
8dd9df88f647
some support for session-qualified theories: allow to refer to resources via qualified name instead of odd file-system path;
wenzelm
parents:
56782
diff
changeset
|
150 |
val name = dep.name |
|
8dd9df88f647
some support for session-qualified theories: allow to refer to resources via qualified name instead of odd file-system path;
wenzelm
parents:
56782
diff
changeset
|
151 |
known.get(name.theory) match {
|
|
8dd9df88f647
some support for session-qualified theories: allow to refer to resources via qualified name instead of odd file-system path;
wenzelm
parents:
56782
diff
changeset
|
152 |
case Some(name1) if name != name1 => |
|
8dd9df88f647
some support for session-qualified theories: allow to refer to resources via qualified name instead of odd file-system path;
wenzelm
parents:
56782
diff
changeset
|
153 |
error("Duplicate theory " + quote(name.node) + " vs. " + quote(name1.node))
|
|
8dd9df88f647
some support for session-qualified theories: allow to refer to resources via qualified name instead of odd file-system path;
wenzelm
parents:
56782
diff
changeset
|
154 |
case _ => |
|
8dd9df88f647
some support for session-qualified theories: allow to refer to resources via qualified name instead of odd file-system path;
wenzelm
parents:
56782
diff
changeset
|
155 |
known + (name.theory -> name) + (Long_Name.base_name(name.theory) -> name) |
|
8dd9df88f647
some support for session-qualified theories: allow to refer to resources via qualified name instead of odd file-system path;
wenzelm
parents:
56782
diff
changeset
|
156 |
} |
|
8dd9df88f647
some support for session-qualified theories: allow to refer to resources via qualified name instead of odd file-system path;
wenzelm
parents:
56782
diff
changeset
|
157 |
}) |
|
8dd9df88f647
some support for session-qualified theories: allow to refer to resources via qualified name instead of odd file-system path;
wenzelm
parents:
56782
diff
changeset
|
158 |
|
|
54549
2a3053472ec3
actually expose errors of cumulative theory dependencies;
wenzelm
parents:
54514
diff
changeset
|
159 |
val loaded_theories = thy_deps.loaded_theories |
|
2a3053472ec3
actually expose errors of cumulative theory dependencies;
wenzelm
parents:
54514
diff
changeset
|
160 |
val keywords = thy_deps.keywords |
| 63865 | 161 |
val syntax = thy_deps.syntax |
|
48660
730ca503e955
static outer syntax based on session specifications;
wenzelm
parents:
48650
diff
changeset
|
162 |
|
|
59891
9ce697050455
added isabelle build option -k, for fast off-line checking of theory sources;
wenzelm
parents:
59811
diff
changeset
|
163 |
val theory_files = thy_deps.deps.map(dep => Path.explode(dep.name.node)) |
|
62866
d20262cd20e8
read Pure file dependencies directly from ROOT.ML;
wenzelm
parents:
62840
diff
changeset
|
164 |
val loaded_files = |
|
62902
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
wenzelm
parents:
62883
diff
changeset
|
165 |
if (inlined_files) {
|
|
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
wenzelm
parents:
62883
diff
changeset
|
166 |
val pure_files = |
|
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
wenzelm
parents:
62883
diff
changeset
|
167 |
if (Sessions.pure_name(name)) Sessions.pure_files(resources, syntax, info.dir) |
|
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
wenzelm
parents:
62883
diff
changeset
|
168 |
else Nil |
|
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
wenzelm
parents:
62883
diff
changeset
|
169 |
pure_files ::: thy_deps.loaded_files |
|
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
wenzelm
parents:
62883
diff
changeset
|
170 |
} |
|
62866
d20262cd20e8
read Pure file dependencies directly from ROOT.ML;
wenzelm
parents:
62840
diff
changeset
|
171 |
else Nil |
| 48903 | 172 |
|
|
54549
2a3053472ec3
actually expose errors of cumulative theory dependencies;
wenzelm
parents:
54514
diff
changeset
|
173 |
val all_files = |
|
59891
9ce697050455
added isabelle build option -k, for fast off-line checking of theory sources;
wenzelm
parents:
59811
diff
changeset
|
174 |
(theory_files ::: loaded_files ::: |
|
56533
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
wenzelm
parents:
56464
diff
changeset
|
175 |
info.files.map(file => info.dir + file) ::: |
|
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
wenzelm
parents:
56464
diff
changeset
|
176 |
info.document_files.map(file => info.dir + file._1 + file._2)).map(_.expand) |
|
54549
2a3053472ec3
actually expose errors of cumulative theory dependencies;
wenzelm
parents:
54514
diff
changeset
|
177 |
|
|
56824
5ae68f53b7c2
discontinued adhoc check (see also ea8343187225);
wenzelm
parents:
56801
diff
changeset
|
178 |
if (list_files) |
|
54549
2a3053472ec3
actually expose errors of cumulative theory dependencies;
wenzelm
parents:
54514
diff
changeset
|
179 |
progress.echo(cat_lines(all_files.map(_.implode).sorted.map(" " + _)))
|
|
2a3053472ec3
actually expose errors of cumulative theory dependencies;
wenzelm
parents:
54514
diff
changeset
|
180 |
|
| 59895 | 181 |
if (check_keywords.nonEmpty) |
182 |
Check_Keywords.check_keywords(progress, syntax.keywords, check_keywords, theory_files) |
|
|
59891
9ce697050455
added isabelle build option -k, for fast off-line checking of theory sources;
wenzelm
parents:
59811
diff
changeset
|
183 |
|
|
54549
2a3053472ec3
actually expose errors of cumulative theory dependencies;
wenzelm
parents:
54514
diff
changeset
|
184 |
val sources = all_files.map(p => (p, SHA1.digest(p.file))) |
| 48903 | 185 |
|
| 60077 | 186 |
val session_graph = |
| 64856 | 187 |
Present.session_graph(info.parent getOrElse "", |
188 |
resources.base.loaded_theories, thy_deps.deps) |
|
| 59444 | 189 |
|
| 64856 | 190 |
val base = |
191 |
Sessions.Base( |
|
192 |
loaded_theories, known_theories, keywords, syntax, sources, session_graph) |
|
193 |
deps + (name -> base) |
|
|
54549
2a3053472ec3
actually expose errors of cumulative theory dependencies;
wenzelm
parents:
54514
diff
changeset
|
194 |
} |
|
2a3053472ec3
actually expose errors of cumulative theory dependencies;
wenzelm
parents:
54514
diff
changeset
|
195 |
catch {
|
|
2a3053472ec3
actually expose errors of cumulative theory dependencies;
wenzelm
parents:
54514
diff
changeset
|
196 |
case ERROR(msg) => |
|
2a3053472ec3
actually expose errors of cumulative theory dependencies;
wenzelm
parents:
54514
diff
changeset
|
197 |
cat_error(msg, "The error(s) above occurred in session " + |
|
2a3053472ec3
actually expose errors of cumulative theory dependencies;
wenzelm
parents:
54514
diff
changeset
|
198 |
quote(name) + Position.here(info.pos)) |
|
2a3053472ec3
actually expose errors of cumulative theory dependencies;
wenzelm
parents:
54514
diff
changeset
|
199 |
} |
|
48423
0ccf143a2a69
maintain set of source digests, including relevant parts of session entry;
wenzelm
parents:
48422
diff
changeset
|
200 |
})) |
|
48422
9613780a805b
determine source dependencies, relatively to preloaded theories;
wenzelm
parents:
48421
diff
changeset
|
201 |
|
| 65211 | 202 |
def session_base(options: Options, session: String, dirs: List[Path] = Nil): Sessions.Base = |
|
48710
5b51ccdc8623
prefer static Build.session_content for loaded theories -- discontinued incremental protocol;
wenzelm
parents:
48708
diff
changeset
|
203 |
{
|
| 65211 | 204 |
val (_, tree) = Sessions.load(options, dirs = dirs).selection(sessions = List(session)) |
205 |
dependencies(tree = tree)(session) |
|
|
52439
4cf3f6153eb8
improved "isabelle keywords" and "isabelle update_keywords" based on Isabelle/Scala, without requiring to build sessions first;
wenzelm
parents:
52115
diff
changeset
|
206 |
} |
|
4cf3f6153eb8
improved "isabelle keywords" and "isabelle update_keywords" based on Isabelle/Scala, without requiring to build sessions first;
wenzelm
parents:
52115
diff
changeset
|
207 |
|
|
48422
9613780a805b
determine source dependencies, relatively to preloaded theories;
wenzelm
parents:
48421
diff
changeset
|
208 |
|
| 48424 | 209 |
/* jobs */ |
| 48341 | 210 |
|
|
62637
0189fe0f6452
support for Poly/ML heap hierarchy, which saves a lot of disk space;
wenzelm
parents:
62636
diff
changeset
|
211 |
private class Job(progress: Progress, name: String, val info: Sessions.Info, tree: Sessions.Tree, |
| 64265 | 212 |
store: Sessions.Store, do_output: Boolean, verbose: Boolean, val numa_node: Option[Int], |
| 62633 | 213 |
session_graph: Graph_Display.Graph, command_timings: List[Properties.T]) |
| 48418 | 214 |
{
|
| 62633 | 215 |
val output = store.output_dir + Path.basic(name) |
| 48674 | 216 |
def output_path: Option[Path] = if (do_output) Some(output) else None |
| 62630 | 217 |
def output_save_state: String = |
|
62825
e6e80a8bf624
structure PolyML is sealed after bootstrap: all ML system access is managed by Isabelle;
wenzelm
parents:
62714
diff
changeset
|
218 |
if (do_output) "ML_Heap.save_child " + ML_Syntax.print_string0(File.platform_path(output)) |
| 62630 | 219 |
else "" |
| 62633 | 220 |
output.file.delete |
| 48674 | 221 |
|
222 |
private val parent = info.parent.getOrElse("")
|
|
| 48418 | 223 |
|
| 59445 | 224 |
private val graph_file = Isabelle_System.tmp_file("session_graph", "pdf")
|
225 |
try { isabelle.graphview.Graph_File.write(info.options, graph_file, session_graph) }
|
|
226 |
catch { case ERROR(_) => /*error should be exposed in ML*/ }
|
|
227 |
||
|
62610
4c89504c76fb
more uniform signature for various process invocations;
wenzelm
parents:
62603
diff
changeset
|
228 |
private val env = |
|
4c89504c76fb
more uniform signature for various process invocations;
wenzelm
parents:
62603
diff
changeset
|
229 |
Isabelle_System.settings() + |
|
4c89504c76fb
more uniform signature for various process invocations;
wenzelm
parents:
62603
diff
changeset
|
230 |
("ISABELLE_ML_DEBUGGER" -> info.options.bool("ML_debugger").toString)
|
|
48698
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
wenzelm
parents:
48684
diff
changeset
|
231 |
|
| 62573 | 232 |
private val future_result: Future[Process_Result] = |
|
61559
313eca3fa847
more direct task future implementation, with proper cancel operation;
wenzelm
parents:
61556
diff
changeset
|
233 |
Future.thread("build") {
|
|
62944
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
234 |
val args_file = Isabelle_System.tmp_file("build")
|
|
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
235 |
File.write(args_file, YXML.string_of_body( |
|
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
236 |
{
|
|
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
237 |
val theories = info.theories.map(x => (x._2, x._3)) |
|
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
238 |
import XML.Encode._ |
|
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
239 |
pair(list(pair(string, int)), pair(list(properties), pair(bool, pair(bool, |
|
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
240 |
pair(Path.encode, pair(list(pair(Path.encode, Path.encode)), pair(string, |
|
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
241 |
pair(string, pair(string, pair(string, |
|
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
242 |
list(pair(Options.encode, list(Path.encode)))))))))))))( |
|
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
243 |
(Symbol.codes, (command_timings, (do_output, (verbose, |
|
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
244 |
(store.browser_info, (info.document_files, (File.standard_path(graph_file), |
|
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
245 |
(parent, (info.chapter, (name, |
|
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
246 |
theories))))))))))) |
|
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
247 |
})) |
|
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
248 |
|
|
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
249 |
val eval = |
|
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
250 |
"Command_Line.tool0 (fn () => (" +
|
|
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
251 |
"Build.build " + ML_Syntax.print_string0(File.standard_path(args_file)) + |
|
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
252 |
(if (do_output) "; ML_Heap.share_common_data (); " + output_save_state |
|
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
253 |
else "") + "));" |
| 64265 | 254 |
|
255 |
val process_options = |
|
256 |
numa_node match {
|
|
257 |
case None => info.options |
|
258 |
case Some(n) => info.options.string("ML_process_policy") = NUMA.policy(n)
|
|
259 |
} |
|
| 62573 | 260 |
val process = |
| 62883 | 261 |
if (Sessions.pure_name(name)) {
|
| 64265 | 262 |
ML_Process(process_options, raw_ml_system = true, cwd = info.dir.file, |
| 62946 | 263 |
args = |
264 |
(for ((root, _) <- Thy_Header.ml_roots) yield List("--use", root)).flatten :::
|
|
265 |
List("--eval", eval),
|
|
|
62944
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
266 |
env = env, tree = Some(tree), store = store, cleanup = () => args_file.delete) |
| 62573 | 267 |
} |
268 |
else {
|
|
| 64265 | 269 |
ML_Process(process_options, parent, List("--eval", eval), cwd = info.dir.file,
|
|
62637
0189fe0f6452
support for Poly/ML heap hierarchy, which saves a lot of disk space;
wenzelm
parents:
62636
diff
changeset
|
270 |
env = env, tree = Some(tree), store = store, cleanup = () => args_file.delete) |
| 62573 | 271 |
} |
| 64265 | 272 |
|
| 62573 | 273 |
process.result( |
|
51962
016cb7d8f297
limit build process output, to avoid bombing Isabelle/Scala process by ill-behaved jobs (e.g. Containers in AFP/9025435b29cf);
wenzelm
parents:
51663
diff
changeset
|
274 |
progress_stdout = (line: String) => |
| 50847 | 275 |
Library.try_unprefix("\floading_theory = ", line) match {
|
276 |
case Some(theory) => progress.theory(name, theory) |
|
277 |
case None => |
|
|
51962
016cb7d8f297
limit build process output, to avoid bombing Isabelle/Scala process by ill-behaved jobs (e.g. Containers in AFP/9025435b29cf);
wenzelm
parents:
51663
diff
changeset
|
278 |
}, |
|
016cb7d8f297
limit build process output, to avoid bombing Isabelle/Scala process by ill-behaved jobs (e.g. Containers in AFP/9025435b29cf);
wenzelm
parents:
51663
diff
changeset
|
279 |
progress_limit = |
|
016cb7d8f297
limit build process output, to avoid bombing Isabelle/Scala process by ill-behaved jobs (e.g. Containers in AFP/9025435b29cf);
wenzelm
parents:
51663
diff
changeset
|
280 |
info.options.int("process_output_limit") match {
|
|
016cb7d8f297
limit build process output, to avoid bombing Isabelle/Scala process by ill-behaved jobs (e.g. Containers in AFP/9025435b29cf);
wenzelm
parents:
51663
diff
changeset
|
281 |
case 0 => None |
|
016cb7d8f297
limit build process output, to avoid bombing Isabelle/Scala process by ill-behaved jobs (e.g. Containers in AFP/9025435b29cf);
wenzelm
parents:
51663
diff
changeset
|
282 |
case m => Some(m * 1000000L) |
|
56871
d06ff36b4fa7
expose interrupts more like ML version, but not in managed bash processes of Build;
wenzelm
parents:
56861
diff
changeset
|
283 |
}, |
|
d06ff36b4fa7
expose interrupts more like ML version, but not in managed bash processes of Build;
wenzelm
parents:
56861
diff
changeset
|
284 |
strict = false) |
| 50845 | 285 |
} |
| 48674 | 286 |
|
| 62572 | 287 |
def terminate: Unit = future_result.cancel |
288 |
def is_finished: Boolean = future_result.is_finished |
|
| 48674 | 289 |
|
| 62569 | 290 |
@volatile private var was_timeout = false |
| 56779 | 291 |
private val timeout_request: Option[Event_Timer.Request] = |
292 |
{
|
|
| 62569 | 293 |
if (info.timeout > Time.zero) |
294 |
Some(Event_Timer.request(Time.now() + info.timeout) { terminate; was_timeout = true })
|
|
| 48674 | 295 |
else None |
| 56779 | 296 |
} |
| 48674 | 297 |
|
| 62400 | 298 |
def join: Process_Result = |
| 50845 | 299 |
{
|
| 62572 | 300 |
val result = future_result.join |
| 50845 | 301 |
|
|
62944
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
302 |
if (result.ok) |
| 62633 | 303 |
Present.finish(progress, store.browser_info, graph_file, info, name) |
| 61372 | 304 |
|
| 59445 | 305 |
graph_file.delete |
| 56779 | 306 |
timeout_request.foreach(_.cancel) |
| 48674 | 307 |
|
| 62572 | 308 |
if (result.interrupted) {
|
309 |
if (was_timeout) result.error(Output.error_text("Timeout")).was_timeout
|
|
310 |
else result.error(Output.error_text("Interrupt"))
|
|
|
52063
fd533ac64390
timeout counts as regular error, with rc = 1 (cf. special Exn.Interrupt vs. regular TimeLimit.TimeOut in Isabelle/ML);
wenzelm
parents:
51987
diff
changeset
|
311 |
} |
| 62572 | 312 |
else result |
| 48674 | 313 |
} |
| 48364 | 314 |
} |
315 |
||
| 48424 | 316 |
|
| 50777 | 317 |
/* sources and heaps */ |
318 |
||
| 62628 | 319 |
private val SOURCES = "sources: " |
320 |
private val INPUT_HEAP = "input_heap: " |
|
321 |
private val OUTPUT_HEAP = "output_heap: " |
|
322 |
private val LOG_START = "log:" |
|
323 |
private val line_prefixes = List(SOURCES, INPUT_HEAP, OUTPUT_HEAP, LOG_START) |
|
|
48504
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
wenzelm
parents:
48494
diff
changeset
|
324 |
|
| 62628 | 325 |
private def sources_stamp(digests: List[SHA1.Digest]): String = |
326 |
digests.map(_.toString).sorted.mkString(SOURCES, " ", "") |
|
|
48639
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
wenzelm
parents:
48626
diff
changeset
|
327 |
|
| 62628 | 328 |
private def read_stamps(path: Path): Option[(String, List[String], List[String])] = |
|
48639
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
wenzelm
parents:
48626
diff
changeset
|
329 |
if (path.is_file) {
|
| 62628 | 330 |
val stream = new GZIPInputStream(new BufferedInputStream(new FileInputStream(path.file))) |
| 50203 | 331 |
val reader = new BufferedReader(new InputStreamReader(stream, UTF8.charset)) |
| 62628 | 332 |
val lines = |
333 |
{
|
|
334 |
val lines = new mutable.ListBuffer[String] |
|
335 |
try {
|
|
336 |
var finished = false |
|
337 |
while (!finished) {
|
|
338 |
val line = reader.readLine |
|
339 |
if (line != null && line_prefixes.exists(line.startsWith(_))) |
|
340 |
lines += line |
|
341 |
else finished = true |
|
342 |
} |
|
343 |
} |
|
|
48639
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
wenzelm
parents:
48626
diff
changeset
|
344 |
finally { reader.close }
|
| 62628 | 345 |
lines.toList |
346 |
} |
|
347 |
||
348 |
if (!lines.isEmpty && lines.last.startsWith(LOG_START)) {
|
|
349 |
lines.find(_.startsWith(SOURCES)).map(s => |
|
350 |
(s, lines.filter(_.startsWith(INPUT_HEAP)), lines.filter(_.startsWith(OUTPUT_HEAP)))) |
|
351 |
} |
|
|
48504
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
wenzelm
parents:
48494
diff
changeset
|
352 |
else None |
|
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
wenzelm
parents:
48494
diff
changeset
|
353 |
} |
|
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
wenzelm
parents:
48494
diff
changeset
|
354 |
else None |
|
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
wenzelm
parents:
48494
diff
changeset
|
355 |
|
|
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
wenzelm
parents:
48494
diff
changeset
|
356 |
|
| 62631 | 357 |
|
| 62641 | 358 |
/** build with results **/ |
| 48424 | 359 |
|
| 63996 | 360 |
class Results private[Build](results: Map[String, (Option[Process_Result], Sessions.Info)]) |
| 62403 | 361 |
{
|
362 |
def sessions: Set[String] = results.keySet |
|
|
63082
6af03422535a
expose Sessions.Info in Build.Results
Lars Hupel <lars.hupel@mytum.de>
parents:
62946
diff
changeset
|
363 |
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
|
364 |
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
|
365 |
def info(name: String): Sessions.Info = results(name)._2 |
|
6af03422535a
expose Sessions.Info in Build.Results
Lars Hupel <lars.hupel@mytum.de>
parents:
62946
diff
changeset
|
366 |
val rc = (0 /: results.iterator.map({ case (_, (Some(r), _)) => r.rc case (_, (None, _)) => 1 }))(_ max _)
|
| 62641 | 367 |
def ok: Boolean = rc == 0 |
| 62406 | 368 |
|
369 |
override def toString: String = rc.toString |
|
| 62403 | 370 |
} |
371 |
||
| 62641 | 372 |
def build( |
|
50404
898cac1dad5e
avoid startup within GUI thread -- it is only required later for dialog;
wenzelm
parents:
50367
diff
changeset
|
373 |
options: Options, |
| 64909 | 374 |
progress: Progress = No_Progress, |
|
48511
37999ee01156
remove old output heaps, to ensure that result is valid wrt. check_stamps;
wenzelm
parents:
48509
diff
changeset
|
375 |
build_heap: Boolean = false, |
| 48595 | 376 |
clean_build: Boolean = false, |
| 56890 | 377 |
dirs: List[Path] = Nil, |
378 |
select_dirs: List[Path] = Nil, |
|
| 64265 | 379 |
numa_shuffling: Boolean = false, |
| 48509 | 380 |
max_jobs: Int = 1, |
| 48903 | 381 |
list_files: Boolean = false, |
|
59891
9ce697050455
added isabelle build option -k, for fast off-line checking of theory sources;
wenzelm
parents:
59811
diff
changeset
|
382 |
check_keywords: Set[String] = Set.empty, |
| 48509 | 383 |
no_build: Boolean = false, |
384 |
system_mode: Boolean = false, |
|
385 |
verbose: Boolean = false, |
|
| 63224 | 386 |
requirements: Boolean = false, |
387 |
all_sessions: Boolean = false, |
|
388 |
exclude_session_groups: List[String] = Nil, |
|
|
59892
2a616319c171
added isabelle build option -x, to exclude sessions;
wenzelm
parents:
59891
diff
changeset
|
389 |
exclude_sessions: List[String] = Nil, |
| 63224 | 390 |
session_groups: List[String] = Nil, |
| 62641 | 391 |
sessions: List[String] = Nil): Results = |
| 48341 | 392 |
{
|
| 63224 | 393 |
build_selection( |
394 |
options = options, |
|
395 |
progress = progress, |
|
396 |
build_heap = build_heap, |
|
397 |
clean_build = clean_build, |
|
398 |
dirs = dirs, |
|
399 |
select_dirs = select_dirs, |
|
| 64265 | 400 |
numa_shuffling = numa_shuffling, |
| 63224 | 401 |
max_jobs = max_jobs, |
402 |
list_files = list_files, |
|
403 |
check_keywords = check_keywords, |
|
404 |
no_build = no_build, |
|
405 |
system_mode = system_mode, |
|
406 |
verbose = verbose, |
|
407 |
selection = { full_tree =>
|
|
408 |
full_tree.selection(requirements, all_sessions, |
|
409 |
exclude_session_groups, exclude_sessions, session_groups, sessions) }) |
|
410 |
} |
|
411 |
||
412 |
def build_selection( |
|
413 |
options: Options, |
|
| 64909 | 414 |
progress: Progress = No_Progress, |
| 63224 | 415 |
build_heap: Boolean = false, |
416 |
clean_build: Boolean = false, |
|
417 |
dirs: List[Path] = Nil, |
|
418 |
select_dirs: List[Path] = Nil, |
|
| 64265 | 419 |
numa_shuffling: Boolean = false, |
| 63224 | 420 |
max_jobs: Int = 1, |
421 |
list_files: Boolean = false, |
|
422 |
check_keywords: Set[String] = Set.empty, |
|
423 |
no_build: Boolean = false, |
|
424 |
system_mode: Boolean = false, |
|
425 |
verbose: Boolean = false, |
|
426 |
selection: Sessions.Tree => (List[String], Sessions.Tree) = |
|
427 |
(_.selection(all_sessions = true))): Results = |
|
428 |
{
|
|
429 |
/* session selection and dependencies */ |
|
| 51220 | 430 |
|
| 62714 | 431 |
val build_options = options.int.update("completion_limit", 0).bool.update("ML_statistics", true)
|
432 |
val full_tree = Sessions.load(build_options, dirs, select_dirs) |
|
| 63224 | 433 |
val (selected, selected_tree) = selection(full_tree) |
|
59891
9ce697050455
added isabelle build option -k, for fast off-line checking of theory sources;
wenzelm
parents:
59811
diff
changeset
|
434 |
val deps = dependencies(progress, true, verbose, list_files, check_keywords, selected_tree) |
| 48368 | 435 |
|
| 62628 | 436 |
def session_sources_stamp(name: String): String = |
| 62631 | 437 |
sources_stamp(selected_tree(name).meta_digest :: deps.sources(name)) |
|
48504
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
wenzelm
parents:
48494
diff
changeset
|
438 |
|
| 62632 | 439 |
val store = Sessions.store(system_mode) |
| 51220 | 440 |
|
441 |
||
442 |
/* queue with scheduling information */ |
|
443 |
||
|
51230
19192615911e
option parallel_proofs_reuse_timing controls reuse of log information -- since it is not always beneficial for performance;
wenzelm
parents:
51229
diff
changeset
|
444 |
def load_timings(name: String): (List[Properties.T], Double) = |
| 51221 | 445 |
{
|
446 |
val (path, text) = |
|
| 62632 | 447 |
store.find_log_gz(name) match {
|
448 |
case Some(path) => (path, File.read_gzip(path)) |
|
| 51221 | 449 |
case None => |
| 62632 | 450 |
store.find_log(name) match {
|
451 |
case Some(path) => (path, File.read(path)) |
|
| 51221 | 452 |
case None => (Path.current, "") |
453 |
} |
|
454 |
} |
|
|
51244
d8ca566b22b3
more robust load_timings: ignore XML.Decode errors as well;
wenzelm
parents:
51230
diff
changeset
|
455 |
|
|
d8ca566b22b3
more robust load_timings: ignore XML.Decode errors as well;
wenzelm
parents:
51230
diff
changeset
|
456 |
def ignore_error(msg: String): (List[Properties.T], Double) = |
|
d8ca566b22b3
more robust load_timings: ignore XML.Decode errors as well;
wenzelm
parents:
51230
diff
changeset
|
457 |
{
|
|
56782
433cf57550fa
more systematic Isabelle output, like in classic Isabelle/ML (without markup);
wenzelm
parents:
56780
diff
changeset
|
458 |
Output.warning("Ignoring bad log file: " + path + (if (msg == "") "" else "\n" + msg))
|
|
51244
d8ca566b22b3
more robust load_timings: ignore XML.Decode errors as well;
wenzelm
parents:
51230
diff
changeset
|
459 |
(Nil, 0.0) |
|
d8ca566b22b3
more robust load_timings: ignore XML.Decode errors as well;
wenzelm
parents:
51230
diff
changeset
|
460 |
} |
|
d8ca566b22b3
more robust load_timings: ignore XML.Decode errors as well;
wenzelm
parents:
51230
diff
changeset
|
461 |
|
| 51221 | 462 |
try {
|
| 64082 | 463 |
val info = Build_Log.Log_File(name, text).parse_session_info(name, command_timings = true) |
| 51221 | 464 |
val session_timing = Markup.Elapsed.unapply(info.session_timing) getOrElse 0.0 |
465 |
(info.command_timings, session_timing) |
|
| 51220 | 466 |
} |
| 51221 | 467 |
catch {
|
|
51244
d8ca566b22b3
more robust load_timings: ignore XML.Decode errors as well;
wenzelm
parents:
51230
diff
changeset
|
468 |
case ERROR(msg) => ignore_error(msg) |
|
51986
5fdca5bfc0b4
more robust load_timings: ignore JVM errors such as java.lang.OutOfMemoryError;
wenzelm
parents:
51983
diff
changeset
|
469 |
case exn: java.lang.Error => ignore_error(Exn.message(exn)) |
| 51987 | 470 |
case _: XML.Error => ignore_error("")
|
| 51221 | 471 |
} |
472 |
} |
|
| 51220 | 473 |
|
|
51230
19192615911e
option parallel_proofs_reuse_timing controls reuse of log information -- since it is not always beneficial for performance;
wenzelm
parents:
51229
diff
changeset
|
474 |
val queue = Queue(selected_tree, load_timings) |
| 51220 | 475 |
|
476 |
||
477 |
/* main build process */ |
|
478 |
||
| 62632 | 479 |
store.prepare_output() |
| 48373 | 480 |
|
| 48595 | 481 |
// optional cleanup |
482 |
if (clean_build) {
|
|
| 49131 | 483 |
for (name <- full_tree.graph.all_succs(selected)) {
|
| 48595 | 484 |
val files = |
| 62632 | 485 |
List(Path.basic(name), Sessions.log(name), Sessions.log_gz(name)). |
486 |
map(store.output_dir + _).filter(_.is_file) |
|
| 59319 | 487 |
if (files.nonEmpty) progress.echo("Cleaning " + name + " ...")
|
| 50366 | 488 |
if (!files.forall(p => p.file.delete)) progress.echo(name + " FAILED to delete") |
| 48595 | 489 |
} |
490 |
} |
|
491 |
||
| 48425 | 492 |
// scheduler loop |
|
63082
6af03422535a
expose Sessions.Info in Build.Results
Lars Hupel <lars.hupel@mytum.de>
parents:
62946
diff
changeset
|
493 |
case class Result( |
|
6af03422535a
expose Sessions.Info in Build.Results
Lars Hupel <lars.hupel@mytum.de>
parents:
62946
diff
changeset
|
494 |
current: Boolean, heap_stamp: Option[String], |
|
6af03422535a
expose Sessions.Info in Build.Results
Lars Hupel <lars.hupel@mytum.de>
parents:
62946
diff
changeset
|
495 |
process: Option[Process_Result], info: Sessions.Info) |
| 62402 | 496 |
{
|
497 |
def ok: Boolean = |
|
498 |
process match {
|
|
499 |
case None => false |
|
500 |
case Some(res) => res.rc == 0 |
|
501 |
} |
|
502 |
} |
|
|
48639
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
wenzelm
parents:
48626
diff
changeset
|
503 |
|
|
56837
5a598f1eecfd
more robust interrupt handling for Scala_Console, which uses JVM Thread.interrupt instead of POSIX SIGINT;
wenzelm
parents:
56831
diff
changeset
|
504 |
def sleep() |
|
5a598f1eecfd
more robust interrupt handling for Scala_Console, which uses JVM Thread.interrupt instead of POSIX SIGINT;
wenzelm
parents:
56831
diff
changeset
|
505 |
{
|
|
5a598f1eecfd
more robust interrupt handling for Scala_Console, which uses JVM Thread.interrupt instead of POSIX SIGINT;
wenzelm
parents:
56831
diff
changeset
|
506 |
try { Thread.sleep(500) }
|
| 56861 | 507 |
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
|
508 |
} |
| 50366 | 509 |
|
| 64265 | 510 |
val numa_nodes = new NUMA.Nodes(numa_shuffling) |
511 |
||
| 48425 | 512 |
@tailrec def loop( |
|
48676
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
wenzelm
parents:
48675
diff
changeset
|
513 |
pending: Queue, |
| 62628 | 514 |
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
|
515 |
results: Map[String, Result]): Map[String, Result] = |
| 48425 | 516 |
{
|
| 64265 | 517 |
def used_node(i: Int): Boolean = |
518 |
running.iterator.exists( |
|
519 |
{ case (_, (_, job)) => job.numa_node.isDefined && job.numa_node.get == i })
|
|
520 |
||
| 48425 | 521 |
if (pending.is_empty) results |
| 51253 | 522 |
else {
|
523 |
if (progress.stopped) |
|
524 |
for ((_, (_, job)) <- running) job.terminate |
|
525 |
||
| 48674 | 526 |
running.find({ case (_, (_, job)) => job.is_finished }) match {
|
| 62628 | 527 |
case Some((name, (input_heaps, job))) => |
| 50367 | 528 |
//{{{ finish job
|
| 48424 | 529 |
|
| 62401 | 530 |
val process_result = job.join |
| 62573 | 531 |
process_result.err_lines.foreach(progress.echo(_)) |
532 |
if (process_result.ok) |
|
533 |
progress.echo("Finished " + name + " (" + process_result.timing.message_resources + ")")
|
|
| 48373 | 534 |
|
|
62404
13a0f537e232
retain tail out_lines as printed, but not the whole log content;
wenzelm
parents:
62403
diff
changeset
|
535 |
val process_result_tail = |
|
13a0f537e232
retain tail out_lines as printed, but not the whole log content;
wenzelm
parents:
62403
diff
changeset
|
536 |
{
|
|
13a0f537e232
retain tail out_lines as printed, but not the whole log content;
wenzelm
parents:
62403
diff
changeset
|
537 |
val lines = process_result.out_lines.filterNot(_.startsWith("\f"))
|
|
62409
e391528eff3b
proper option process_output_tail, more generous default;
wenzelm
parents:
62406
diff
changeset
|
538 |
val tail = job.info.options.int("process_output_tail")
|
|
e391528eff3b
proper option process_output_tail, more generous default;
wenzelm
parents:
62406
diff
changeset
|
539 |
val lines1 = if (tail == 0) lines else lines.drop(lines.length - tail max 0) |
| 62632 | 540 |
process_result.copy( |
541 |
out_lines = |
|
542 |
"(see also " + (store.output_dir + Sessions.log(name)).file.toString + ")" :: |
|
543 |
lines1) |
|
|
62404
13a0f537e232
retain tail out_lines as printed, but not the whole log content;
wenzelm
parents:
62403
diff
changeset
|
544 |
} |
|
13a0f537e232
retain tail out_lines as printed, but not the whole log content;
wenzelm
parents:
62403
diff
changeset
|
545 |
|
| 62636 | 546 |
val heap_stamp = |
| 62401 | 547 |
if (process_result.ok) {
|
| 62632 | 548 |
(store.output_dir + Sessions.log(name)).file.delete |
| 62636 | 549 |
val heap_stamp = |
|
62704
478b49f0d726
proper SHA1 digest as annex to heap file: Poly/ML reads precise segment length;
wenzelm
parents:
62643
diff
changeset
|
550 |
for (path <- job.output_path if path.is_file) |
|
478b49f0d726
proper SHA1 digest as annex to heap file: Poly/ML reads precise segment length;
wenzelm
parents:
62643
diff
changeset
|
551 |
yield Sessions.write_heap_digest(path) |
|
48639
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
wenzelm
parents:
48626
diff
changeset
|
552 |
|
| 62632 | 553 |
File.write_gzip(store.output_dir + Sessions.log_gz(name), |
| 64173 | 554 |
terminate_lines( |
| 62628 | 555 |
session_sources_stamp(name) :: |
556 |
input_heaps.map(INPUT_HEAP + _) ::: |
|
| 62636 | 557 |
heap_stamp.toList.map(OUTPUT_HEAP + _) ::: |
| 62628 | 558 |
List(LOG_START) ::: process_result.out_lines)) |
|
48639
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
wenzelm
parents:
48626
diff
changeset
|
559 |
|
| 62636 | 560 |
heap_stamp |
|
48639
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
wenzelm
parents:
48626
diff
changeset
|
561 |
} |
|
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
wenzelm
parents:
48626
diff
changeset
|
562 |
else {
|
| 62632 | 563 |
(store.output_dir + Path.basic(name)).file.delete |
564 |
(store.output_dir + Sessions.log_gz(name)).file.delete |
|
|
48639
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
wenzelm
parents:
48626
diff
changeset
|
565 |
|
| 62632 | 566 |
File.write(store.output_dir + Sessions.log(name), |
| 64173 | 567 |
terminate_lines(process_result.out_lines)) |
| 50366 | 568 |
progress.echo(name + " FAILED") |
|
62404
13a0f537e232
retain tail out_lines as printed, but not the whole log content;
wenzelm
parents:
62403
diff
changeset
|
569 |
if (!process_result.interrupted) progress.echo(process_result_tail.out) |
|
48639
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
wenzelm
parents:
48626
diff
changeset
|
570 |
|
| 62636 | 571 |
None |
|
48639
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
wenzelm
parents:
48626
diff
changeset
|
572 |
} |
|
50707
5b2bf7611662
maintain session index on Scala side, for more determistic results;
wenzelm
parents:
50686
diff
changeset
|
573 |
loop(pending - name, running - name, |
|
63082
6af03422535a
expose Sessions.Info in Build.Results
Lars Hupel <lars.hupel@mytum.de>
parents:
62946
diff
changeset
|
574 |
results + (name -> Result(false, heap_stamp, Some(process_result_tail), job.info))) |
| 50367 | 575 |
//}}} |
| 60215 | 576 |
case None if running.size < (max_jobs max 1) => |
| 50367 | 577 |
//{{{ check/start next job
|
| 48547 | 578 |
pending.dequeue(running.isDefinedAt(_)) match {
|
579 |
case Some((name, info)) => |
|
| 62628 | 580 |
val ancestor_results = selected_tree.ancestors(name).map(results(_)) |
| 62636 | 581 |
val ancestor_heaps = ancestor_results.flatMap(_.heap_stamp) |
| 62628 | 582 |
|
| 62883 | 583 |
val do_output = build_heap || Sessions.pure_name(name) || queue.is_inner(name) |
| 48547 | 584 |
|
| 62636 | 585 |
val (current, heap_stamp) = |
| 48547 | 586 |
{
|
| 62632 | 587 |
store.find(name) match {
|
| 62636 | 588 |
case Some((log_gz, heap_stamp)) => |
| 62632 | 589 |
read_stamps(log_gz) match {
|
| 62628 | 590 |
case Some((sources, input_heaps, output_heaps)) => |
591 |
val current = |
|
592 |
sources == session_sources_stamp(name) && |
|
593 |
input_heaps == ancestor_heaps.map(INPUT_HEAP + _) && |
|
| 62636 | 594 |
output_heaps == heap_stamp.toList.map(OUTPUT_HEAP + _) && |
595 |
!(do_output && heap_stamp.isEmpty) |
|
596 |
(current, heap_stamp) |
|
597 |
case None => (false, None) |
|
| 48547 | 598 |
} |
| 62636 | 599 |
case None => (false, None) |
|
48504
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
wenzelm
parents:
48494
diff
changeset
|
600 |
} |
| 48547 | 601 |
} |
| 62628 | 602 |
val all_current = current && ancestor_results.forall(_.current) |
|
48528
784c6f63d79c
proper all_current, which regards parent status as well;
wenzelm
parents:
48511
diff
changeset
|
603 |
|
| 48547 | 604 |
if (all_current) |
| 62402 | 605 |
loop(pending - name, running, |
|
63082
6af03422535a
expose Sessions.Info in Build.Results
Lars Hupel <lars.hupel@mytum.de>
parents:
62946
diff
changeset
|
606 |
results + (name -> Result(true, heap_stamp, Some(Process_Result(0)), info))) |
|
48678
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
607 |
else if (no_build) {
|
| 50366 | 608 |
if (verbose) progress.echo("Skipping " + name + " ...")
|
| 62402 | 609 |
loop(pending - name, running, |
|
63082
6af03422535a
expose Sessions.Info in Build.Results
Lars Hupel <lars.hupel@mytum.de>
parents:
62946
diff
changeset
|
610 |
results + (name -> Result(false, heap_stamp, Some(Process_Result(1)), info))) |
|
48678
ff27af15530c
queue ordering by descending outdegree and timeout;
wenzelm
parents:
48676
diff
changeset
|
611 |
} |
| 62628 | 612 |
else if (ancestor_results.forall(_.ok) && !progress.stopped) {
|
| 64265 | 613 |
val numa_node = numa_nodes.next(used_node(_)) |
| 50366 | 614 |
progress.echo((if (do_output) "Building " else "Running ") + name + " ...") |
| 51220 | 615 |
val job = |
|
62637
0189fe0f6452
support for Poly/ML heap hierarchy, which saves a lot of disk space;
wenzelm
parents:
62636
diff
changeset
|
616 |
new Job(progress, name, info, selected_tree, store, do_output, verbose, |
| 64265 | 617 |
numa_node, deps(name).session_graph, queue.command_timings(name)) |
| 62628 | 618 |
loop(pending, running + (name -> (ancestor_heaps, job)), results) |
| 48547 | 619 |
} |
620 |
else {
|
|
| 50366 | 621 |
progress.echo(name + " CANCELLED") |
|
63082
6af03422535a
expose Sessions.Info in Build.Results
Lars Hupel <lars.hupel@mytum.de>
parents:
62946
diff
changeset
|
622 |
loop(pending - name, running, results + (name -> Result(false, heap_stamp, None, info))) |
| 48547 | 623 |
} |
624 |
case None => sleep(); loop(pending, running, results) |
|
| 48425 | 625 |
} |
| 50367 | 626 |
///}}} |
| 48425 | 627 |
case None => sleep(); loop(pending, running, results) |
| 48373 | 628 |
} |
| 51253 | 629 |
} |
| 48425 | 630 |
} |
631 |
||
| 51220 | 632 |
|
633 |
/* build results */ |
|
634 |
||
| 62641 | 635 |
val results0 = |
| 48583 | 636 |
if (deps.is_empty) {
|
|
56782
433cf57550fa
more systematic Isabelle output, like in classic Isabelle/ML (without markup);
wenzelm
parents:
56780
diff
changeset
|
637 |
progress.echo(Output.warning_text("Nothing to build"))
|
|
50707
5b2bf7611662
maintain session index on Scala side, for more determistic results;
wenzelm
parents:
50686
diff
changeset
|
638 |
Map.empty[String, Result] |
| 48583 | 639 |
} |
640 |
else loop(queue, Map.empty, Map.empty) |
|
641 |
||
| 62641 | 642 |
val results = |
| 64265 | 643 |
new Results( |
644 |
(for ((name, result) <- results0.iterator) |
|
645 |
yield (name, (result.process, result.info))).toMap) |
|
| 62641 | 646 |
|
647 |
if (results.rc != 0 && (verbose || !no_build)) {
|
|
648 |
val unfinished = |
|
649 |
(for {
|
|
650 |
name <- results.sessions.iterator |
|
651 |
if !results(name).ok |
|
652 |
} yield name).toList.sorted |
|
653 |
progress.echo("Unfinished session(s): " + commas(unfinished))
|
|
654 |
} |
|
655 |
||
|
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
|
656 |
|
|
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
|
657 |
/* 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
|
658 |
|
|
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
|
659 |
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
|
660 |
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
|
661 |
(for {
|
| 62641 | 662 |
(name, result) <- results0.iterator |
|
62944
3ee643c5ed00
more standard session build process, including browser_info;
wenzelm
parents:
62902
diff
changeset
|
663 |
if result.ok |
|
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
|
664 |
info = full_tree(name) |
|
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
|
665 |
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
|
666 |
} 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
|
667 |
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
|
668 |
|
|
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
|
669 |
for ((chapter, entries) <- browser_chapters) |
| 62632 | 670 |
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
|
671 |
|
| 62632 | 672 |
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
|
673 |
} |
|
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
|
674 |
|
| 62641 | 675 |
results |
| 48341 | 676 |
} |
677 |
||
678 |
||
| 62833 | 679 |
/* Isabelle tool wrapper */ |
| 48341 | 680 |
|
| 62833 | 681 |
val isabelle_tool = Isabelle_Tool("build", "build and manage Isabelle sessions", args =>
|
| 48341 | 682 |
{
|
| 62833 | 683 |
val build_options = Word.explode(Isabelle_System.getenv("ISABELLE_BUILD_OPTIONS"))
|
| 62590 | 684 |
|
| 62833 | 685 |
var select_dirs: List[Path] = Nil |
| 64265 | 686 |
var numa_shuffling = false |
| 62833 | 687 |
var requirements = false |
688 |
var exclude_session_groups: List[String] = Nil |
|
689 |
var all_sessions = false |
|
690 |
var build_heap = false |
|
691 |
var clean_build = false |
|
692 |
var dirs: List[Path] = Nil |
|
693 |
var session_groups: List[String] = Nil |
|
694 |
var max_jobs = 1 |
|
695 |
var check_keywords: Set[String] = Set.empty |
|
696 |
var list_files = false |
|
697 |
var no_build = false |
|
698 |
var options = (Options.init() /: build_options)(_ + _) |
|
699 |
var system_mode = false |
|
700 |
var verbose = false |
|
701 |
var exclude_sessions: List[String] = Nil |
|
| 62590 | 702 |
|
| 62833 | 703 |
val getopts = Getopts("""
|
| 62590 | 704 |
Usage: isabelle build [OPTIONS] [SESSIONS ...] |
705 |
||
706 |
Options are: |
|
707 |
-D DIR include session directory and select its sessions |
|
| 64265 | 708 |
-N cyclic shuffling of NUMA CPU nodes (performance tuning) |
| 62590 | 709 |
-R operate on requirements of selected sessions |
710 |
-X NAME exclude sessions from group NAME and all descendants |
|
711 |
-a select all sessions |
|
712 |
-b build heap images |
|
713 |
-c clean build |
|
714 |
-d DIR include session directory |
|
715 |
-g NAME select session group NAME |
|
716 |
-j INT maximum number of parallel jobs (default 1) |
|
717 |
-k KEYWORD check theory sources for conflicts with proposed keywords |
|
718 |
-l list session source files |
|
719 |
-n no build -- test dependencies only |
|
720 |
-o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) |
|
721 |
-s system build mode: produce output in ISABELLE_HOME |
|
722 |
-v verbose |
|
723 |
-x NAME exclude session NAME and all descendants |
|
724 |
||
| 62596 | 725 |
Build and manage Isabelle sessions, depending on implicit settings: |
726 |
||
| 64455 | 727 |
""" + Library.prefix_lines(" ", Build_Log.Settings.show()) + "\n",
|
| 62833 | 728 |
"D:" -> (arg => select_dirs = select_dirs ::: List(Path.explode(arg))), |
| 64265 | 729 |
"N" -> (_ => numa_shuffling = true), |
| 62833 | 730 |
"R" -> (_ => requirements = true), |
731 |
"X:" -> (arg => exclude_session_groups = exclude_session_groups ::: List(arg)), |
|
732 |
"a" -> (_ => all_sessions = true), |
|
733 |
"b" -> (_ => build_heap = true), |
|
734 |
"c" -> (_ => clean_build = true), |
|
735 |
"d:" -> (arg => dirs = dirs ::: List(Path.explode(arg))), |
|
736 |
"g:" -> (arg => session_groups = session_groups ::: List(arg)), |
|
| 63805 | 737 |
"j:" -> (arg => max_jobs = Value.Int.parse(arg)), |
| 62833 | 738 |
"k:" -> (arg => check_keywords = check_keywords + arg), |
739 |
"l" -> (_ => list_files = true), |
|
740 |
"n" -> (_ => no_build = true), |
|
741 |
"o:" -> (arg => options = options + arg), |
|
742 |
"s" -> (_ => system_mode = true), |
|
743 |
"v" -> (_ => verbose = true), |
|
744 |
"x:" -> (arg => exclude_sessions = exclude_sessions ::: List(arg))) |
|
| 62590 | 745 |
|
| 62833 | 746 |
val sessions = getopts(args) |
| 62590 | 747 |
|
| 64115 | 748 |
val progress = new Console_Progress(verbose = verbose) |
| 62590 | 749 |
|
| 64140 | 750 |
val start_date = Date.now() |
751 |
||
| 62833 | 752 |
if (verbose) {
|
753 |
progress.echo( |
|
| 64155 | 754 |
"Started at " + Build_Log.print_date(start_date) + |
| 64140 | 755 |
" (" + Isabelle_System.getenv("ML_IDENTIFIER") + " on " + Isabelle_System.hostname() +")")
|
| 64081 | 756 |
progress.echo(Build_Log.Settings.show() + "\n") |
| 62833 | 757 |
} |
| 62590 | 758 |
|
| 62833 | 759 |
val results = |
760 |
progress.interrupt_handler {
|
|
| 63224 | 761 |
build(options, progress, |
762 |
build_heap = build_heap, |
|
763 |
clean_build = clean_build, |
|
764 |
dirs = dirs, |
|
765 |
select_dirs = select_dirs, |
|
| 64265 | 766 |
numa_shuffling = NUMA.enabled_warning(numa_shuffling), |
| 63224 | 767 |
max_jobs = max_jobs, |
768 |
list_files = list_files, |
|
769 |
check_keywords = check_keywords, |
|
770 |
no_build = no_build, |
|
771 |
system_mode = system_mode, |
|
772 |
verbose = verbose, |
|
773 |
requirements = requirements, |
|
774 |
all_sessions = all_sessions, |
|
775 |
exclude_session_groups = exclude_session_groups, |
|
776 |
exclude_sessions = exclude_sessions, |
|
777 |
session_groups = session_groups, |
|
778 |
sessions = sessions) |
|
| 62833 | 779 |
} |
| 64140 | 780 |
val end_date = Date.now() |
781 |
val elapsed_time = end_date.time - start_date.time |
|
| 62590 | 782 |
|
| 62833 | 783 |
if (verbose) {
|
| 64155 | 784 |
progress.echo("\nFinished at " + Build_Log.print_date(end_date))
|
| 62833 | 785 |
} |
| 62590 | 786 |
|
| 62833 | 787 |
val total_timing = |
788 |
(Timing.zero /: results.sessions.iterator.map(a => results(a).timing))(_ + _). |
|
789 |
copy(elapsed = elapsed_time) |
|
790 |
progress.echo(total_timing.message_resources) |
|
| 62590 | 791 |
|
| 62833 | 792 |
sys.exit(results.rc) |
793 |
}) |
|
|
59366
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
794 |
|
|
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
795 |
|
|
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
796 |
/* PIDE protocol */ |
|
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
797 |
|
|
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
798 |
def build_theories( |
|
59369
7090199d3f78
more informative build_theories_result: cumulative Runtime.exn_message;
wenzelm
parents:
59367
diff
changeset
|
799 |
session: Session, master_dir: Path, theories: List[(Options, List[Path])]): Promise[XML.Body] = |
| 59367 | 800 |
session.get_protocol_handler(classOf[Handler].getName) match {
|
|
59366
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
801 |
case Some(handler: Handler) => handler.build_theories(session, master_dir, theories) |
|
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
802 |
case _ => error("Cannot invoke build_theories: bad protocol handler")
|
|
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
803 |
} |
|
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
804 |
|
| 59367 | 805 |
class Handler(progress: Progress, session_name: String) extends Session.Protocol_Handler |
|
59366
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
806 |
{
|
|
59369
7090199d3f78
more informative build_theories_result: cumulative Runtime.exn_message;
wenzelm
parents:
59367
diff
changeset
|
807 |
private val pending = Synchronized(Map.empty[String, Promise[XML.Body]]) |
|
59366
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
808 |
|
| 65220 | 809 |
override def exit(): Unit = |
810 |
pending.change(promises => { for ((_, promise) <- promises) promise.cancel; Map.empty })
|
|
811 |
||
|
59366
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
812 |
def build_theories( |
|
59369
7090199d3f78
more informative build_theories_result: cumulative Runtime.exn_message;
wenzelm
parents:
59367
diff
changeset
|
813 |
session: Session, master_dir: Path, theories: List[(Options, List[Path])]): Promise[XML.Body] = |
|
59366
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
814 |
{
|
|
59369
7090199d3f78
more informative build_theories_result: cumulative Runtime.exn_message;
wenzelm
parents:
59367
diff
changeset
|
815 |
val promise = Future.promise[XML.Body] |
|
59366
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
816 |
val id = Document_ID.make().toString |
|
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
817 |
pending.change(promises => promises + (id -> promise)) |
|
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
818 |
session.build_theories(id, master_dir, theories) |
|
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
819 |
promise |
|
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
820 |
} |
|
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
821 |
|
| 65219 | 822 |
private def loading_theory(msg: Prover.Protocol_Output): Boolean = |
| 59367 | 823 |
msg.properties match {
|
824 |
case Markup.Loading_Theory(name) => progress.theory(session_name, name); true |
|
825 |
case _ => false |
|
826 |
} |
|
827 |
||
| 65219 | 828 |
private def build_theories_result(msg: Prover.Protocol_Output): Boolean = |
|
59366
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
829 |
msg.properties match {
|
|
59369
7090199d3f78
more informative build_theories_result: cumulative Runtime.exn_message;
wenzelm
parents:
59367
diff
changeset
|
830 |
case Markup.Build_Theories_Result(id) => |
|
59366
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
831 |
pending.change_result(promises => |
|
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
832 |
promises.get(id) match {
|
|
59369
7090199d3f78
more informative build_theories_result: cumulative Runtime.exn_message;
wenzelm
parents:
59367
diff
changeset
|
833 |
case Some(promise) => |
|
7090199d3f78
more informative build_theories_result: cumulative Runtime.exn_message;
wenzelm
parents:
59367
diff
changeset
|
834 |
val error_message = |
|
7090199d3f78
more informative build_theories_result: cumulative Runtime.exn_message;
wenzelm
parents:
59367
diff
changeset
|
835 |
try { YXML.parse_body(Symbol.decode(msg.text)) }
|
|
7090199d3f78
more informative build_theories_result: cumulative Runtime.exn_message;
wenzelm
parents:
59367
diff
changeset
|
836 |
catch { case exn: Throwable => List(XML.Text(Exn.message(exn))) }
|
|
7090199d3f78
more informative build_theories_result: cumulative Runtime.exn_message;
wenzelm
parents:
59367
diff
changeset
|
837 |
promise.fulfill(error_message) |
|
7090199d3f78
more informative build_theories_result: cumulative Runtime.exn_message;
wenzelm
parents:
59367
diff
changeset
|
838 |
(true, promises - id) |
|
7090199d3f78
more informative build_theories_result: cumulative Runtime.exn_message;
wenzelm
parents:
59367
diff
changeset
|
839 |
case None => |
|
7090199d3f78
more informative build_theories_result: cumulative Runtime.exn_message;
wenzelm
parents:
59367
diff
changeset
|
840 |
(false, promises) |
|
59366
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
841 |
}) |
|
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
842 |
case _ => false |
|
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
843 |
} |
|
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
844 |
|
| 59367 | 845 |
val functions = |
| 65219 | 846 |
List( |
| 59367 | 847 |
Markup.BUILD_THEORIES_RESULT -> build_theories_result _, |
848 |
Markup.LOADING_THEORY -> loading_theory _) |
|
|
59366
e94df7f6b608
clarified build_theories: proper protocol handler;
wenzelm
parents:
59319
diff
changeset
|
849 |
} |
| 48276 | 850 |
} |