| author | wenzelm | 
| Fri, 19 Aug 2022 20:07:41 +0200 | |
| changeset 75907 | 091edca12219 | 
| parent 75899 | d50c2129e73a | 
| child 75908 | 6b979348455e | 
| 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  | 
||
| 73364 | 11  | 
import scala.collection.immutable.SortedSet  | 
| 
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  | 
|
| 75393 | 15  | 
object Build {
 | 
| 62631 | 16  | 
/** auxiliary **/  | 
| 48424 | 17  | 
|
| 
65291
 
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
 
wenzelm 
parents: 
65289 
diff
changeset
 | 
18  | 
/* persistent build info */  | 
| 
 
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
 
wenzelm 
parents: 
65289 
diff
changeset
 | 
19  | 
|
| 
 
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
 
wenzelm 
parents: 
65289 
diff
changeset
 | 
20  | 
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
 | 
21  | 
sources: String,  | 
| 
65291
 
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
 
wenzelm 
parents: 
65289 
diff
changeset
 | 
22  | 
input_heaps: List[String],  | 
| 
 
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
 
wenzelm 
parents: 
65289 
diff
changeset
 | 
23  | 
output_heap: Option[String],  | 
| 75393 | 24  | 
return_code: Int  | 
25  | 
  ) {
 | 
|
| 66747 | 26  | 
def ok: Boolean = return_code == 0  | 
27  | 
}  | 
|
| 
65291
 
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
 
wenzelm 
parents: 
65289 
diff
changeset
 | 
28  | 
|
| 
 
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
 
wenzelm 
parents: 
65289 
diff
changeset
 | 
29  | 
|
| 65289 | 30  | 
/* 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
 | 
31  | 
|
| 75393 | 32  | 
  private object Queue {
 | 
| 
68086
 
9e1c670301b8
cleanup session output before starting build job;
 
wenzelm 
parents: 
67852 
diff
changeset
 | 
33  | 
type Timings = (List[Properties.T], Double)  | 
| 
 
9e1c670301b8
cleanup session output before starting build job;
 
wenzelm 
parents: 
67852 
diff
changeset
 | 
34  | 
|
| 75393 | 35  | 
    def load_timings(progress: Progress, store: Sessions.Store, session_name: String): Timings = {
 | 
| 
68086
 
9e1c670301b8
cleanup session output before starting build job;
 
wenzelm 
parents: 
67852 
diff
changeset
 | 
36  | 
val no_timings: Timings = (Nil, 0.0)  | 
| 65289 | 37  | 
|
| 
72634
 
5cea0993ee4f
clarified access to single database server vs. collection of database files;
 
wenzelm 
parents: 
72624 
diff
changeset
 | 
38  | 
      store.try_open_database(session_name) match {
 | 
| 
65291
 
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
 
wenzelm 
parents: 
65289 
diff
changeset
 | 
39  | 
case None => no_timings  | 
| 68214 | 40  | 
case Some(db) =>  | 
| 75393 | 41  | 
          def ignore_error(msg: String) = {
 | 
| 75793 | 42  | 
            progress.echo_warning("Ignoring bad database " + db +
 | 
43  | 
" for session " + quote(session_name) + (if (msg == "") "" else ":\n" + msg))  | 
|
| 
65291
 
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
 
wenzelm 
parents: 
65289 
diff
changeset
 | 
44  | 
no_timings  | 
| 
 
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
 
wenzelm 
parents: 
65289 
diff
changeset
 | 
45  | 
}  | 
| 
 
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
 
wenzelm 
parents: 
65289 
diff
changeset
 | 
46  | 
          try {
 | 
| 71669 | 47  | 
val command_timings = store.read_command_timings(db, session_name)  | 
| 68214 | 48  | 
val session_timing =  | 
| 71669 | 49  | 
              store.read_session_timing(db, session_name) match {
 | 
| 68214 | 50  | 
case Markup.Elapsed(t) => t  | 
51  | 
case _ => 0.0  | 
|
52  | 
}  | 
|
53  | 
(command_timings, session_timing)  | 
|
| 
65291
 
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
 
wenzelm 
parents: 
65289 
diff
changeset
 | 
54  | 
}  | 
| 
 
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
 
wenzelm 
parents: 
65289 
diff
changeset
 | 
55  | 
          catch {
 | 
| 
 
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
 
wenzelm 
parents: 
65289 
diff
changeset
 | 
56  | 
case ERROR(msg) => ignore_error(msg)  | 
| 
 
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
 
wenzelm 
parents: 
65289 
diff
changeset
 | 
57  | 
case exn: java.lang.Error => ignore_error(Exn.message(exn))  | 
| 75783 | 58  | 
            case _: XML.Error => ignore_error("XML.Error")
 | 
| 
65291
 
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
 
wenzelm 
parents: 
65289 
diff
changeset
 | 
59  | 
}  | 
| 73367 | 60  | 
          finally { db.close() }
 | 
| 65289 | 61  | 
}  | 
62  | 
}  | 
|
63  | 
||
| 75393 | 64  | 
def make_session_timing(  | 
65  | 
sessions_structure: Sessions.Structure,  | 
|
66  | 
timing: Map[String, Double]  | 
|
67  | 
    ) : Map[String, Double] = {
 | 
|
| 68731 | 68  | 
val maximals = sessions_structure.build_graph.maximals.toSet  | 
| 75393 | 69  | 
      def desc_timing(session_name: String): Double = {
 | 
| 71669 | 70  | 
if (maximals.contains(session_name)) timing(session_name)  | 
| 68487 | 71  | 
        else {
 | 
| 71669 | 72  | 
val descendants = sessions_structure.build_descendants(List(session_name)).toSet  | 
| 68731 | 73  | 
val g = sessions_structure.build_graph.restrict(descendants)  | 
| 75394 | 74  | 
          (0.0 :: g.maximals.flatMap { desc =>
 | 
| 68487 | 75  | 
val ps = g.all_preds(List(desc))  | 
| 71895 | 76  | 
if (ps.exists(p => !timing.isDefinedAt(p))) None  | 
| 68487 | 77  | 
else Some(ps.map(timing(_)).sum)  | 
| 75394 | 78  | 
}).max  | 
| 68487 | 79  | 
}  | 
| 
68486
 
6984a55f3cba
clarified queue ordering: take session descendants into account, notably for "slow" AFP sessions;
 
wenzelm 
parents: 
68331 
diff
changeset
 | 
80  | 
}  | 
| 
 
6984a55f3cba
clarified queue ordering: take session descendants into account, notably for "slow" AFP sessions;
 
wenzelm 
parents: 
68331 
diff
changeset
 | 
81  | 
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
 | 
82  | 
}  | 
| 
 
6984a55f3cba
clarified queue ordering: take session descendants into account, notably for "slow" AFP sessions;
 
wenzelm 
parents: 
68331 
diff
changeset
 | 
83  | 
|
| 75393 | 84  | 
def apply(  | 
85  | 
progress: Progress,  | 
|
86  | 
sessions_structure: Sessions.Structure,  | 
|
87  | 
store: Sessions.Store  | 
|
88  | 
    ) : Queue = {
 | 
|
| 68731 | 89  | 
val graph = sessions_structure.build_graph  | 
| 
65415
 
8cd54b18b68b
clarified signature: tree structure is not essential;
 
wenzelm 
parents: 
65406 
diff
changeset
 | 
90  | 
val names = graph.keys  | 
| 51220 | 91  | 
|
| 65831 | 92  | 
val timings = names.map(name => (name, load_timings(progress, store, name)))  | 
| 51220 | 93  | 
val command_timings =  | 
| 
68486
 
6984a55f3cba
clarified queue ordering: take session descendants into account, notably for "slow" AFP sessions;
 
wenzelm 
parents: 
68331 
diff
changeset
 | 
94  | 
        timings.map({ case (name, (ts, _)) => (name, ts) }).toMap.withDefaultValue(Nil)
 | 
| 51220 | 95  | 
val session_timing =  | 
| 68731 | 96  | 
make_session_timing(sessions_structure,  | 
97  | 
          timings.map({ case (name, (_, t)) => (name, t) }).toMap)
 | 
|
| 
48678
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
98  | 
|
| 75393 | 99  | 
      object Ordering extends scala.math.Ordering[String] {
 | 
| 
48678
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
100  | 
def compare(name1: String, name2: String): Int =  | 
| 73365 | 101  | 
          session_timing(name2) compare session_timing(name1) match {
 | 
| 
48678
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
102  | 
case 0 =>  | 
| 68731 | 103  | 
              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
 | 
104  | 
case 0 => name1 compare name2  | 
| 
48678
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
105  | 
case ord => ord  | 
| 
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
106  | 
}  | 
| 
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
107  | 
case ord => ord  | 
| 
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
108  | 
}  | 
| 
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
109  | 
}  | 
| 
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
110  | 
|
| 
65415
 
8cd54b18b68b
clarified signature: tree structure is not essential;
 
wenzelm 
parents: 
65406 
diff
changeset
 | 
111  | 
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
 | 
112  | 
}  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
113  | 
}  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
114  | 
|
| 65314 | 115  | 
private class Queue(  | 
| 62631 | 116  | 
graph: Graph[String, Sessions.Info],  | 
| 51220 | 117  | 
order: SortedSet[String],  | 
| 75393 | 118  | 
val command_timings: String => List[Properties.T]  | 
119  | 
  ) {
 | 
|
| 
48676
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
120  | 
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
 | 
121  | 
|
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
122  | 
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
 | 
123  | 
|
| 
51227
 
88c96e836ed6
prefer comparison of session timing, if this is known already;
 
wenzelm 
parents: 
51223 
diff
changeset
 | 
124  | 
def - (name: String): Queue =  | 
| 73365 | 125  | 
new Queue(graph.del_node(name), order - name, command_timings)  | 
| 
48676
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
126  | 
|
| 75393 | 127  | 
    def dequeue(skip: String => Boolean): Option[(String, Sessions.Info)] = {
 | 
| 73365 | 128  | 
val it = order.iterator.dropWhile(name => skip(name) || !graph.is_minimal(name))  | 
| 73344 | 129  | 
      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
 | 
130  | 
else None  | 
| 
 
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  | 
}  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
133  | 
|
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
134  | 
|
| 62631 | 135  | 
|
| 62641 | 136  | 
/** build with results **/  | 
| 48424 | 137  | 
|
| 75393 | 138  | 
  class Results private[Build](results: Map[String, (Option[Process_Result], Sessions.Info)]) {
 | 
| 62403 | 139  | 
def sessions: Set[String] = results.keySet  | 
| 
72673
 
8ff7a0e394f9
clarified PDF/HTML presentation, based on pdf blobs from session database (e.g. from earlier builds);
 
wenzelm 
parents: 
72662 
diff
changeset
 | 
140  | 
def infos: List[Sessions.Info] = results.values.map(_._2).toList  | 
| 
63082
 
6af03422535a
expose Sessions.Info in Build.Results
 
Lars Hupel <lars.hupel@mytum.de> 
parents: 
62946 
diff
changeset
 | 
141  | 
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
 | 
142  | 
def apply(name: String): Process_Result = results(name)._1.getOrElse(Process_Result(1))  | 
| 
74701
 
2bc24136bdeb
clarified order: prefer bottom-up construction of partial content;
 
wenzelm 
parents: 
74700 
diff
changeset
 | 
143  | 
def get_info(name: String): Option[Sessions.Info] = results.get(name).map(_._2)  | 
| 
 
2bc24136bdeb
clarified order: prefer bottom-up construction of partial content;
 
wenzelm 
parents: 
74700 
diff
changeset
 | 
144  | 
def info(name: String): Sessions.Info = get_info(name).get  | 
| 71601 | 145  | 
val rc: Int =  | 
| 73359 | 146  | 
      results.iterator.map({ case (_, (Some(r), _)) => r.rc case (_, (None, _)) => 1 }).
 | 
| 74306 | 147  | 
foldLeft(Process_Result.RC.ok)(_ max _)  | 
148  | 
def ok: Boolean = rc == Process_Result.RC.ok  | 
|
| 62406 | 149  | 
|
150  | 
override def toString: String = rc.toString  | 
|
| 62403 | 151  | 
}  | 
152  | 
||
| 72021 | 153  | 
def session_finished(session_name: String, process_result: Process_Result): String =  | 
154  | 
    "Finished " + session_name + " (" + process_result.timing.message_resources + ")"
 | 
|
| 72013 | 155  | 
|
| 75393 | 156  | 
  def session_timing(session_name: String, build_log: Build_Log.Session_Info): String = {
 | 
| 72021 | 157  | 
val props = build_log.session_timing  | 
158  | 
val threads = Markup.Session_Timing.Threads.unapply(props) getOrElse 1  | 
|
| 74782 | 159  | 
val timing = Markup.Timing_Properties.get(props)  | 
| 72015 | 160  | 
    "Timing " + session_name + " (" + threads + " threads, " + timing.message_factor + ")"
 | 
| 72021 | 161  | 
}  | 
| 72018 | 162  | 
|
| 62641 | 163  | 
def build(  | 
| 
50404
 
898cac1dad5e
avoid startup within GUI thread -- it is only required later for dialog;
 
wenzelm 
parents: 
50367 
diff
changeset
 | 
164  | 
options: Options,  | 
| 71981 | 165  | 
selection: Sessions.Selection = Sessions.Selection.empty,  | 
| 72652 | 166  | 
presentation: Presentation.Context = Presentation.Context.none,  | 
| 
71726
 
a5fda30edae2
clarified signature: more uniform treatment of stopped/interrupted state;
 
wenzelm 
parents: 
71718 
diff
changeset
 | 
167  | 
progress: Progress = new Progress,  | 
| 
65832
 
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
 
wenzelm 
parents: 
65831 
diff
changeset
 | 
168  | 
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
 | 
169  | 
build_heap: Boolean = false,  | 
| 48595 | 170  | 
clean_build: Boolean = false,  | 
| 56890 | 171  | 
dirs: List[Path] = Nil,  | 
172  | 
select_dirs: List[Path] = Nil,  | 
|
| 
66968
 
9991671c98aa
allow to augment session context via explicit session infos;
 
wenzelm 
parents: 
66962 
diff
changeset
 | 
173  | 
infos: List[Sessions.Info] = Nil,  | 
| 64265 | 174  | 
numa_shuffling: Boolean = false,  | 
| 48509 | 175  | 
max_jobs: Int = 1,  | 
| 48903 | 176  | 
list_files: Boolean = false,  | 
| 
59891
 
9ce697050455
added isabelle build option -k, for fast off-line checking of theory sources;
 
wenzelm 
parents: 
59811 
diff
changeset
 | 
177  | 
check_keywords: Set[String] = Set.empty,  | 
| 66841 | 178  | 
fresh_build: Boolean = false,  | 
| 48509 | 179  | 
no_build: Boolean = false,  | 
| 66745 | 180  | 
soft_build: Boolean = false,  | 
| 48509 | 181  | 
verbose: Boolean = false,  | 
| 
73805
 
b73777a0c076
allow build session setup, e.g. for protocol handlers;
 
wenzelm 
parents: 
73804 
diff
changeset
 | 
182  | 
export_files: Boolean = false,  | 
| 75393 | 183  | 
session_setup: (String, Session) => Unit = (_, _) => ()  | 
184  | 
  ): Results = {
 | 
|
| 71677 | 185  | 
val build_options =  | 
| 
71679
 
eeaa4021f080
proper "editor_tracing_messages=0" as in "isabelle dump";
 
wenzelm 
parents: 
71677 
diff
changeset
 | 
186  | 
options +  | 
| 
 
eeaa4021f080
proper "editor_tracing_messages=0" as in "isabelle dump";
 
wenzelm 
parents: 
71677 
diff
changeset
 | 
187  | 
"completion_limit=0" +  | 
| 
 
eeaa4021f080
proper "editor_tracing_messages=0" as in "isabelle dump";
 
wenzelm 
parents: 
71677 
diff
changeset
 | 
188  | 
"editor_tracing_messages=0" +  | 
| 
72728
 
caa182bdab7a
clarified options: batch-build has pide_reports disabled by default (requires significant resources);
 
wenzelm 
parents: 
72693 
diff
changeset
 | 
189  | 
        ("pide_reports=" + options.bool("build_pide_reports"))
 | 
| 51220 | 190  | 
|
| 
69854
 
cc0b3e177b49
system option "system_heaps" supersedes various command-line options for "system build mode";
 
wenzelm 
parents: 
69811 
diff
changeset
 | 
191  | 
val store = Sessions.store(build_options)  | 
| 66745 | 192  | 
|
| 
69369
 
6ecc85955e04
prefer "Isabelle DejaVu Sans", even for headless batch-build (session_graph.pdf);
 
wenzelm 
parents: 
68957 
diff
changeset
 | 
193  | 
Isabelle_Fonts.init()  | 
| 
 
6ecc85955e04
prefer "Isabelle DejaVu Sans", even for headless batch-build (session_graph.pdf);
 
wenzelm 
parents: 
68957 
diff
changeset
 | 
194  | 
|
| 66745 | 195  | 
|
196  | 
/* session selection and dependencies */  | 
|
| 65422 | 197  | 
|
| 66961 | 198  | 
val full_sessions =  | 
| 67026 | 199  | 
Sessions.load_structure(build_options, dirs = dirs, select_dirs = select_dirs, infos = infos)  | 
| 65422 | 200  | 
|
| 
73012
 
238ddf525da4
clarified HTML presentation, e.g. avoid bulky jobs like HOL or HOL-Analysis in applications;
 
wenzelm 
parents: 
72993 
diff
changeset
 | 
201  | 
val full_sessions_selection = full_sessions.imports_selection(selection)  | 
| 74808 | 202  | 
val full_sessions_selected = full_sessions_selection.toSet  | 
| 
73012
 
238ddf525da4
clarified HTML presentation, e.g. avoid bulky jobs like HOL or HOL-Analysis in applications;
 
wenzelm 
parents: 
72993 
diff
changeset
 | 
203  | 
|
| 75393 | 204  | 
    def sources_stamp(deps: Sessions.Deps, session_name: String): String = {
 | 
| 
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
 | 
205  | 
val digests =  | 
| 71669 | 206  | 
full_sessions(session_name).meta_digest ::  | 
| 75741 | 207  | 
deps.session_sources(session_name) :::  | 
| 71669 | 208  | 
deps.imported_sources(session_name)  | 
| 72654 | 209  | 
SHA1.digest_set(digests).toString  | 
| 
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
 | 
210  | 
}  | 
| 66745 | 211  | 
|
| 75393 | 212  | 
    val deps = {
 | 
| 66745 | 213  | 
val deps0 =  | 
| 
71896
 
ce06d6456cc8
clarified signature --- fit within limit of 22 arguments;
 
wenzelm 
parents: 
71895 
diff
changeset
 | 
214  | 
Sessions.deps(full_sessions.selection(selection),  | 
| 
66962
 
e1bde71bace6
clarified signature: global_theories is always required;
 
wenzelm 
parents: 
66961 
diff
changeset
 | 
215  | 
progress = progress, inlined_files = true, verbose = verbose,  | 
| 
 
e1bde71bace6
clarified signature: global_theories is always required;
 
wenzelm 
parents: 
66961 
diff
changeset
 | 
216  | 
list_files = list_files, check_keywords = check_keywords).check_errors  | 
| 66745 | 217  | 
|
| 66841 | 218  | 
      if (soft_build && !fresh_build) {
 | 
| 66745 | 219  | 
val outdated =  | 
| 68204 | 220  | 
deps0.sessions_structure.build_topological_order.flatMap(name =>  | 
| 
72634
 
5cea0993ee4f
clarified access to single database server vs. collection of database files;
 
wenzelm 
parents: 
72624 
diff
changeset
 | 
221  | 
            store.try_open_database(name) match {
 | 
| 68214 | 222  | 
case Some(db) =>  | 
| 68216 | 223  | 
                using(db)(store.read_build(_, name)) match {
 | 
| 66745 | 224  | 
case Some(build)  | 
| 66747 | 225  | 
if build.ok && build.sources == sources_stamp(deps0, name) => None  | 
| 66745 | 226  | 
case _ => Some(name)  | 
227  | 
}  | 
|
228  | 
case None => Some(name)  | 
|
229  | 
})  | 
|
| 68204 | 230  | 
|
231  | 
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
 | 
232  | 
progress = progress, inlined_files = true).check_errors  | 
| 66745 | 233  | 
}  | 
| 68204 | 234  | 
else deps0  | 
| 66745 | 235  | 
}  | 
236  | 
||
| 74800 | 237  | 
val presentation_sessions =  | 
238  | 
      (for {
 | 
|
| 
74813
 
2ad892ac749a
present only selected session theories (as in Isabelle2021), in contrast to 2bc24136bdeb, eb89b3a37826;
 
wenzelm 
parents: 
74808 
diff
changeset
 | 
239  | 
session_name <- deps.sessions_structure.build_topological_order.iterator  | 
| 74808 | 240  | 
info <- deps.sessions_structure.get(session_name)  | 
241  | 
if full_sessions_selected(session_name) && presentation.enabled(info) }  | 
|
| 74800 | 242  | 
yield info).toList  | 
243  | 
||
| 66745 | 244  | 
|
245  | 
/* check unknown files */  | 
|
| 
48504
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
246  | 
|
| 
65832
 
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
 
wenzelm 
parents: 
65831 
diff
changeset
 | 
247  | 
    if (check_unknown_files) {
 | 
| 
 
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
 
wenzelm 
parents: 
65831 
diff
changeset
 | 
248  | 
val source_files =  | 
| 
 
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
 
wenzelm 
parents: 
65831 
diff
changeset
 | 
249  | 
        (for {
 | 
| 
 
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
 
wenzelm 
parents: 
65831 
diff
changeset
 | 
250  | 
(_, base) <- deps.session_bases.iterator  | 
| 75741 | 251  | 
(path, _) <- base.session_sources.iterator  | 
| 
65832
 
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
 
wenzelm 
parents: 
65831 
diff
changeset
 | 
252  | 
} yield path).toList  | 
| 67098 | 253  | 
      val exclude_files = List(Path.explode("$POLYML_EXE")).map(_.canonical_file)
 | 
254  | 
val unknown_files =  | 
|
| 
67782
 
7e223a05e6d8
clarified notion of unknown files: ignore files outside of a Mercurial repository;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
255  | 
Mercurial.check_files(source_files)._2.  | 
| 67098 | 256  | 
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
 | 
257  | 
      if (unknown_files.nonEmpty) {
 | 
| 
67782
 
7e223a05e6d8
clarified notion of unknown files: ignore files outside of a Mercurial repository;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
258  | 
        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
 | 
259  | 
          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
 | 
260  | 
}  | 
| 
 
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
 
wenzelm 
parents: 
65831 
diff
changeset
 | 
261  | 
}  | 
| 
 
2fb85623c386
implicitly check for unknown files (not part of a Mercurial repository);
 
wenzelm 
parents: 
65831 
diff
changeset
 | 
262  | 
|
| 51220 | 263  | 
|
264  | 
/* main build process */  | 
|
265  | 
||
| 68204 | 266  | 
val queue = Queue(progress, deps.sessions_structure, store)  | 
| 65289 | 267  | 
|
| 68219 | 268  | 
store.prepare_output_dir()  | 
| 48373 | 269  | 
|
| 48595 | 270  | 
    if (clean_build) {
 | 
| 
73012
 
238ddf525da4
clarified HTML presentation, e.g. avoid bulky jobs like HOL or HOL-Analysis in applications;
 
wenzelm 
parents: 
72993 
diff
changeset
 | 
271  | 
      for (name <- full_sessions.imports_descendants(full_sessions_selection)) {
 | 
| 
68220
 
8fc4e3d1df86
clarified store.clean_output: cleanup user_output_dir even in system_mode;
 
wenzelm 
parents: 
68219 
diff
changeset
 | 
272  | 
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
 | 
273  | 
        if (relevant) {
 | 
| 
 
8fc4e3d1df86
clarified store.clean_output: cleanup user_output_dir even in system_mode;
 
wenzelm 
parents: 
68219 
diff
changeset
 | 
274  | 
          if (ok) progress.echo("Cleaned " + name)
 | 
| 
 
8fc4e3d1df86
clarified store.clean_output: cleanup user_output_dir even in system_mode;
 
wenzelm 
parents: 
68219 
diff
changeset
 | 
275  | 
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
 | 
276  | 
}  | 
| 48595 | 277  | 
}  | 
278  | 
}  | 
|
279  | 
||
| 48425 | 280  | 
// scheduler loop  | 
| 
63082
 
6af03422535a
expose Sessions.Info in Build.Results
 
Lars Hupel <lars.hupel@mytum.de> 
parents: 
62946 
diff
changeset
 | 
281  | 
case class Result(  | 
| 66594 | 282  | 
current: Boolean,  | 
| 68213 | 283  | 
heap_digest: Option[String],  | 
| 66594 | 284  | 
process: Option[Process_Result],  | 
| 75393 | 285  | 
info: Sessions.Info  | 
286  | 
    ) {
 | 
|
| 62402 | 287  | 
def ok: Boolean =  | 
288  | 
        process match {
 | 
|
289  | 
case None => false  | 
|
| 74306 | 290  | 
case Some(res) => res.ok  | 
| 62402 | 291  | 
}  | 
292  | 
}  | 
|
| 
48639
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
293  | 
|
| 73340 | 294  | 
def sleep(): Unit =  | 
| 
73702
 
7202e12cb324
tuned signature --- following hints by IntelliJ IDEA;
 
wenzelm 
parents: 
73522 
diff
changeset
 | 
295  | 
      Isabelle_Thread.interrupt_handler(_ => progress.stop()) { Time.seconds(0.5).sleep() }
 | 
| 50366 | 296  | 
|
| 
73777
 
52e43a93d51f
clarified system_log: make this work independently of the particular "isabelle build" command-line (e.g. "isabelle mirabelle");
 
wenzelm 
parents: 
73774 
diff
changeset
 | 
297  | 
val log =  | 
| 
 
52e43a93d51f
clarified system_log: make this work independently of the particular "isabelle build" command-line (e.g. "isabelle mirabelle");
 
wenzelm 
parents: 
73774 
diff
changeset
 | 
298  | 
      build_options.string("system_log") match {
 | 
| 
 
52e43a93d51f
clarified system_log: make this work independently of the particular "isabelle build" command-line (e.g. "isabelle mirabelle");
 
wenzelm 
parents: 
73774 
diff
changeset
 | 
299  | 
case "" => No_Logger  | 
| 
74827
 
c1b5d6e6ff74
clarified system option standard values: avoid oddities like "isabelle build -o document_output" producing directories named "true";
 
wenzelm 
parents: 
74818 
diff
changeset
 | 
300  | 
case "-" => Logger.make(progress)  | 
| 
73777
 
52e43a93d51f
clarified system_log: make this work independently of the particular "isabelle build" command-line (e.g. "isabelle mirabelle");
 
wenzelm 
parents: 
73774 
diff
changeset
 | 
301  | 
case log_file => Logger.make(Some(Path.explode(log_file)))  | 
| 
 
52e43a93d51f
clarified system_log: make this work independently of the particular "isabelle build" command-line (e.g. "isabelle mirabelle");
 
wenzelm 
parents: 
73774 
diff
changeset
 | 
302  | 
}  | 
| 
 
52e43a93d51f
clarified system_log: make this work independently of the particular "isabelle build" command-line (e.g. "isabelle mirabelle");
 
wenzelm 
parents: 
73774 
diff
changeset
 | 
303  | 
|
| 64265 | 304  | 
val numa_nodes = new NUMA.Nodes(numa_shuffling)  | 
305  | 
||
| 48425 | 306  | 
@tailrec def loop(  | 
| 
48676
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
307  | 
pending: Queue,  | 
| 72662 | 308  | 
running: Map[String, (List[String], Build_Job)],  | 
| 75393 | 309  | 
results: Map[String, Result]  | 
310  | 
    ): Map[String, Result] = {
 | 
|
| 64265 | 311  | 
def used_node(i: Int): Boolean =  | 
312  | 
running.iterator.exists(  | 
|
313  | 
          { case (_, (_, job)) => job.numa_node.isDefined && job.numa_node.get == i })
 | 
|
314  | 
||
| 48425 | 315  | 
if (pending.is_empty) results  | 
| 51253 | 316  | 
      else {
 | 
| 
71726
 
a5fda30edae2
clarified signature: more uniform treatment of stopped/interrupted state;
 
wenzelm 
parents: 
71718 
diff
changeset
 | 
317  | 
        if (progress.stopped) {
 | 
| 73367 | 318  | 
for ((_, (_, job)) <- running) job.terminate()  | 
| 
71726
 
a5fda30edae2
clarified signature: more uniform treatment of stopped/interrupted state;
 
wenzelm 
parents: 
71718 
diff
changeset
 | 
319  | 
}  | 
| 51253 | 320  | 
|
| 48674 | 321  | 
        running.find({ case (_, (_, job)) => job.is_finished }) match {
 | 
| 71669 | 322  | 
case Some((session_name, (input_heaps, job))) =>  | 
| 50367 | 323  | 
            //{{{ finish job
 | 
| 48424 | 324  | 
|
| 68217 | 325  | 
val (process_result, heap_digest) = job.join  | 
| 48373 | 326  | 
|
| 
71623
 
b3bddebe44ca
clarified signature: more explicit type Protocol_Message.Marker;
 
wenzelm 
parents: 
71614 
diff
changeset
 | 
327  | 
val log_lines = process_result.out_lines.filterNot(Protocol_Message.Marker.test)  | 
| 75393 | 328  | 
            val process_result_tail = {
 | 
| 
62409
 
e391528eff3b
proper option process_output_tail, more generous default;
 
wenzelm 
parents: 
62406 
diff
changeset
 | 
329  | 
              val tail = job.info.options.int("process_output_tail")
 | 
| 62632 | 330  | 
process_result.copy(  | 
331  | 
out_lines =  | 
|
| 71669 | 332  | 
"(see also " + store.output_log(session_name).file.toString + ")" ::  | 
| 65294 | 333  | 
(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
 | 
334  | 
}  | 
| 
 
13a0f537e232
retain tail out_lines as printed, but not the whole log content;
 
wenzelm 
parents: 
62403 
diff
changeset
 | 
335  | 
|
| 72017 | 336  | 
val build_log =  | 
337  | 
Build_Log.Log_File(session_name, process_result.out_lines).  | 
|
338  | 
parse_session_info(  | 
|
339  | 
command_timings = true,  | 
|
340  | 
theory_timings = true,  | 
|
341  | 
ml_statistics = true,  | 
|
342  | 
task_statistics = true)  | 
|
| 
66666
 
1a620647285c
clarified messages: after writing all files (see also 27f90319a499 and 57c85c83c11b);
 
wenzelm 
parents: 
66595 
diff
changeset
 | 
343  | 
|
| 
 
1a620647285c
clarified messages: after writing all files (see also 27f90319a499 and 57c85c83c11b);
 
wenzelm 
parents: 
66595 
diff
changeset
 | 
344  | 
// write log file  | 
| 68217 | 345  | 
            if (process_result.ok) {
 | 
| 71669 | 346  | 
File.write_gzip(store.output_log_gz(session_name), terminate_lines(log_lines))  | 
| 68217 | 347  | 
}  | 
| 71669 | 348  | 
else File.write(store.output_log(session_name), terminate_lines(log_lines))  | 
| 
65291
 
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
 
wenzelm 
parents: 
65289 
diff
changeset
 | 
349  | 
|
| 
 
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
 
wenzelm 
parents: 
65289 
diff
changeset
 | 
350  | 
// write database  | 
| 72017 | 351  | 
using(store.open_database(session_name, output = true))(db =>  | 
352  | 
store.write_session_info(db, session_name,  | 
|
353  | 
build_log =  | 
|
354  | 
                  if (process_result.timeout) build_log.error("Timeout") else build_log,
 | 
|
355  | 
build =  | 
|
356  | 
Session_Info(sources_stamp(deps, session_name), input_heaps, heap_digest,  | 
|
357  | 
process_result.rc)))  | 
|
| 
65291
 
57c85c83c11b
maintain persistent session info in SQLite database instead of log file;
 
wenzelm 
parents: 
65289 
diff
changeset
 | 
358  | 
|
| 
66666
 
1a620647285c
clarified messages: after writing all files (see also 27f90319a499 and 57c85c83c11b);
 
wenzelm 
parents: 
66595 
diff
changeset
 | 
359  | 
// messages  | 
| 71601 | 360  | 
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
 | 
361  | 
|
| 72013 | 362  | 
            if (process_result.ok) {
 | 
| 
72022
 
45865bb06182
clarified message --- as in former ML version (see 940195fbb282);
 
wenzelm 
parents: 
72021 
diff
changeset
 | 
363  | 
if (verbose) progress.echo(session_timing(session_name, build_log))  | 
| 72021 | 364  | 
progress.echo(session_finished(session_name, process_result))  | 
| 72013 | 365  | 
}  | 
| 
66666
 
1a620647285c
clarified messages: after writing all files (see also 27f90319a499 and 57c85c83c11b);
 
wenzelm 
parents: 
66595 
diff
changeset
 | 
366  | 
            else {
 | 
| 71669 | 367  | 
progress.echo(session_name + " FAILED")  | 
| 
66666
 
1a620647285c
clarified messages: after writing all files (see also 27f90319a499 and 57c85c83c11b);
 
wenzelm 
parents: 
66595 
diff
changeset
 | 
368  | 
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
 | 
369  | 
}  | 
| 
 
1a620647285c
clarified messages: after writing all files (see also 27f90319a499 and 57c85c83c11b);
 
wenzelm 
parents: 
66595 
diff
changeset
 | 
370  | 
|
| 71669 | 371  | 
loop(pending - session_name, running - session_name,  | 
372  | 
results +  | 
|
373  | 
(session_name -> Result(false, heap_digest, Some(process_result_tail), job.info)))  | 
|
| 50367 | 374  | 
//}}}  | 
| 60215 | 375  | 
case None if running.size < (max_jobs max 1) =>  | 
| 50367 | 376  | 
            //{{{ check/start next job
 | 
| 71601 | 377  | 
            pending.dequeue(running.isDefinedAt) match {
 | 
| 71669 | 378  | 
case Some((session_name, info)) =>  | 
| 66848 | 379  | 
val ancestor_results =  | 
| 71669 | 380  | 
deps.sessions_structure.build_requirements(List(session_name)).  | 
381  | 
filterNot(_ == session_name).map(results(_))  | 
|
| 68213 | 382  | 
val ancestor_heaps = ancestor_results.flatMap(_.heap_digest)  | 
| 62628 | 383  | 
|
| 71669 | 384  | 
val do_store =  | 
385  | 
build_heap || Sessions.is_pure(session_name) || queue.is_inner(session_name)  | 
|
| 48547 | 386  | 
|
| 75393 | 387  | 
                val (current, heap_digest) = {
 | 
| 
72634
 
5cea0993ee4f
clarified access to single database server vs. collection of database files;
 
wenzelm 
parents: 
72624 
diff
changeset
 | 
388  | 
                  store.try_open_database(session_name) match {
 | 
| 
68212
 
5a59fded83c7
clarified heap vs. database operations: discontinued correlation of directory;
 
wenzelm 
parents: 
68209 
diff
changeset
 | 
389  | 
case Some(db) =>  | 
| 71669 | 390  | 
                      using(db)(store.read_build(_, session_name)) match {
 | 
| 68216 | 391  | 
case Some(build) =>  | 
| 71669 | 392  | 
val heap_digest = store.find_heap_digest(session_name)  | 
| 68216 | 393  | 
val current =  | 
394  | 
!fresh_build &&  | 
|
395  | 
build.ok &&  | 
|
| 71669 | 396  | 
build.sources == sources_stamp(deps, session_name) &&  | 
| 68216 | 397  | 
build.input_heaps == ancestor_heaps &&  | 
398  | 
build.output_heap == heap_digest &&  | 
|
| 71614 | 399  | 
!(do_store && heap_digest.isEmpty)  | 
| 68216 | 400  | 
(current, heap_digest)  | 
401  | 
case None => (false, None)  | 
|
402  | 
}  | 
|
| 62636 | 403  | 
case None => (false, None)  | 
| 
48504
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
404  | 
}  | 
| 48547 | 405  | 
}  | 
| 62628 | 406  | 
val all_current = current && ancestor_results.forall(_.current)  | 
| 
48528
 
784c6f63d79c
proper all_current, which regards parent status as well;
 
wenzelm 
parents: 
48511 
diff
changeset
 | 
407  | 
|
| 48547 | 408  | 
if (all_current)  | 
| 71669 | 409  | 
loop(pending - session_name, running,  | 
410  | 
results +  | 
|
411  | 
(session_name -> Result(true, heap_digest, Some(Process_Result(0)), info)))  | 
|
| 
48678
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
412  | 
                else if (no_build) {
 | 
| 71894 | 413  | 
progress.echo_if(verbose, "Skipping " + session_name + " ...")  | 
| 71669 | 414  | 
loop(pending - session_name, running,  | 
415  | 
results +  | 
|
416  | 
(session_name -> Result(false, heap_digest, Some(Process_Result(1)), info)))  | 
|
| 
48678
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
417  | 
}  | 
| 62628 | 418  | 
                else if (ancestor_results.forall(_.ok) && !progress.stopped) {
 | 
| 71669 | 419  | 
progress.echo((if (do_store) "Building " else "Running ") + session_name + " ...")  | 
| 
68086
 
9e1c670301b8
cleanup session output before starting build job;
 
wenzelm 
parents: 
67852 
diff
changeset
 | 
420  | 
|
| 71669 | 421  | 
store.clean_output(session_name)  | 
422  | 
using(store.open_database(session_name, output = true))(  | 
|
423  | 
store.init_session_info(_, session_name))  | 
|
| 
68086
 
9e1c670301b8
cleanup session output before starting build job;
 
wenzelm 
parents: 
67852 
diff
changeset
 | 
424  | 
|
| 71601 | 425  | 
val numa_node = numa_nodes.next(used_node)  | 
| 51220 | 426  | 
val job =  | 
| 72693 | 427  | 
new Build_Job(progress, session_name, info, deps, store, do_store,  | 
| 
73805
 
b73777a0c076
allow build session setup, e.g. for protocol handlers;
 
wenzelm 
parents: 
73804 
diff
changeset
 | 
428  | 
log, session_setup, numa_node, queue.command_timings(session_name))  | 
| 71669 | 429  | 
loop(pending, running + (session_name -> (ancestor_heaps, job)), results)  | 
| 48547 | 430  | 
}  | 
431  | 
                else {
 | 
|
| 71669 | 432  | 
progress.echo(session_name + " CANCELLED")  | 
433  | 
loop(pending - session_name, running,  | 
|
434  | 
results + (session_name -> Result(false, heap_digest, None, info)))  | 
|
| 48547 | 435  | 
}  | 
436  | 
case None => sleep(); loop(pending, running, results)  | 
|
| 48425 | 437  | 
}  | 
| 50367 | 438  | 
///}}}  | 
| 48425 | 439  | 
case None => sleep(); loop(pending, running, results)  | 
| 48373 | 440  | 
}  | 
| 51253 | 441  | 
}  | 
| 48425 | 442  | 
}  | 
443  | 
||
| 51220 | 444  | 
|
445  | 
/* build results */  | 
|
446  | 
||
| 75393 | 447  | 
    val results = {
 | 
| 
72673
 
8ff7a0e394f9
clarified PDF/HTML presentation, based on pdf blobs from session database (e.g. from earlier builds);
 
wenzelm 
parents: 
72662 
diff
changeset
 | 
448  | 
val results0 =  | 
| 
 
8ff7a0e394f9
clarified PDF/HTML presentation, based on pdf blobs from session database (e.g. from earlier builds);
 
wenzelm 
parents: 
72662 
diff
changeset
 | 
449  | 
        if (deps.is_empty) {
 | 
| 
 
8ff7a0e394f9
clarified PDF/HTML presentation, based on pdf blobs from session database (e.g. from earlier builds);
 
wenzelm 
parents: 
72662 
diff
changeset
 | 
450  | 
          progress.echo_warning("Nothing to build")
 | 
| 
 
8ff7a0e394f9
clarified PDF/HTML presentation, based on pdf blobs from session database (e.g. from earlier builds);
 
wenzelm 
parents: 
72662 
diff
changeset
 | 
451  | 
Map.empty[String, Result]  | 
| 
 
8ff7a0e394f9
clarified PDF/HTML presentation, based on pdf blobs from session database (e.g. from earlier builds);
 
wenzelm 
parents: 
72662 
diff
changeset
 | 
452  | 
}  | 
| 
 
8ff7a0e394f9
clarified PDF/HTML presentation, based on pdf blobs from session database (e.g. from earlier builds);
 
wenzelm 
parents: 
72662 
diff
changeset
 | 
453  | 
        else Isabelle_Thread.uninterruptible { loop(queue, Map.empty, Map.empty) }
 | 
| 48583 | 454  | 
|
| 64265 | 455  | 
new Results(  | 
456  | 
(for ((name, result) <- results0.iterator)  | 
|
457  | 
yield (name, (result.process, result.info))).toMap)  | 
|
| 
72673
 
8ff7a0e394f9
clarified PDF/HTML presentation, based on pdf blobs from session database (e.g. from earlier builds);
 
wenzelm 
parents: 
72662 
diff
changeset
 | 
458  | 
}  | 
| 62641 | 459  | 
|
| 
69811
 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 
wenzelm 
parents: 
69777 
diff
changeset
 | 
460  | 
    if (export_files) {
 | 
| 
71896
 
ce06d6456cc8
clarified signature --- fit within limit of 22 arguments;
 
wenzelm 
parents: 
71895 
diff
changeset
 | 
461  | 
      for (name <- full_sessions.imports_selection(selection).iterator if results(name).ok) {
 | 
| 
69811
 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 
wenzelm 
parents: 
69777 
diff
changeset
 | 
462  | 
val info = results.info(name)  | 
| 
 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 
wenzelm 
parents: 
69777 
diff
changeset
 | 
463  | 
        if (info.export_files.nonEmpty) {
 | 
| 
 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 
wenzelm 
parents: 
69777 
diff
changeset
 | 
464  | 
          progress.echo("Exporting " + info.name + " ...")
 | 
| 
 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 
wenzelm 
parents: 
69777 
diff
changeset
 | 
465  | 
          for ((dir, prune, pats) <- info.export_files) {
 | 
| 
 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 
wenzelm 
parents: 
69777 
diff
changeset
 | 
466  | 
Export.export_files(store, name, info.dir + dir,  | 
| 
71726
 
a5fda30edae2
clarified signature: more uniform treatment of stopped/interrupted state;
 
wenzelm 
parents: 
71718 
diff
changeset
 | 
467  | 
progress = if (verbose) progress else new Progress,  | 
| 
69811
 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 
wenzelm 
parents: 
69777 
diff
changeset
 | 
468  | 
export_prune = prune,  | 
| 
 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 
wenzelm 
parents: 
69777 
diff
changeset
 | 
469  | 
export_patterns = pats)  | 
| 
 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 
wenzelm 
parents: 
69777 
diff
changeset
 | 
470  | 
}  | 
| 
 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 
wenzelm 
parents: 
69777 
diff
changeset
 | 
471  | 
}  | 
| 
 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 
wenzelm 
parents: 
69777 
diff
changeset
 | 
472  | 
}  | 
| 
 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 
wenzelm 
parents: 
69777 
diff
changeset
 | 
473  | 
}  | 
| 
 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 
wenzelm 
parents: 
69777 
diff
changeset
 | 
474  | 
|
| 74306 | 475  | 
    if (results.rc != Process_Result.RC.ok && (verbose || !no_build)) {
 | 
| 62641 | 476  | 
val unfinished =  | 
477  | 
        (for {
 | 
|
478  | 
name <- results.sessions.iterator  | 
|
479  | 
if !results(name).ok  | 
|
480  | 
} yield name).toList.sorted  | 
|
481  | 
      progress.echo("Unfinished session(s): " + commas(unfinished))
 | 
|
482  | 
}  | 
|
483  | 
||
| 
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
 | 
484  | 
|
| 
72673
 
8ff7a0e394f9
clarified PDF/HTML presentation, based on pdf blobs from session database (e.g. from earlier builds);
 
wenzelm 
parents: 
72662 
diff
changeset
 | 
485  | 
/* PDF/HTML presentation */  | 
| 
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
 | 
486  | 
|
| 72677 | 487  | 
    if (!no_build && !progress.stopped && results.ok) {
 | 
| 74698 | 488  | 
      if (presentation_sessions.nonEmpty) {
 | 
| 72676 | 489  | 
val presentation_dir = presentation.dir(store)  | 
| 74709 | 490  | 
Presentation.update_root(presentation_dir)  | 
| 
74701
 
2bc24136bdeb
clarified order: prefer bottom-up construction of partial content;
 
wenzelm 
parents: 
74700 
diff
changeset
 | 
491  | 
|
| 
 
2bc24136bdeb
clarified order: prefer bottom-up construction of partial content;
 
wenzelm 
parents: 
74700 
diff
changeset
 | 
492  | 
        for ((chapter, infos) <- presentation_sessions.groupBy(_.chapter).iterator) {
 | 
| 
 
2bc24136bdeb
clarified order: prefer bottom-up construction of partial content;
 
wenzelm 
parents: 
74700 
diff
changeset
 | 
493  | 
val entries = infos.map(info => (info.name, info.description))  | 
| 74709 | 494  | 
Presentation.update_chapter(presentation_dir, chapter, entries)  | 
| 
74701
 
2bc24136bdeb
clarified order: prefer bottom-up construction of partial content;
 
wenzelm 
parents: 
74700 
diff
changeset
 | 
495  | 
}  | 
| 72676 | 496  | 
|
| 75786 | 497  | 
        using(Export.open_database_context(store)) { database_context =>
 | 
| 75907 | 498  | 
val document_info =  | 
499  | 
Document_Info.read(database_context, deps, presentation_sessions.map(_.name))  | 
|
| 
74798
 
507f50dbeb79
just one Presentation.State for all sessions: avoid duplication of already presented theories (very slow) and cached theory export (not very slow);
 
wenzelm 
parents: 
74782 
diff
changeset
 | 
500  | 
|
| 75394 | 501  | 
          Par_List.map({ (session: String) =>
 | 
| 72683 | 502  | 
progress.expose_interrupt()  | 
| 
74770
 
32c2587cda4f
clarified HTML_Context: more explicit directory structure;
 
wenzelm 
parents: 
74767 
diff
changeset
 | 
503  | 
|
| 
 
32c2587cda4f
clarified HTML_Context: more explicit directory structure;
 
wenzelm 
parents: 
74767 
diff
changeset
 | 
504  | 
val html_context =  | 
| 75897 | 505  | 
Presentation.html_context(deps.sessions_structure, Presentation.elements1,  | 
| 75907 | 506  | 
root_dir = presentation_dir, document_info = document_info)  | 
| 
75778
 
d18c96b9b955
prefer Export.Context/Session_Context/Theory_Context over Sessions.Database_Context;
 
wenzelm 
parents: 
75741 
diff
changeset
 | 
507  | 
|
| 75786 | 508  | 
            using(database_context.open_session(deps.base_info(session))) { session_context =>
 | 
| 75897 | 509  | 
Presentation.session_html(html_context, session_context,  | 
| 
75887
 
e5c0116a5c9f
clarified signature: replaced Sessions.Deps by Sessions.Structure from HTML_Context;
 
wenzelm 
parents: 
75884 
diff
changeset
 | 
510  | 
progress = progress, verbose = verbose)  | 
| 
75778
 
d18c96b9b955
prefer Export.Context/Session_Context/Theory_Context over Sessions.Database_Context;
 
wenzelm 
parents: 
75741 
diff
changeset
 | 
511  | 
}  | 
| 
74818
 
3064e165c660
clarified HTML_Context.theory_exports: prefer value-oriented parallelism;
 
wenzelm 
parents: 
74815 
diff
changeset
 | 
512  | 
}, presentation_sessions.map(_.name))  | 
| 75394 | 513  | 
}  | 
| 72676 | 514  | 
}  | 
| 
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
 | 
515  | 
}  | 
| 
 
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
 | 
516  | 
|
| 62641 | 517  | 
results  | 
| 48341 | 518  | 
}  | 
519  | 
||
520  | 
||
| 62833 | 521  | 
/* Isabelle tool wrapper */  | 
| 48341 | 522  | 
|
| 72763 | 523  | 
  val isabelle_tool = Isabelle_Tool("build", "build and manage Isabelle sessions",
 | 
| 75394 | 524  | 
Scala_Project.here,  | 
525  | 
    { args =>
 | 
|
526  | 
      val build_options = Word.explode(Isabelle_System.getenv("ISABELLE_BUILD_OPTIONS"))
 | 
|
| 62590 | 527  | 
|
| 75394 | 528  | 
var base_sessions: List[String] = Nil  | 
529  | 
var select_dirs: List[Path] = Nil  | 
|
530  | 
var numa_shuffling = false  | 
|
531  | 
var presentation = Presentation.Context.none  | 
|
532  | 
var requirements = false  | 
|
533  | 
var soft_build = false  | 
|
534  | 
var exclude_session_groups: List[String] = Nil  | 
|
535  | 
var all_sessions = false  | 
|
536  | 
var build_heap = false  | 
|
537  | 
var clean_build = false  | 
|
538  | 
var dirs: List[Path] = Nil  | 
|
539  | 
var export_files = false  | 
|
540  | 
var fresh_build = false  | 
|
541  | 
var session_groups: List[String] = Nil  | 
|
542  | 
var max_jobs = 1  | 
|
543  | 
var check_keywords: Set[String] = Set.empty  | 
|
544  | 
var list_files = false  | 
|
545  | 
var no_build = false  | 
|
546  | 
var options = Options.init(opts = build_options)  | 
|
547  | 
var verbose = false  | 
|
548  | 
var exclude_sessions: List[String] = Nil  | 
|
| 62590 | 549  | 
|
| 75394 | 550  | 
      val getopts = Getopts("""
 | 
| 62590 | 551  | 
Usage: isabelle build [OPTIONS] [SESSIONS ...]  | 
552  | 
||
553  | 
Options are:  | 
|
| 
66737
 
2edc0c42c883
option -B for "isabelle build" and "isabelle imports";
 
wenzelm 
parents: 
66736 
diff
changeset
 | 
554  | 
-B NAME include session NAME and all descendants  | 
| 62590 | 555  | 
-D DIR include session directory and select its sessions  | 
| 64265 | 556  | 
-N cyclic shuffling of NUMA CPU nodes (performance tuning)  | 
| 72648 | 557  | 
    -P DIR       enable HTML/PDF presentation in directory (":" for default)
 | 
| 71807 | 558  | 
-R refer to requirements of selected sessions  | 
| 66745 | 559  | 
-S soft build: only observe changes of sources, not heap images  | 
| 62590 | 560  | 
-X NAME exclude sessions from group NAME and all descendants  | 
561  | 
-a select all sessions  | 
|
562  | 
-b build heap images  | 
|
563  | 
-c clean build  | 
|
564  | 
-d DIR include session directory  | 
|
| 
69811
 
18f61ce86425
clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
 
wenzelm 
parents: 
69777 
diff
changeset
 | 
565  | 
-e export files from session specification into file-system  | 
| 66841 | 566  | 
-f fresh build  | 
| 62590 | 567  | 
-g NAME select session group NAME  | 
568  | 
-j INT maximum number of parallel jobs (default 1)  | 
|
569  | 
-k KEYWORD check theory sources for conflicts with proposed keywords  | 
|
570  | 
-l list session source files  | 
|
571  | 
-n no build -- test dependencies only  | 
|
572  | 
-o OPTION override Isabelle system OPTION (via NAME=VAL or NAME)  | 
|
573  | 
-v verbose  | 
|
574  | 
-x NAME exclude session NAME and all descendants  | 
|
575  | 
||
| 62596 | 576  | 
Build and manage Isabelle sessions, depending on implicit settings:  | 
577  | 
||
| 73736 | 578  | 
""" + Library.indent_lines(2, Build_Log.Settings.show()) + "\n",  | 
| 75394 | 579  | 
"B:" -> (arg => base_sessions = base_sessions ::: List(arg)),  | 
580  | 
"D:" -> (arg => select_dirs = select_dirs ::: List(Path.explode(arg))),  | 
|
581  | 
"N" -> (_ => numa_shuffling = true),  | 
|
582  | 
"P:" -> (arg => presentation = Presentation.Context.make(arg)),  | 
|
583  | 
"R" -> (_ => requirements = true),  | 
|
584  | 
"S" -> (_ => soft_build = true),  | 
|
585  | 
"X:" -> (arg => exclude_session_groups = exclude_session_groups ::: List(arg)),  | 
|
586  | 
"a" -> (_ => all_sessions = true),  | 
|
587  | 
"b" -> (_ => build_heap = true),  | 
|
588  | 
"c" -> (_ => clean_build = true),  | 
|
589  | 
"d:" -> (arg => dirs = dirs ::: List(Path.explode(arg))),  | 
|
590  | 
"e" -> (_ => export_files = true),  | 
|
591  | 
"f" -> (_ => fresh_build = true),  | 
|
592  | 
"g:" -> (arg => session_groups = session_groups ::: List(arg)),  | 
|
593  | 
"j:" -> (arg => max_jobs = Value.Int.parse(arg)),  | 
|
594  | 
"k:" -> (arg => check_keywords = check_keywords + arg),  | 
|
595  | 
"l" -> (_ => list_files = true),  | 
|
596  | 
"n" -> (_ => no_build = true),  | 
|
597  | 
"o:" -> (arg => options = options + arg),  | 
|
598  | 
"v" -> (_ => verbose = true),  | 
|
599  | 
"x:" -> (arg => exclude_sessions = exclude_sessions ::: List(arg)))  | 
|
| 62590 | 600  | 
|
| 75394 | 601  | 
val sessions = getopts(args)  | 
| 62590 | 602  | 
|
| 75394 | 603  | 
val progress = new Console_Progress(verbose = verbose)  | 
| 62590 | 604  | 
|
| 75394 | 605  | 
val start_date = Date.now()  | 
| 64140 | 606  | 
|
| 75394 | 607  | 
      if (verbose) {
 | 
608  | 
progress.echo(  | 
|
609  | 
"Started at " + Build_Log.print_date(start_date) +  | 
|
610  | 
            " (" + Isabelle_System.getenv("ML_IDENTIFIER") + " on " + Isabelle_System.hostname() +")")
 | 
|
611  | 
progress.echo(Build_Log.Settings.show() + "\n")  | 
|
612  | 
}  | 
|
| 62590 | 613  | 
|
| 75394 | 614  | 
val results =  | 
615  | 
        progress.interrupt_handler {
 | 
|
616  | 
build(options,  | 
|
617  | 
selection = Sessions.Selection(  | 
|
618  | 
requirements = requirements,  | 
|
619  | 
all_sessions = all_sessions,  | 
|
620  | 
base_sessions = base_sessions,  | 
|
621  | 
exclude_session_groups = exclude_session_groups,  | 
|
622  | 
exclude_sessions = exclude_sessions,  | 
|
623  | 
session_groups = session_groups,  | 
|
624  | 
sessions = sessions),  | 
|
625  | 
presentation = presentation,  | 
|
626  | 
progress = progress,  | 
|
627  | 
check_unknown_files = Mercurial.is_repository(Path.ISABELLE_HOME),  | 
|
628  | 
build_heap = build_heap,  | 
|
629  | 
clean_build = clean_build,  | 
|
630  | 
dirs = dirs,  | 
|
631  | 
select_dirs = select_dirs,  | 
|
632  | 
numa_shuffling = NUMA.enabled_warning(progress, numa_shuffling),  | 
|
633  | 
max_jobs = max_jobs,  | 
|
634  | 
list_files = list_files,  | 
|
635  | 
check_keywords = check_keywords,  | 
|
636  | 
fresh_build = fresh_build,  | 
|
637  | 
no_build = no_build,  | 
|
638  | 
soft_build = soft_build,  | 
|
639  | 
verbose = verbose,  | 
|
640  | 
export_files = export_files)  | 
|
641  | 
}  | 
|
642  | 
val end_date = Date.now()  | 
|
643  | 
val elapsed_time = end_date.time - start_date.time  | 
|
644  | 
||
645  | 
      if (verbose) {
 | 
|
646  | 
        progress.echo("\nFinished at " + Build_Log.print_date(end_date))
 | 
|
| 62833 | 647  | 
}  | 
| 62590 | 648  | 
|
| 75394 | 649  | 
val total_timing =  | 
650  | 
results.sessions.iterator.map(a => results(a).timing).foldLeft(Timing.zero)(_ + _).  | 
|
651  | 
copy(elapsed = elapsed_time)  | 
|
652  | 
progress.echo(total_timing.message_resources)  | 
|
| 62590 | 653  | 
|
| 75394 | 654  | 
sys.exit(results.rc)  | 
655  | 
})  | 
|
| 68305 | 656  | 
|
657  | 
||
658  | 
/* build logic image */  | 
|
659  | 
||
660  | 
def build_logic(options: Options, logic: String,  | 
|
| 
71726
 
a5fda30edae2
clarified signature: more uniform treatment of stopped/interrupted state;
 
wenzelm 
parents: 
71718 
diff
changeset
 | 
661  | 
progress: Progress = new Progress,  | 
| 68305 | 662  | 
build_heap: Boolean = false,  | 
663  | 
dirs: List[Path] = Nil,  | 
|
| 70791 | 664  | 
fresh: Boolean = false,  | 
| 75393 | 665  | 
strict: Boolean = false  | 
666  | 
  ): Int = {
 | 
|
| 
71896
 
ce06d6456cc8
clarified signature --- fit within limit of 22 arguments;
 
wenzelm 
parents: 
71895 
diff
changeset
 | 
667  | 
val selection = Sessions.Selection.session(logic)  | 
| 69540 | 668  | 
val rc =  | 
| 71981 | 669  | 
if (!fresh && build(options, selection = selection,  | 
| 74306 | 670  | 
build_heap = build_heap, no_build = true, dirs = dirs).ok) Process_Result.RC.ok  | 
| 69540 | 671  | 
      else {
 | 
672  | 
        progress.echo("Build started for Isabelle/" + logic + " ...")
 | 
|
| 71981 | 673  | 
Build.build(options, selection = selection, progress = progress,  | 
674  | 
build_heap = build_heap, fresh_build = fresh, dirs = dirs).rc  | 
|
| 69540 | 675  | 
}  | 
| 74306 | 676  | 
    if (strict && rc != Process_Result.RC.ok) error("Failed to build Isabelle/" + logic) else rc
 | 
| 68305 | 677  | 
}  | 
| 48276 | 678  | 
}  |