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