| author | wenzelm | 
| Sun, 02 Sep 2012 19:26:05 +0200 | |
| changeset 49065 | 8ead9e8b15fb | 
| parent 48992 | 0518bf89c777 | 
| child 49098 | 673e0ed547af | 
| permissions | -rw-r--r-- | 
| 48276 | 1  | 
/* Title: Pure/System/build.scala  | 
2  | 
Author: Makarius  | 
|
3  | 
||
4  | 
Build and manage Isabelle sessions.  | 
|
5  | 
*/  | 
|
6  | 
||
7  | 
package isabelle  | 
|
8  | 
||
9  | 
||
| 48661 | 10  | 
import java.util.{Timer, TimerTask}
 | 
| 48548 | 11  | 
import java.io.{BufferedInputStream, FileInputStream,
 | 
| 48494 | 12  | 
BufferedReader, InputStreamReader, IOException}  | 
13  | 
import java.util.zip.GZIPInputStream  | 
|
| 48335 | 14  | 
|
| 
48676
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
15  | 
import scala.collection.SortedSet  | 
| 
48340
 
6f4fc030882a
allow explicit specification of additional session directories;
 
wenzelm 
parents: 
48339 
diff
changeset
 | 
16  | 
import scala.annotation.tailrec  | 
| 
48337
 
9c7f8e5805b4
cumulate semantic Session_Info, based on syntactic Session_Entry;
 
wenzelm 
parents: 
48336 
diff
changeset
 | 
17  | 
|
| 48335 | 18  | 
|
| 48276 | 19  | 
object Build  | 
20  | 
{
 | 
|
| 
48337
 
9c7f8e5805b4
cumulate semantic Session_Info, based on syntactic Session_Entry;
 
wenzelm 
parents: 
48336 
diff
changeset
 | 
21  | 
/** session information **/  | 
| 48334 | 22  | 
|
| 
48675
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
23  | 
// external version  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
24  | 
sealed case class Session_Entry(  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
25  | 
pos: Position.T,  | 
| 
48738
 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 
wenzelm 
parents: 
48737 
diff
changeset
 | 
26  | 
name: String,  | 
| 
48675
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
27  | 
groups: List[String],  | 
| 
48738
 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 
wenzelm 
parents: 
48737 
diff
changeset
 | 
28  | 
path: String,  | 
| 
48675
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
29  | 
parent: Option[String],  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
30  | 
description: String,  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
31  | 
options: List[Options.Spec],  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
32  | 
theories: List[(List[Options.Spec], List[String])],  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
33  | 
files: List[String])  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
34  | 
|
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
35  | 
// internal version  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
36  | 
sealed case class Session_Info(  | 
| 
48737
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
37  | 
select: Boolean,  | 
| 
48675
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
38  | 
pos: Position.T,  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
39  | 
groups: List[String],  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
40  | 
dir: Path,  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
41  | 
parent: Option[String],  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
42  | 
description: String,  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
43  | 
options: Options,  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
44  | 
theories: List[(Options, List[Path])],  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
45  | 
files: List[Path],  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
46  | 
entry_digest: SHA1.Digest)  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
47  | 
|
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
48  | 
def is_pure(name: String): Boolean = name == "RAW" || name == "Pure"  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
49  | 
|
| 
48737
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
50  | 
def session_info(options: Options, select: Boolean, dir: Path, entry: Session_Entry)  | 
| 
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
51  | 
: (String, Session_Info) =  | 
| 
48675
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
52  | 
    try {
 | 
| 
48738
 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 
wenzelm 
parents: 
48737 
diff
changeset
 | 
53  | 
val name = entry.name  | 
| 
48675
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
54  | 
|
| 
48738
 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 
wenzelm 
parents: 
48737 
diff
changeset
 | 
55  | 
      if (name == "") error("Bad session name")
 | 
| 
 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 
wenzelm 
parents: 
48737 
diff
changeset
 | 
56  | 
      if (is_pure(name) && entry.parent.isDefined) error("Illegal parent session")
 | 
| 
 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 
wenzelm 
parents: 
48737 
diff
changeset
 | 
57  | 
      if (!is_pure(name) && !entry.parent.isDefined) error("Missing parent session")
 | 
| 
48675
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
58  | 
|
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
59  | 
val session_options = options ++ entry.options  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
60  | 
|
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
61  | 
val theories =  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
62  | 
        entry.theories.map({ case (opts, thys) =>
 | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
63  | 
(session_options ++ opts, thys.map(Path.explode(_))) })  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
64  | 
val files = entry.files.map(Path.explode(_))  | 
| 
48738
 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 
wenzelm 
parents: 
48737 
diff
changeset
 | 
65  | 
val entry_digest = SHA1.digest((name, entry.parent, entry.options, entry.theories).toString)  | 
| 
48675
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
66  | 
|
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
67  | 
val info =  | 
| 
48738
 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 
wenzelm 
parents: 
48737 
diff
changeset
 | 
68  | 
Session_Info(select, entry.pos, entry.groups, dir + Path.explode(entry.path),  | 
| 
 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 
wenzelm 
parents: 
48737 
diff
changeset
 | 
69  | 
entry.parent, entry.description, session_options, theories, files, entry_digest)  | 
| 
48675
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
70  | 
|
| 
48738
 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 
wenzelm 
parents: 
48737 
diff
changeset
 | 
71  | 
(name, info)  | 
| 
48675
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
72  | 
}  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
73  | 
    catch {
 | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
74  | 
case ERROR(msg) =>  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
75  | 
error(msg + "\nThe error(s) above occurred in session entry " +  | 
| 48992 | 76  | 
quote(entry.name) + Position.here(entry.pos))  | 
| 
48675
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
77  | 
}  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
78  | 
|
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
79  | 
|
| 
48676
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
80  | 
/* session tree */  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
81  | 
|
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
82  | 
object Session_Tree  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
83  | 
  {
 | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
84  | 
def apply(infos: Seq[(String, Session_Info)]): Session_Tree =  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
85  | 
    {
 | 
| 48680 | 86  | 
val graph1 =  | 
87  | 
        (Graph.string[Session_Info] /: infos) {
 | 
|
| 
48684
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
88  | 
case (graph, (name, info)) =>  | 
| 48680 | 89  | 
if (graph.defined(name))  | 
| 48992 | 90  | 
              error("Duplicate session " + quote(name) + Position.here(info.pos))
 | 
| 48680 | 91  | 
else graph.new_node(name, info)  | 
92  | 
}  | 
|
93  | 
val graph2 =  | 
|
94  | 
        (graph1 /: graph1.entries) {
 | 
|
95  | 
case (graph, (name, (info, _))) =>  | 
|
96  | 
            info.parent match {
 | 
|
97  | 
case None => graph  | 
|
98  | 
case Some(parent) =>  | 
|
99  | 
if (!graph.defined(parent))  | 
|
100  | 
                  error("Bad parent session " + quote(parent) + " for " +
 | 
|
| 48992 | 101  | 
quote(name) + Position.here(info.pos))  | 
| 
48676
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
102  | 
|
| 48680 | 103  | 
                try { graph.add_edge_acyclic(parent, name) }
 | 
104  | 
                catch {
 | 
|
105  | 
case exn: Graph.Cycles[_] =>  | 
|
106  | 
error(cat_lines(exn.cycles.map(cycle =>  | 
|
107  | 
"Cyclic session dependency of " +  | 
|
108  | 
                        cycle.map(c => quote(c.toString)).mkString(" via "))) +
 | 
|
| 48992 | 109  | 
Position.here(info.pos))  | 
| 48680 | 110  | 
}  | 
| 
48676
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
111  | 
}  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
112  | 
}  | 
| 48680 | 113  | 
new Session_Tree(graph2)  | 
| 
48676
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
114  | 
}  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
115  | 
}  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
116  | 
|
| 48680 | 117  | 
final class Session_Tree private(val graph: Graph[String, Session_Info])  | 
| 
48676
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
118  | 
extends PartialFunction[String, Session_Info]  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
119  | 
  {
 | 
| 48680 | 120  | 
def apply(name: String): Session_Info = graph.get_node(name)  | 
121  | 
def isDefinedAt(name: String): Boolean = graph.defined(name)  | 
|
| 
48676
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
122  | 
|
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
123  | 
def required(all_sessions: Boolean, session_groups: List[String], sessions: List[String])  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
124  | 
: (List[String], Session_Tree) =  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
125  | 
    {
 | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
126  | 
val bad_sessions = sessions.filterNot(isDefinedAt(_))  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
127  | 
      if (!bad_sessions.isEmpty) error("Undefined session(s): " + commas_quote(bad_sessions))
 | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
128  | 
|
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
129  | 
val selected =  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
130  | 
      {
 | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
131  | 
if (all_sessions) graph.keys.toList  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
132  | 
        else {
 | 
| 
48737
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
133  | 
val select_group = session_groups.toSet  | 
| 
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
134  | 
val select = sessions.toSet  | 
| 
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
135  | 
          (for {
 | 
| 
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
136  | 
(name, (info, _)) <- graph.entries  | 
| 
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
137  | 
if info.select || select(name) || apply(name).groups.exists(select_group)  | 
| 
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
138  | 
} yield name).toList  | 
| 
48676
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
139  | 
}  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
140  | 
}  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
141  | 
val descendants = graph.all_succs(selected)  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
142  | 
val tree1 = new Session_Tree(graph.restrict(graph.all_preds(selected).toSet))  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
143  | 
(descendants, tree1)  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
144  | 
}  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
145  | 
|
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
146  | 
def topological_order: List[(String, Session_Info)] =  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
147  | 
graph.topological_order.map(name => (name, apply(name)))  | 
| 48794 | 148  | 
|
149  | 
    override def toString: String = graph.entries.map(_._1).toList.sorted.mkString(",")
 | 
|
| 
48676
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
150  | 
}  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
151  | 
|
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
152  | 
|
| 
48675
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
153  | 
/* parser */  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
154  | 
|
| 
48713
 
de26cf3191a3
more token markers, based on actual outer syntax;
 
wenzelm 
parents: 
48710 
diff
changeset
 | 
155  | 
private val SESSION = "session"  | 
| 
 
de26cf3191a3
more token markers, based on actual outer syntax;
 
wenzelm 
parents: 
48710 
diff
changeset
 | 
156  | 
private val IN = "in"  | 
| 
 
de26cf3191a3
more token markers, based on actual outer syntax;
 
wenzelm 
parents: 
48710 
diff
changeset
 | 
157  | 
private val DESCRIPTION = "description"  | 
| 
 
de26cf3191a3
more token markers, based on actual outer syntax;
 
wenzelm 
parents: 
48710 
diff
changeset
 | 
158  | 
private val OPTIONS = "options"  | 
| 
 
de26cf3191a3
more token markers, based on actual outer syntax;
 
wenzelm 
parents: 
48710 
diff
changeset
 | 
159  | 
private val THEORIES = "theories"  | 
| 
 
de26cf3191a3
more token markers, based on actual outer syntax;
 
wenzelm 
parents: 
48710 
diff
changeset
 | 
160  | 
private val FILES = "files"  | 
| 
 
de26cf3191a3
more token markers, based on actual outer syntax;
 
wenzelm 
parents: 
48710 
diff
changeset
 | 
161  | 
|
| 
 
de26cf3191a3
more token markers, based on actual outer syntax;
 
wenzelm 
parents: 
48710 
diff
changeset
 | 
162  | 
lazy val root_syntax =  | 
| 
48738
 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 
wenzelm 
parents: 
48737 
diff
changeset
 | 
163  | 
    Outer_Syntax.init() + "(" + ")" + "+" + "," + "=" + "[" + "]" +
 | 
| 48718 | 164  | 
(SESSION, Keyword.THY_DECL) + IN + DESCRIPTION + OPTIONS + THEORIES + FILES  | 
| 
48713
 
de26cf3191a3
more token markers, based on actual outer syntax;
 
wenzelm 
parents: 
48710 
diff
changeset
 | 
165  | 
|
| 
48675
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
166  | 
private object Parser extends Parse.Parser  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
167  | 
  {
 | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
168  | 
def session_entry(pos: Position.T): Parser[Session_Entry] =  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
169  | 
    {
 | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
170  | 
      val session_name = atom("session name", _.is_name)
 | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
171  | 
|
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
172  | 
val option =  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
173  | 
        name ~ opt(keyword("=") ~! name ^^ { case _ ~ x => x }) ^^ { case x ~ y => (x, y) }
 | 
| 48862 | 174  | 
      val options = keyword("[") ~> rep1sep(option, keyword(",")) <~ keyword("]")
 | 
| 
48675
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
175  | 
|
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
176  | 
val theories =  | 
| 48739 | 177  | 
keyword(THEORIES) ~! ((options | success(Nil)) ~ rep(theory_name)) ^^  | 
| 
48675
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
178  | 
          { case _ ~ (x ~ y) => (x, y) }
 | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
179  | 
|
| 48912 | 180  | 
command(SESSION) ~!  | 
181  | 
(session_name ~  | 
|
182  | 
          ((keyword("(") ~! (rep1(name) <~ keyword(")")) ^^ { case _ ~ x => x }) | success(Nil)) ~
 | 
|
183  | 
          ((keyword(IN) ~! path ^^ { case _ ~ x => x }) | success(".")) ~
 | 
|
184  | 
          (keyword("=") ~!
 | 
|
185  | 
            (opt(session_name ~! keyword("+") ^^ { case x ~ _ => x }) ~
 | 
|
186  | 
              ((keyword(DESCRIPTION) ~! text ^^ { case _ ~ x => x }) | success("")) ~
 | 
|
187  | 
              ((keyword(OPTIONS) ~! options ^^ { case _ ~ x => x }) | success(Nil)) ~
 | 
|
| 48916 | 188  | 
rep1(theories) ~  | 
| 48912 | 189  | 
              ((keyword(FILES) ~! rep1(path) ^^ { case _ ~ x => x }) | success(Nil))))) ^^
 | 
190  | 
        { case _ ~ (a ~ b ~ c ~ (_ ~ (d ~ e ~ f ~ g ~ h))) =>
 | 
|
191  | 
Session_Entry(pos, a, b, c, d, e, f, g, h) }  | 
|
| 
48675
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
192  | 
}  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
193  | 
|
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
194  | 
def parse_entries(root: Path): List[Session_Entry] =  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
195  | 
    {
 | 
| 
48713
 
de26cf3191a3
more token markers, based on actual outer syntax;
 
wenzelm 
parents: 
48710 
diff
changeset
 | 
196  | 
val toks = root_syntax.scan(File.read(root))  | 
| 
48675
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
197  | 
      parse_all(rep(session_entry(root.position)), Token.reader(toks, root.implode)) match {
 | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
198  | 
case Success(result, _) => result  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
199  | 
case bad => error(bad.toString)  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
200  | 
}  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
201  | 
}  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
202  | 
}  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
203  | 
|
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
204  | 
|
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
205  | 
/* find sessions within certain directories */  | 
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
206  | 
|
| 
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
207  | 
  private val ROOT = Path.explode("ROOT")
 | 
| 
48684
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
208  | 
  private val ROOTS = Path.explode("ROOTS")
 | 
| 
48675
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
209  | 
|
| 
48684
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
210  | 
private def is_session_dir(dir: Path): Boolean =  | 
| 
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
211  | 
(dir + ROOT).is_file || (dir + ROOTS).is_file  | 
| 
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
212  | 
|
| 
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
213  | 
private def check_session_dir(dir: Path): Path =  | 
| 
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
214  | 
if (is_session_dir(dir)) dir  | 
| 
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
215  | 
    else error("Bad session root directory: " + dir.toString)
 | 
| 
48675
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
216  | 
|
| 
48737
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
217  | 
def find_sessions(options: Options, more_dirs: List[(Boolean, Path)]): Session_Tree =  | 
| 
48675
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
218  | 
  {
 | 
| 
48737
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
219  | 
def find_dir(select: Boolean, dir: Path): List[(String, Session_Info)] =  | 
| 
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
220  | 
find_root(select, dir) ::: find_roots(select, dir)  | 
| 
48684
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
221  | 
|
| 
48737
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
222  | 
def find_root(select: Boolean, dir: Path): List[(String, Session_Info)] =  | 
| 
48684
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
223  | 
    {
 | 
| 
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
224  | 
val root = dir + ROOT  | 
| 
48737
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
225  | 
if (root.is_file) Parser.parse_entries(root).map(session_info(options, select, dir, _))  | 
| 
48684
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
226  | 
else Nil  | 
| 
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
227  | 
}  | 
| 
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
228  | 
|
| 
48737
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
229  | 
def find_roots(select: Boolean, dir: Path): List[(String, Session_Info)] =  | 
| 
48684
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
230  | 
    {
 | 
| 
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
231  | 
val roots = dir + ROOTS  | 
| 
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
232  | 
      if (roots.is_file) {
 | 
| 
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
233  | 
        for {
 | 
| 
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
234  | 
line <- split_lines(File.read(roots))  | 
| 
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
235  | 
          if !(line == "" || line.startsWith("#"))
 | 
| 
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
236  | 
dir1 =  | 
| 
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
237  | 
            try { check_session_dir(dir + Path.explode(line)) }
 | 
| 
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
238  | 
            catch {
 | 
| 
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
239  | 
case ERROR(msg) =>  | 
| 
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
240  | 
error(msg + "\nThe error(s) above occurred in session catalog " + roots.toString)  | 
| 
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
241  | 
}  | 
| 
48737
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
242  | 
info <- find_dir(select, dir1)  | 
| 
48684
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
243  | 
} yield info  | 
| 
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
244  | 
}  | 
| 
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
245  | 
else Nil  | 
| 
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
246  | 
}  | 
| 
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
247  | 
|
| 
48737
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
248  | 
val default_dirs = Isabelle_System.components().filter(is_session_dir(_)).map((false, _))  | 
| 
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
249  | 
|
| 
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
250  | 
    more_dirs foreach { case (_, dir) => check_session_dir(dir) }
 | 
| 
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
251  | 
|
| 
48684
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
252  | 
Session_Tree(  | 
| 
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
253  | 
      for {
 | 
| 
48737
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
254  | 
(select, dir) <- default_dirs ::: more_dirs  | 
| 
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
255  | 
info <- find_dir(select, dir)  | 
| 
48684
 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 
wenzelm 
parents: 
48680 
diff
changeset
 | 
256  | 
} yield info)  | 
| 
48349
 
a78e5d399599
support Session.Queue with ordering and dependencies;
 
wenzelm 
parents: 
48347 
diff
changeset
 | 
257  | 
}  | 
| 48334 | 258  | 
|
| 
48337
 
9c7f8e5805b4
cumulate semantic Session_Info, based on syntactic Session_Entry;
 
wenzelm 
parents: 
48336 
diff
changeset
 | 
259  | 
|
| 48424 | 260  | 
|
261  | 
/** build **/  | 
|
262  | 
||
| 48478 | 263  | 
  private def echo(msg: String) { java.lang.System.out.println(msg) }
 | 
264  | 
private def sleep(): Unit = Thread.sleep(500)  | 
|
265  | 
||
266  | 
||
| 
48676
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
267  | 
/* queue */  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
268  | 
|
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
269  | 
object Queue  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
270  | 
  {
 | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
271  | 
def apply(tree: Session_Tree): Queue =  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
272  | 
    {
 | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
273  | 
val graph = tree.graph  | 
| 
48678
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
274  | 
|
| 
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
275  | 
def outdegree(name: String): Int = graph.imm_succs(name).size  | 
| 
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
276  | 
      def timeout(name: String): Double = tree(name).options.real("timeout")
 | 
| 
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
277  | 
|
| 
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
278  | 
object Ordering extends scala.math.Ordering[String]  | 
| 
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
279  | 
      {
 | 
| 
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
280  | 
def compare(name1: String, name2: String): Int =  | 
| 
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
281  | 
          outdegree(name2) compare outdegree(name1) match {
 | 
| 
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
282  | 
case 0 =>  | 
| 
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
283  | 
              timeout(name2) compare timeout(name1) match {
 | 
| 
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
284  | 
case 0 => name1 compare name2  | 
| 
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
285  | 
case ord => ord  | 
| 
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
286  | 
}  | 
| 
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
287  | 
case ord => ord  | 
| 
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
288  | 
}  | 
| 
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
289  | 
}  | 
| 
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
290  | 
|
| 
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
291  | 
new Queue(graph, SortedSet(graph.keys.toList: _*)(Ordering))  | 
| 
48676
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
292  | 
}  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
293  | 
}  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
294  | 
|
| 48680 | 295  | 
final class Queue private(graph: Graph[String, Session_Info], order: SortedSet[String])  | 
| 
48676
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
296  | 
  {
 | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
297  | 
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
 | 
298  | 
|
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
299  | 
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
 | 
300  | 
|
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
301  | 
def - (name: String): Queue = new Queue(graph.del_node(name), order - name)  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
302  | 
|
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
303  | 
def dequeue(skip: String => Boolean): Option[(String, Session_Info)] =  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
304  | 
    {
 | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
305  | 
val it = order.iterator.dropWhile(name => skip(name) || !graph.is_minimal(name))  | 
| 48680 | 306  | 
      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
 | 
307  | 
else None  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
308  | 
}  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
309  | 
}  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
310  | 
|
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
311  | 
|
| 
48710
 
5b51ccdc8623
prefer static Build.session_content for loaded theories -- discontinued incremental protocol;
 
wenzelm 
parents: 
48708 
diff
changeset
 | 
312  | 
/* source dependencies and static content */  | 
| 
48422
 
9613780a805b
determine source dependencies, relatively to preloaded theories;
 
wenzelm 
parents: 
48421 
diff
changeset
 | 
313  | 
|
| 
48710
 
5b51ccdc8623
prefer static Build.session_content for loaded theories -- discontinued incremental protocol;
 
wenzelm 
parents: 
48708 
diff
changeset
 | 
314  | 
sealed case class Session_Content(  | 
| 
48422
 
9613780a805b
determine source dependencies, relatively to preloaded theories;
 
wenzelm 
parents: 
48421 
diff
changeset
 | 
315  | 
loaded_theories: Set[String],  | 
| 
48660
 
730ca503e955
static outer syntax based on session specifications;
 
wenzelm 
parents: 
48650 
diff
changeset
 | 
316  | 
syntax: Outer_Syntax,  | 
| 48794 | 317  | 
sources: List[(Path, SHA1.Digest)],  | 
318  | 
errors: List[String])  | 
|
319  | 
  {
 | 
|
320  | 
def check_errors: Session_Content =  | 
|
321  | 
    {
 | 
|
322  | 
if (errors.isEmpty) this  | 
|
323  | 
else error(cat_lines(errors))  | 
|
324  | 
}  | 
|
325  | 
}  | 
|
| 
48422
 
9613780a805b
determine source dependencies, relatively to preloaded theories;
 
wenzelm 
parents: 
48421 
diff
changeset
 | 
326  | 
|
| 
48710
 
5b51ccdc8623
prefer static Build.session_content for loaded theories -- discontinued incremental protocol;
 
wenzelm 
parents: 
48708 
diff
changeset
 | 
327  | 
sealed case class Deps(deps: Map[String, Session_Content])  | 
| 
48423
 
0ccf143a2a69
maintain set of source digests, including relevant parts of session entry;
 
wenzelm 
parents: 
48422 
diff
changeset
 | 
328  | 
  {
 | 
| 48583 | 329  | 
def is_empty: Boolean = deps.isEmpty  | 
| 
48710
 
5b51ccdc8623
prefer static Build.session_content for loaded theories -- discontinued incremental protocol;
 
wenzelm 
parents: 
48708 
diff
changeset
 | 
330  | 
def apply(name: String): Session_Content = deps(name)  | 
| 
48504
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
331  | 
def sources(name: String): List[SHA1.Digest] = deps(name).sources.map(_._2)  | 
| 
48423
 
0ccf143a2a69
maintain set of source digests, including relevant parts of session entry;
 
wenzelm 
parents: 
48422 
diff
changeset
 | 
332  | 
}  | 
| 
 
0ccf143a2a69
maintain set of source digests, including relevant parts of session entry;
 
wenzelm 
parents: 
48422 
diff
changeset
 | 
333  | 
|
| 48903 | 334  | 
def dependencies(verbose: Boolean, list_files: Boolean, tree: Session_Tree): Deps =  | 
| 
48710
 
5b51ccdc8623
prefer static Build.session_content for loaded theories -- discontinued incremental protocol;
 
wenzelm 
parents: 
48708 
diff
changeset
 | 
335  | 
Deps((Map.empty[String, Session_Content] /: tree.topological_order)(  | 
| 
48422
 
9613780a805b
determine source dependencies, relatively to preloaded theories;
 
wenzelm 
parents: 
48421 
diff
changeset
 | 
336  | 
      { case (deps, (name, info)) =>
 | 
| 48794 | 337  | 
val (preloaded, parent_syntax, parent_errors) =  | 
| 
48422
 
9613780a805b
determine source dependencies, relatively to preloaded theories;
 
wenzelm 
parents: 
48421 
diff
changeset
 | 
338  | 
            info.parent match {
 | 
| 
48870
 
4accee106f0f
clarified initialization of Thy_Load, Thy_Info, Session;
 
wenzelm 
parents: 
48867 
diff
changeset
 | 
339  | 
case None =>  | 
| 
 
4accee106f0f
clarified initialization of Thy_Load, Thy_Info, Session;
 
wenzelm 
parents: 
48867 
diff
changeset
 | 
340  | 
(Set.empty[String], Outer_Syntax.init_pure(), Nil)  | 
| 48794 | 341  | 
case Some(parent_name) =>  | 
342  | 
val parent = deps(parent_name)  | 
|
343  | 
(parent.loaded_theories, parent.syntax, parent.errors)  | 
|
| 
48422
 
9613780a805b
determine source dependencies, relatively to preloaded theories;
 
wenzelm 
parents: 
48421 
diff
changeset
 | 
344  | 
}  | 
| 
48870
 
4accee106f0f
clarified initialization of Thy_Load, Thy_Info, Session;
 
wenzelm 
parents: 
48867 
diff
changeset
 | 
345  | 
val thy_info = new Thy_Info(new Thy_Load(preloaded, parent_syntax))  | 
| 
48422
 
9613780a805b
determine source dependencies, relatively to preloaded theories;
 
wenzelm 
parents: 
48421 
diff
changeset
 | 
346  | 
|
| 48903 | 347  | 
          if (verbose || list_files) {
 | 
| 48583 | 348  | 
val groups =  | 
349  | 
if (info.groups.isEmpty) ""  | 
|
350  | 
              else info.groups.mkString(" (", " ", ")")
 | 
|
| 48674 | 351  | 
            echo("Session " + name + groups)
 | 
| 48583 | 352  | 
}  | 
| 48478 | 353  | 
|
| 
48422
 
9613780a805b
determine source dependencies, relatively to preloaded theories;
 
wenzelm 
parents: 
48421 
diff
changeset
 | 
354  | 
val thy_deps =  | 
| 
 
9613780a805b
determine source dependencies, relatively to preloaded theories;
 
wenzelm 
parents: 
48421 
diff
changeset
 | 
355  | 
thy_info.dependencies(  | 
| 
 
9613780a805b
determine source dependencies, relatively to preloaded theories;
 
wenzelm 
parents: 
48421 
diff
changeset
 | 
356  | 
info.theories.map(_._2).flatten.  | 
| 
 
9613780a805b
determine source dependencies, relatively to preloaded theories;
 
wenzelm 
parents: 
48421 
diff
changeset
 | 
357  | 
map(thy => Document.Node.Name(info.dir + Thy_Load.thy_path(thy))))  | 
| 
 
9613780a805b
determine source dependencies, relatively to preloaded theories;
 
wenzelm 
parents: 
48421 
diff
changeset
 | 
358  | 
|
| 48872 | 359  | 
val loaded_theories = thy_deps.loaded_theories  | 
| 48883 | 360  | 
val syntax = thy_deps.make_syntax  | 
| 
48660
 
730ca503e955
static outer syntax based on session specifications;
 
wenzelm 
parents: 
48650 
diff
changeset
 | 
361  | 
|
| 
48423
 
0ccf143a2a69
maintain set of source digests, including relevant parts of session entry;
 
wenzelm 
parents: 
48422 
diff
changeset
 | 
362  | 
val all_files =  | 
| 48904 | 363  | 
            (thy_deps.deps.map({ case (n, h) =>
 | 
364  | 
val thy = Path.explode(n.node)  | 
|
365  | 
val uses = h.uses.map(p => Path.explode(n.dir) + Path.explode(p._1))  | 
|
| 
48422
 
9613780a805b
determine source dependencies, relatively to preloaded theories;
 
wenzelm 
parents: 
48421 
diff
changeset
 | 
366  | 
thy :: uses  | 
| 48904 | 367  | 
}).flatten ::: info.files.map(file => info.dir + file)).map(_.expand)  | 
| 48903 | 368  | 
|
369  | 
if (list_files)  | 
|
370  | 
            echo(cat_lines(all_files.map(_.implode).sorted.map("  " + _)))
 | 
|
371  | 
||
| 48485 | 372  | 
val sources =  | 
373  | 
            try { all_files.map(p => (p, SHA1.digest(p))) }
 | 
|
374  | 
            catch {
 | 
|
375  | 
case ERROR(msg) =>  | 
|
| 
48675
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
376  | 
error(msg + "\nThe error(s) above occurred in session " +  | 
| 48992 | 377  | 
quote(name) + Position.here(info.pos))  | 
| 48485 | 378  | 
}  | 
| 48872 | 379  | 
val errors = parent_errors ::: thy_deps.deps.map(d => d._2.errors).flatten  | 
| 
48422
 
9613780a805b
determine source dependencies, relatively to preloaded theories;
 
wenzelm 
parents: 
48421 
diff
changeset
 | 
380  | 
|
| 48794 | 381  | 
deps + (name -> Session_Content(loaded_theories, syntax, sources, errors))  | 
| 
48423
 
0ccf143a2a69
maintain set of source digests, including relevant parts of session entry;
 
wenzelm 
parents: 
48422 
diff
changeset
 | 
382  | 
}))  | 
| 
48422
 
9613780a805b
determine source dependencies, relatively to preloaded theories;
 
wenzelm 
parents: 
48421 
diff
changeset
 | 
383  | 
|
| 48791 | 384  | 
def session_content(dirs: List[Path], session: String): Session_Content =  | 
| 
48710
 
5b51ccdc8623
prefer static Build.session_content for loaded theories -- discontinued incremental protocol;
 
wenzelm 
parents: 
48708 
diff
changeset
 | 
385  | 
  {
 | 
| 48791 | 386  | 
val (_, tree) =  | 
387  | 
find_sessions(Options.init(), dirs.map((false, _))).required(false, Nil, List(session))  | 
|
| 48903 | 388  | 
dependencies(false, false, tree)(session)  | 
| 
48710
 
5b51ccdc8623
prefer static Build.session_content for loaded theories -- discontinued incremental protocol;
 
wenzelm 
parents: 
48708 
diff
changeset
 | 
389  | 
}  | 
| 
 
5b51ccdc8623
prefer static Build.session_content for loaded theories -- discontinued incremental protocol;
 
wenzelm 
parents: 
48708 
diff
changeset
 | 
390  | 
|
| 48794 | 391  | 
def outer_syntax(session: String): Outer_Syntax =  | 
392  | 
session_content(Nil, session).check_errors.syntax  | 
|
| 
48710
 
5b51ccdc8623
prefer static Build.session_content for loaded theories -- discontinued incremental protocol;
 
wenzelm 
parents: 
48708 
diff
changeset
 | 
393  | 
|
| 
48422
 
9613780a805b
determine source dependencies, relatively to preloaded theories;
 
wenzelm 
parents: 
48421 
diff
changeset
 | 
394  | 
|
| 48424 | 395  | 
/* jobs */  | 
| 48341 | 396  | 
|
| 
48675
 
10f5303f86e5
clarified Session_Entry vs. Session_Info with related parsing operations;
 
wenzelm 
parents: 
48674 
diff
changeset
 | 
397  | 
private class Job(name: String, info: Session_Info, output: Path, do_output: Boolean,  | 
| 48674 | 398  | 
verbose: Boolean, browser_info: Path)  | 
| 48418 | 399  | 
  {
 | 
| 
48467
 
a4318c36a829
more precise propagation of options: build, session, theories;
 
wenzelm 
parents: 
48462 
diff
changeset
 | 
400  | 
// global browser info dir  | 
| 48674 | 401  | 
    if (info.options.bool("browser_info") && !(browser_info + Path.explode("index.html")).is_file)
 | 
| 
48467
 
a4318c36a829
more precise propagation of options: build, session, theories;
 
wenzelm 
parents: 
48462 
diff
changeset
 | 
402  | 
    {
 | 
| 
 
a4318c36a829
more precise propagation of options: build, session, theories;
 
wenzelm 
parents: 
48462 
diff
changeset
 | 
403  | 
browser_info.file.mkdirs()  | 
| 
 
a4318c36a829
more precise propagation of options: build, session, theories;
 
wenzelm 
parents: 
48462 
diff
changeset
 | 
404  | 
      File.copy(Path.explode("~~/lib/logo/isabelle.gif"),
 | 
| 
 
a4318c36a829
more precise propagation of options: build, session, theories;
 
wenzelm 
parents: 
48462 
diff
changeset
 | 
405  | 
        browser_info + Path.explode("isabelle.gif"))
 | 
| 
 
a4318c36a829
more precise propagation of options: build, session, theories;
 
wenzelm 
parents: 
48462 
diff
changeset
 | 
406  | 
      File.write(browser_info + Path.explode("index.html"),
 | 
| 
 
a4318c36a829
more precise propagation of options: build, session, theories;
 
wenzelm 
parents: 
48462 
diff
changeset
 | 
407  | 
        File.read(Path.explode("~~/lib/html/library_index_header.template")) +
 | 
| 48724 | 408  | 
        File.read(Path.explode("~~/lib/html/library_index_content.template")) +
 | 
| 
48467
 
a4318c36a829
more precise propagation of options: build, session, theories;
 
wenzelm 
parents: 
48462 
diff
changeset
 | 
409  | 
        File.read(Path.explode("~~/lib/html/library_index_footer.template")))
 | 
| 
 
a4318c36a829
more precise propagation of options: build, session, theories;
 
wenzelm 
parents: 
48462 
diff
changeset
 | 
410  | 
}  | 
| 
 
a4318c36a829
more precise propagation of options: build, session, theories;
 
wenzelm 
parents: 
48462 
diff
changeset
 | 
411  | 
|
| 48674 | 412  | 
def output_path: Option[Path] = if (do_output) Some(output) else None  | 
413  | 
||
414  | 
    private val parent = info.parent.getOrElse("")
 | 
|
| 48418 | 415  | 
|
| 
48698
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48684 
diff
changeset
 | 
416  | 
    private val args_file = File.tmp_file("args")
 | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48684 
diff
changeset
 | 
417  | 
File.write(args_file, YXML.string_of_body(  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48684 
diff
changeset
 | 
418  | 
if (is_pure(name)) Options.encode(info.options)  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48684 
diff
changeset
 | 
419  | 
else  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48684 
diff
changeset
 | 
420  | 
        {
 | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48684 
diff
changeset
 | 
421  | 
import XML.Encode._  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48684 
diff
changeset
 | 
422  | 
pair(bool, pair(Options.encode, pair(bool, pair(Path.encode, pair(string,  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48684 
diff
changeset
 | 
423  | 
pair(string, list(pair(Options.encode, list(Path.encode)))))))))(  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48684 
diff
changeset
 | 
424  | 
(do_output, (info.options, (verbose, (browser_info, (parent,  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48684 
diff
changeset
 | 
425  | 
(name, info.theories)))))))  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48684 
diff
changeset
 | 
426  | 
}))  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48684 
diff
changeset
 | 
427  | 
|
| 48674 | 428  | 
private val env =  | 
| 
48698
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48684 
diff
changeset
 | 
429  | 
      Map("INPUT" -> parent, "TARGET" -> name, "OUTPUT" -> Isabelle_System.standard_path(output),
 | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48684 
diff
changeset
 | 
430  | 
(if (is_pure(name)) "ISABELLE_PROCESS_OPTIONS" else "ARGS_FILE") ->  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48684 
diff
changeset
 | 
431  | 
Isabelle_System.posix_path(args_file.getPath))  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48684 
diff
changeset
 | 
432  | 
|
| 48674 | 433  | 
private val script =  | 
| 
48511
 
37999ee01156
remove old output heaps, to ensure that result is valid wrt. check_stamps;
 
wenzelm 
parents: 
48509 
diff
changeset
 | 
434  | 
      if (is_pure(name)) {
 | 
| 
 
37999ee01156
remove old output heaps, to ensure that result is valid wrt. check_stamps;
 
wenzelm 
parents: 
48509 
diff
changeset
 | 
435  | 
if (do_output) "./build " + name + " \"$OUTPUT\""  | 
| 
 
37999ee01156
remove old output heaps, to ensure that result is valid wrt. check_stamps;
 
wenzelm 
parents: 
48509 
diff
changeset
 | 
436  | 
else """ rm -f "$OUTPUT"; ./build """ + name  | 
| 
 
37999ee01156
remove old output heaps, to ensure that result is valid wrt. check_stamps;
 
wenzelm 
parents: 
48509 
diff
changeset
 | 
437  | 
}  | 
| 48418 | 438  | 
      else {
 | 
439  | 
"""  | 
|
440  | 
. "$ISABELLE_HOME/lib/scripts/timestart.bash"  | 
|
441  | 
""" +  | 
|
| 
48511
 
37999ee01156
remove old output heaps, to ensure that result is valid wrt. check_stamps;
 
wenzelm 
parents: 
48509 
diff
changeset
 | 
442  | 
(if (do_output)  | 
| 
 
37999ee01156
remove old output heaps, to ensure that result is valid wrt. check_stamps;
 
wenzelm 
parents: 
48509 
diff
changeset
 | 
443  | 
"""  | 
| 
 
37999ee01156
remove old output heaps, to ensure that result is valid wrt. check_stamps;
 
wenzelm 
parents: 
48509 
diff
changeset
 | 
444  | 
"$ISABELLE_PROCESS" -e "Build.build \"$ARGS_FILE\";" -q -w "$INPUT" "$OUTPUT"  | 
| 
 
37999ee01156
remove old output heaps, to ensure that result is valid wrt. check_stamps;
 
wenzelm 
parents: 
48509 
diff
changeset
 | 
445  | 
"""  | 
| 48418 | 446  | 
else  | 
| 
48511
 
37999ee01156
remove old output heaps, to ensure that result is valid wrt. check_stamps;
 
wenzelm 
parents: 
48509 
diff
changeset
 | 
447  | 
"""  | 
| 
 
37999ee01156
remove old output heaps, to ensure that result is valid wrt. check_stamps;
 
wenzelm 
parents: 
48509 
diff
changeset
 | 
448  | 
rm -f "$OUTPUT"; "$ISABELLE_PROCESS" -e "Build.build \"$ARGS_FILE\";" -r -q "$INPUT"  | 
| 
 
37999ee01156
remove old output heaps, to ensure that result is valid wrt. check_stamps;
 
wenzelm 
parents: 
48509 
diff
changeset
 | 
449  | 
""") +  | 
| 48418 | 450  | 
"""  | 
451  | 
RC="$?"  | 
|
452  | 
||
453  | 
. "$ISABELLE_HOME/lib/scripts/timestop.bash"  | 
|
454  | 
||
455  | 
if [ "$RC" -eq 0 ]; then  | 
|
456  | 
echo "Finished $TARGET ($TIMES_REPORT)" >&2  | 
|
457  | 
fi  | 
|
458  | 
||
459  | 
exit "$RC"  | 
|
460  | 
"""  | 
|
461  | 
}  | 
|
| 48674 | 462  | 
|
463  | 
private val (thread, result) =  | 
|
| 
48698
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48684 
diff
changeset
 | 
464  | 
      Simple_Thread.future("build") { Isabelle_System.bash_env(info.dir.file, env, script) }
 | 
| 48674 | 465  | 
|
466  | 
def terminate: Unit = thread.interrupt  | 
|
467  | 
def is_finished: Boolean = result.is_finished  | 
|
468  | 
||
469  | 
@volatile private var timeout = false  | 
|
470  | 
    private val time = Time.seconds(info.options.real("timeout"))
 | 
|
471  | 
private val timer: Option[Timer] =  | 
|
472  | 
      if (time.seconds > 0.0) {
 | 
|
473  | 
        val t = new Timer("build", true)
 | 
|
474  | 
        t.schedule(new TimerTask { def run = { terminate; timeout = true } }, time.ms)
 | 
|
475  | 
Some(t)  | 
|
476  | 
}  | 
|
477  | 
else None  | 
|
478  | 
||
479  | 
    def join: (String, String, Int) = {
 | 
|
480  | 
val (out, err, rc) = result.join  | 
|
481  | 
args_file.delete  | 
|
482  | 
timer.map(_.cancel())  | 
|
483  | 
||
484  | 
val err1 =  | 
|
485  | 
if (rc == 130)  | 
|
486  | 
          (if (err.isEmpty || err.endsWith("\n")) err else err + "\n") +
 | 
|
487  | 
(if (timeout) "*** Timeout\n" else "*** Interrupt\n")  | 
|
488  | 
else err  | 
|
489  | 
(out, err1, rc)  | 
|
490  | 
}  | 
|
| 48364 | 491  | 
}  | 
492  | 
||
| 48424 | 493  | 
|
| 
48504
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
494  | 
/* log files and corresponding heaps */  | 
| 
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
495  | 
|
| 48505 | 496  | 
  private val LOG = Path.explode("log")
 | 
497  | 
private def log(name: String): Path = LOG + Path.basic(name)  | 
|
498  | 
  private def log_gz(name: String): Path = log(name).ext("gz")
 | 
|
| 
48504
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
499  | 
|
| 48505 | 500  | 
private def sources_stamp(digests: List[SHA1.Digest]): String =  | 
| 
48504
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
501  | 
    digests.map(_.toString).sorted.mkString("sources: ", " ", "")
 | 
| 
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
502  | 
|
| 
48639
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
503  | 
private val no_heap: String = "heap: -"  | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
504  | 
|
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
505  | 
private def heap_stamp(heap: Option[Path]): String =  | 
| 
48504
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
506  | 
  {
 | 
| 
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
507  | 
"heap: " +  | 
| 
48639
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
508  | 
      (heap match {
 | 
| 
48504
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
509  | 
case Some(path) =>  | 
| 
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
510  | 
val file = path.file  | 
| 
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
511  | 
if (file.isFile) file.length.toString + " " + file.lastModified.toString  | 
| 
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
512  | 
else "-"  | 
| 
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
513  | 
case None => "-"  | 
| 
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
514  | 
})  | 
| 
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
515  | 
}  | 
| 
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
516  | 
|
| 
48639
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
517  | 
private def read_stamps(path: Path): Option[(String, String, String)] =  | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
518  | 
    if (path.is_file) {
 | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
519  | 
val stream = new GZIPInputStream (new BufferedInputStream(new FileInputStream(path.file)))  | 
| 
48504
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
520  | 
val reader = new BufferedReader(new InputStreamReader(stream, Standard_System.charset))  | 
| 
48639
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
521  | 
val (s, h1, h2) =  | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
522  | 
        try { (reader.readLine, reader.readLine, reader.readLine) }
 | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
523  | 
        finally { reader.close }
 | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
524  | 
      if (s != null && s.startsWith("sources: ") &&
 | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
525  | 
          h1 != null && h1.startsWith("heap: ") &&
 | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
526  | 
          h2 != null && h2.startsWith("heap: ")) Some((s, h1, h2))
 | 
| 
48504
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
527  | 
else None  | 
| 
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
528  | 
}  | 
| 
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
529  | 
else None  | 
| 
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
530  | 
|
| 
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
531  | 
|
| 48425 | 532  | 
/* build */  | 
| 48424 | 533  | 
|
| 48509 | 534  | 
def build(  | 
535  | 
all_sessions: Boolean = false,  | 
|
| 
48511
 
37999ee01156
remove old output heaps, to ensure that result is valid wrt. check_stamps;
 
wenzelm 
parents: 
48509 
diff
changeset
 | 
536  | 
build_heap: Boolean = false,  | 
| 48595 | 537  | 
clean_build: Boolean = false,  | 
| 
48737
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
538  | 
more_dirs: List[(Boolean, Path)] = Nil,  | 
| 48509 | 539  | 
session_groups: List[String] = Nil,  | 
540  | 
max_jobs: Int = 1,  | 
|
| 48903 | 541  | 
list_files: Boolean = false,  | 
| 48509 | 542  | 
no_build: Boolean = false,  | 
543  | 
build_options: List[String] = Nil,  | 
|
544  | 
system_mode: Boolean = false,  | 
|
545  | 
verbose: Boolean = false,  | 
|
546  | 
sessions: List[String] = Nil): Int =  | 
|
| 48341 | 547  | 
  {
 | 
| 48807 | 548  | 
val options = (Options.init() /: build_options)(_ + _)  | 
| 
48676
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
549  | 
val (descendants, tree) =  | 
| 
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
550  | 
find_sessions(options, more_dirs).required(all_sessions, session_groups, sessions)  | 
| 48903 | 551  | 
val deps = dependencies(verbose, list_files, tree)  | 
| 
48676
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
552  | 
val queue = Queue(tree)  | 
| 48368 | 553  | 
|
| 
48504
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
554  | 
def make_stamp(name: String): String =  | 
| 
48676
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
555  | 
sources_stamp(tree(name).entry_digest :: deps.sources(name))  | 
| 
48504
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
556  | 
|
| 
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
557  | 
val (input_dirs, output_dir, browser_info) =  | 
| 
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
558  | 
      if (system_mode) {
 | 
| 
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
559  | 
        val output_dir = Path.explode("~~/heaps/$ML_IDENTIFIER")
 | 
| 
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
560  | 
        (List(output_dir), output_dir, Path.explode("~~/browser_info"))
 | 
| 
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
561  | 
}  | 
| 
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
562  | 
      else {
 | 
| 
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
563  | 
        val output_dir = Path.explode("$ISABELLE_OUTPUT")
 | 
| 
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
564  | 
(output_dir :: Isabelle_System.find_logics_dirs(), output_dir,  | 
| 
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
565  | 
         Path.explode("$ISABELLE_BROWSER_INFO"))
 | 
| 
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
566  | 
}  | 
| 48363 | 567  | 
|
| 48373 | 568  | 
// prepare log dir  | 
| 
48504
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
569  | 
(output_dir + LOG).file.mkdirs()  | 
| 48373 | 570  | 
|
| 48595 | 571  | 
// optional cleanup  | 
572  | 
    if (clean_build) {
 | 
|
573  | 
      for (name <- descendants) {
 | 
|
574  | 
val files =  | 
|
575  | 
List(Path.basic(name), log(name), log_gz(name)).map(output_dir + _).filter(_.is_file)  | 
|
576  | 
        if (!files.isEmpty) echo("Cleaning " + name + " ...")
 | 
|
577  | 
if (!files.forall(p => p.file.delete)) echo(name + " FAILED to delete")  | 
|
578  | 
}  | 
|
579  | 
}  | 
|
580  | 
||
| 48425 | 581  | 
// scheduler loop  | 
| 
48639
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
582  | 
case class Result(current: Boolean, heap: String, rc: Int)  | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
583  | 
|
| 48425 | 584  | 
@tailrec def loop(  | 
| 
48676
 
3ef82491cdd6
clarified Session_Tree (with proper integrity check) vs. Queue (with provision for alternative ordering);
 
wenzelm 
parents: 
48675 
diff
changeset
 | 
585  | 
pending: Queue,  | 
| 48674 | 586  | 
running: Map[String, (String, Job)],  | 
| 
48639
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
587  | 
results: Map[String, Result]): Map[String, Result] =  | 
| 48425 | 588  | 
    {
 | 
589  | 
if (pending.is_empty) results  | 
|
| 48547 | 590  | 
else  | 
| 48674 | 591  | 
        running.find({ case (_, (_, job)) => job.is_finished }) match {
 | 
592  | 
case Some((name, (parent_heap, job))) =>  | 
|
| 48547 | 593  | 
// finish job  | 
| 48424 | 594  | 
|
| 48547 | 595  | 
val (out, err, rc) = job.join  | 
596  | 
echo(Library.trim_line(err))  | 
|
| 48373 | 597  | 
|
| 
48639
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
598  | 
val heap =  | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
599  | 
              if (rc == 0) {
 | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
600  | 
(output_dir + log(name)).file.delete  | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
601  | 
|
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
602  | 
val sources = make_stamp(name)  | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
603  | 
val heap = heap_stamp(job.output_path)  | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
604  | 
File.write_gzip(output_dir + log_gz(name),  | 
| 48674 | 605  | 
sources + "\n" + parent_heap + "\n" + heap + "\n" + out)  | 
| 
48639
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
606  | 
|
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
607  | 
heap  | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
608  | 
}  | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
609  | 
              else {
 | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
610  | 
(output_dir + Path.basic(name)).file.delete  | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
611  | 
(output_dir + log_gz(name)).file.delete  | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
612  | 
|
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
613  | 
File.write(output_dir + log(name), out)  | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
614  | 
echo(name + " FAILED")  | 
| 48661 | 615  | 
                if (rc != 130) {
 | 
616  | 
                  echo("(see also " + (output_dir + log(name)).file.toString + ")")
 | 
|
617  | 
val lines = split_lines(out)  | 
|
618  | 
val tail = lines.drop(lines.length - 20 max 0)  | 
|
619  | 
                  echo("\n" + cat_lines(tail))
 | 
|
620  | 
}  | 
|
| 
48639
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
621  | 
|
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
622  | 
no_heap  | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
623  | 
}  | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
624  | 
loop(pending - name, running - name, results + (name -> Result(false, heap, rc)))  | 
| 
48528
 
784c6f63d79c
proper all_current, which regards parent status as well;
 
wenzelm 
parents: 
48511 
diff
changeset
 | 
625  | 
|
| 48547 | 626  | 
case None if (running.size < (max_jobs max 1)) =>  | 
627  | 
// check/start next job  | 
|
628  | 
            pending.dequeue(running.isDefinedAt(_)) match {
 | 
|
629  | 
case Some((name, info)) =>  | 
|
| 
48639
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
630  | 
val parent_result =  | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
631  | 
                  info.parent match {
 | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
632  | 
case None => Result(true, no_heap, 0)  | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
633  | 
case Some(parent) => results(parent)  | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
634  | 
}  | 
| 48547 | 635  | 
val output = output_dir + Path.basic(name)  | 
636  | 
val do_output = build_heap || queue.is_inner(name)  | 
|
637  | 
||
| 
48639
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
638  | 
val (current, heap) =  | 
| 48547 | 639  | 
                {
 | 
| 48548 | 640  | 
                  input_dirs.find(dir => (dir + log_gz(name)).is_file) match {
 | 
| 48547 | 641  | 
case Some(dir) =>  | 
| 
48639
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
642  | 
                      read_stamps(dir + log_gz(name)) match {
 | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
643  | 
case Some((s, h1, h2)) =>  | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
644  | 
val heap = heap_stamp(Some(dir + Path.basic(name)))  | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
645  | 
(s == make_stamp(name) && h1 == parent_result.heap && h2 == heap &&  | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
646  | 
!(do_output && heap == no_heap), heap)  | 
| 
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
647  | 
case None => (false, no_heap)  | 
| 48547 | 648  | 
}  | 
| 
48639
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
649  | 
case None => (false, no_heap)  | 
| 
48504
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
650  | 
}  | 
| 48547 | 651  | 
}  | 
| 
48639
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
652  | 
val all_current = current && parent_result.current  | 
| 
48528
 
784c6f63d79c
proper all_current, which regards parent status as well;
 
wenzelm 
parents: 
48511 
diff
changeset
 | 
653  | 
|
| 48547 | 654  | 
if (all_current)  | 
| 
48639
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
655  | 
loop(pending - name, running, results + (name -> Result(true, heap, 0)))  | 
| 
48678
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
656  | 
                else if (no_build) {
 | 
| 
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
657  | 
                  if (verbose) echo("Skipping " + name + " ...")
 | 
| 
48639
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
658  | 
loop(pending - name, running, results + (name -> Result(false, heap, 1)))  | 
| 
48678
 
ff27af15530c
queue ordering by descending outdegree and timeout;
 
wenzelm 
parents: 
48676 
diff
changeset
 | 
659  | 
}  | 
| 
48639
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
660  | 
                else if (parent_result.rc == 0) {
 | 
| 48547 | 661  | 
echo((if (do_output) "Building " else "Running ") + name + " ...")  | 
| 48674 | 662  | 
val job = new Job(name, info, output, do_output, verbose, browser_info)  | 
663  | 
loop(pending, running + (name -> (parent_result.heap, job)), results)  | 
|
| 48547 | 664  | 
}  | 
665  | 
                else {
 | 
|
666  | 
echo(name + " CANCELLED")  | 
|
| 
48639
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
667  | 
loop(pending - name, running, results + (name -> Result(false, heap, 1)))  | 
| 48547 | 668  | 
}  | 
669  | 
case None => sleep(); loop(pending, running, results)  | 
|
| 48425 | 670  | 
}  | 
671  | 
case None => sleep(); loop(pending, running, results)  | 
|
| 48373 | 672  | 
}  | 
| 48425 | 673  | 
}  | 
674  | 
||
| 48583 | 675  | 
val results =  | 
676  | 
      if (deps.is_empty) {
 | 
|
677  | 
        echo("### Nothing to build")
 | 
|
678  | 
Map.empty  | 
|
679  | 
}  | 
|
680  | 
else loop(queue, Map.empty, Map.empty)  | 
|
681  | 
||
| 
48639
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
682  | 
    val rc = (0 /: results)({ case (rc1, (_, res)) => rc1 max res.rc })
 | 
| 
48504
 
21dfd6c04482
actually check source vs. target stamps, based on information from log files;
 
wenzelm 
parents: 
48494 
diff
changeset
 | 
683  | 
    if (rc != 0 && (verbose || !no_build)) {
 | 
| 48552 | 684  | 
val unfinished =  | 
| 
48639
 
675988e64bf9
store parent heap stamp as well -- needs to be propagated through the build hierarchy;
 
wenzelm 
parents: 
48626 
diff
changeset
 | 
685  | 
(for ((name, res) <- results.iterator if res.rc != 0) yield name).toList.sorted  | 
| 48473 | 686  | 
      echo("Unfinished session(s): " + commas(unfinished))
 | 
687  | 
}  | 
|
688  | 
rc  | 
|
| 48341 | 689  | 
}  | 
690  | 
||
691  | 
||
| 
48346
 
e2382bede914
more general support for Isabelle/Scala command line tools;
 
wenzelm 
parents: 
48344 
diff
changeset
 | 
692  | 
/* command line entry point */  | 
| 48341 | 693  | 
|
694  | 
def main(args: Array[String])  | 
|
695  | 
  {
 | 
|
| 
48346
 
e2382bede914
more general support for Isabelle/Scala command line tools;
 
wenzelm 
parents: 
48344 
diff
changeset
 | 
696  | 
    Command_Line.tool {
 | 
| 
 
e2382bede914
more general support for Isabelle/Scala command line tools;
 
wenzelm 
parents: 
48344 
diff
changeset
 | 
697  | 
      args.toList match {
 | 
| 
 
e2382bede914
more general support for Isabelle/Scala command line tools;
 
wenzelm 
parents: 
48344 
diff
changeset
 | 
698  | 
case  | 
| 
 
e2382bede914
more general support for Isabelle/Scala command line tools;
 
wenzelm 
parents: 
48344 
diff
changeset
 | 
699  | 
Properties.Value.Boolean(all_sessions) ::  | 
| 
48511
 
37999ee01156
remove old output heaps, to ensure that result is valid wrt. check_stamps;
 
wenzelm 
parents: 
48509 
diff
changeset
 | 
700  | 
Properties.Value.Boolean(build_heap) ::  | 
| 48595 | 701  | 
Properties.Value.Boolean(clean_build) ::  | 
| 48425 | 702  | 
Properties.Value.Int(max_jobs) ::  | 
| 48903 | 703  | 
Properties.Value.Boolean(list_files) ::  | 
| 48469 | 704  | 
Properties.Value.Boolean(no_build) ::  | 
| 
48447
 
ef600ce4559c
added system build mode: produce output in ISABELLE_HOME;
 
wenzelm 
parents: 
48425 
diff
changeset
 | 
705  | 
Properties.Value.Boolean(system_mode) ::  | 
| 48425 | 706  | 
Properties.Value.Boolean(verbose) ::  | 
| 
48737
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
707  | 
Command_Line.Chunks(select_dirs, include_dirs, session_groups, build_options, sessions) =>  | 
| 
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
708  | 
val dirs =  | 
| 
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
709  | 
select_dirs.map(d => (true, Path.explode(d))) :::  | 
| 
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
710  | 
include_dirs.map(d => (false, Path.explode(d)))  | 
| 
 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 
wenzelm 
parents: 
48724 
diff
changeset
 | 
711  | 
build(all_sessions, build_heap, clean_build, dirs, session_groups,  | 
| 48903 | 712  | 
max_jobs, list_files, no_build, build_options, system_mode, verbose, sessions)  | 
| 
48346
 
e2382bede914
more general support for Isabelle/Scala command line tools;
 
wenzelm 
parents: 
48344 
diff
changeset
 | 
713  | 
        case _ => error("Bad arguments:\n" + cat_lines(args))
 | 
| 48341 | 714  | 
}  | 
| 
48346
 
e2382bede914
more general support for Isabelle/Scala command line tools;
 
wenzelm 
parents: 
48344 
diff
changeset
 | 
715  | 
}  | 
| 48341 | 716  | 
}  | 
| 48276 | 717  | 
}  | 
718  |