| author | wenzelm | 
| Thu, 20 Jul 2023 12:29:57 +0200 | |
| changeset 78416 | f26eba6281b1 | 
| parent 78379 | f6ec57648894 | 
| child 78421 | fd24f380b588 | 
| permissions | -rw-r--r-- | 
| 72662 | 1  | 
/* Title: Pure/Tools/build_job.scala  | 
2  | 
Author: Makarius  | 
|
3  | 
||
4  | 
Build job running prover process, with rudimentary PIDE session.  | 
|
5  | 
*/  | 
|
6  | 
||
7  | 
package isabelle  | 
|
8  | 
||
9  | 
||
10  | 
import scala.collection.mutable  | 
|
11  | 
||
12  | 
||
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
13  | 
trait Build_Job {
 | 
| 
77486
 
032c76e04475
clarified execution context: main work happens within Future.thread;
 
wenzelm 
parents: 
77485 
diff
changeset
 | 
14  | 
def cancel(): Unit = ()  | 
| 77298 | 15  | 
def is_finished: Boolean = false  | 
| 
77486
 
032c76e04475
clarified execution context: main work happens within Future.thread;
 
wenzelm 
parents: 
77485 
diff
changeset
 | 
16  | 
def join: (Process_Result, SHA1.Shasum) = (Process_Result.undefined, SHA1.no_shasum)  | 
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
17  | 
}  | 
| 
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
18  | 
|
| 75393 | 19  | 
object Build_Job {
 | 
| 77396 | 20  | 
/* build session */  | 
21  | 
||
| 77474 | 22  | 
def start_session(  | 
23  | 
build_context: Build_Process.Context,  | 
|
| 
78237
 
c2c59de57df9
clarified static Build_Process.Context vs. dynamic Build_Process.State;
 
wenzelm 
parents: 
78213 
diff
changeset
 | 
24  | 
session_context: Session_Context,  | 
| 77505 | 25  | 
progress: Progress,  | 
26  | 
log: Logger,  | 
|
| 78372 | 27  | 
server: SSH.Server,  | 
| 77474 | 28  | 
session_background: Sessions.Background,  | 
| 
78237
 
c2c59de57df9
clarified static Build_Process.Context vs. dynamic Build_Process.State;
 
wenzelm 
parents: 
78213 
diff
changeset
 | 
29  | 
sources_shasum: SHA1.Shasum,  | 
| 77474 | 30  | 
input_shasum: SHA1.Shasum,  | 
| 
78237
 
c2c59de57df9
clarified static Build_Process.Context vs. dynamic Build_Process.State;
 
wenzelm 
parents: 
78213 
diff
changeset
 | 
31  | 
node_info: Host.Node_Info,  | 
| 
 
c2c59de57df9
clarified static Build_Process.Context vs. dynamic Build_Process.State;
 
wenzelm 
parents: 
78213 
diff
changeset
 | 
32  | 
store_heap: Boolean  | 
| 77505 | 33  | 
  ): Session_Job = {
 | 
| 78372 | 34  | 
new Session_Job(build_context, session_context, progress, log, server,  | 
| 
78237
 
c2c59de57df9
clarified static Build_Process.Context vs. dynamic Build_Process.State;
 
wenzelm 
parents: 
78213 
diff
changeset
 | 
35  | 
session_background, sources_shasum, input_shasum, node_info, store_heap)  | 
| 77505 | 36  | 
}  | 
| 77474 | 37  | 
|
| 77442 | 38  | 
  object Session_Context {
 | 
39  | 
def load(  | 
|
| 78374 | 40  | 
database_server: Option[SQL.Database],  | 
| 
77529
 
40ccee0fe19a
separate static build_uuid from dynamic worker_uuid, to allow multiple worker processes participate in one build process;
 
wenzelm 
parents: 
77521 
diff
changeset
 | 
41  | 
build_uuid: String,  | 
| 
77444
 
0c704aba71e3
clarified signature: reduce explicit access to static Sessions.Structure;
 
wenzelm 
parents: 
77442 
diff
changeset
 | 
42  | 
name: String,  | 
| 
 
0c704aba71e3
clarified signature: reduce explicit access to static Sessions.Structure;
 
wenzelm 
parents: 
77442 
diff
changeset
 | 
43  | 
deps: List[String],  | 
| 
 
0c704aba71e3
clarified signature: reduce explicit access to static Sessions.Structure;
 
wenzelm 
parents: 
77442 
diff
changeset
 | 
44  | 
ancestors: List[String],  | 
| 
77610
 
3b09ae9e40cb
clarified session prefs (or "options" within the database);
 
wenzelm 
parents: 
77587 
diff
changeset
 | 
45  | 
session_prefs: String,  | 
| 77458 | 46  | 
sources_shasum: SHA1.Shasum,  | 
| 77442 | 47  | 
timeout: Time,  | 
| 78178 | 48  | 
store: Store,  | 
| 78374 | 49  | 
progress: Progress = new Progress  | 
| 77442 | 50  | 
    ): Session_Context = {
 | 
| 77450 | 51  | 
def default: Session_Context =  | 
| 77496 | 52  | 
Session_Context(  | 
| 
77610
 
3b09ae9e40cb
clarified session prefs (or "options" within the database);
 
wenzelm 
parents: 
77587 
diff
changeset
 | 
53  | 
name, deps, ancestors, session_prefs, sources_shasum, timeout,  | 
| 
 
3b09ae9e40cb
clarified session prefs (or "options" within the database);
 
wenzelm 
parents: 
77587 
diff
changeset
 | 
54  | 
Time.zero, Bytes.empty, build_uuid)  | 
| 
77444
 
0c704aba71e3
clarified signature: reduce explicit access to static Sessions.Structure;
 
wenzelm 
parents: 
77442 
diff
changeset
 | 
55  | 
|
| 78374 | 56  | 
      def read(db: SQL.Database): Session_Context = {
 | 
57  | 
        def ignore_error(msg: String) = {
 | 
|
58  | 
progress.echo_warning(  | 
|
59  | 
"Ignoring bad database " + db + " for session " + quote(name) +  | 
|
60  | 
if_proper(msg, ":\n" + msg))  | 
|
61  | 
default  | 
|
62  | 
}  | 
|
63  | 
        try {
 | 
|
64  | 
val command_timings = store.read_command_timings(db, name)  | 
|
65  | 
val elapsed =  | 
|
66  | 
            store.read_session_timing(db, name) match {
 | 
|
67  | 
case Markup.Elapsed(s) => Time.seconds(s)  | 
|
68  | 
case _ => Time.zero  | 
|
69  | 
}  | 
|
70  | 
new Session_Context(  | 
|
71  | 
name, deps, ancestors, session_prefs, sources_shasum, timeout,  | 
|
72  | 
elapsed, command_timings, build_uuid)  | 
|
73  | 
}  | 
|
74  | 
        catch {
 | 
|
75  | 
case ERROR(msg) => ignore_error(msg)  | 
|
76  | 
case exn: java.lang.Error => ignore_error(Exn.message(exn))  | 
|
77  | 
          case _: XML.Error => ignore_error("XML.Error")
 | 
|
78  | 
}  | 
|
79  | 
}  | 
|
80  | 
||
81  | 
      database_server match {
 | 
|
82  | 
case Some(db) => if (store.session_info_exists(db)) read(db) else default  | 
|
83  | 
case None => using_option(store.try_open_database(name))(read) getOrElse default  | 
|
| 77442 | 84  | 
}  | 
85  | 
}  | 
|
86  | 
}  | 
|
87  | 
||
| 77496 | 88  | 
sealed case class Session_Context(  | 
89  | 
name: String,  | 
|
90  | 
deps: List[String],  | 
|
91  | 
ancestors: List[String],  | 
|
| 
77610
 
3b09ae9e40cb
clarified session prefs (or "options" within the database);
 
wenzelm 
parents: 
77587 
diff
changeset
 | 
92  | 
session_prefs: String,  | 
| 77496 | 93  | 
sources_shasum: SHA1.Shasum,  | 
94  | 
timeout: Time,  | 
|
95  | 
old_time: Time,  | 
|
96  | 
old_command_timings_blob: Bytes,  | 
|
| 
77529
 
40ccee0fe19a
separate static build_uuid from dynamic worker_uuid, to allow multiple worker processes participate in one build process;
 
wenzelm 
parents: 
77521 
diff
changeset
 | 
97  | 
build_uuid: String  | 
| 
77652
 
5f706f7c624b
more thorough synchronization of internal "_state" vs. external "_database";
 
wenzelm 
parents: 
77651 
diff
changeset
 | 
98  | 
) extends Library.Named  | 
| 77442 | 99  | 
|
| 77474 | 100  | 
class Session_Job private[Build_Job](  | 
| 77455 | 101  | 
build_context: Build_Process.Context,  | 
| 
78237
 
c2c59de57df9
clarified static Build_Process.Context vs. dynamic Build_Process.State;
 
wenzelm 
parents: 
78213 
diff
changeset
 | 
102  | 
session_context: Session_Context,  | 
| 77505 | 103  | 
progress: Progress,  | 
104  | 
log: Logger,  | 
|
| 78372 | 105  | 
server: SSH.Server,  | 
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
106  | 
session_background: Sessions.Background,  | 
| 
78237
 
c2c59de57df9
clarified static Build_Process.Context vs. dynamic Build_Process.State;
 
wenzelm 
parents: 
78213 
diff
changeset
 | 
107  | 
sources_shasum: SHA1.Shasum,  | 
| 
77460
 
ccca589d7027
tuned signature: support general Build_Job instances;
 
wenzelm 
parents: 
77458 
diff
changeset
 | 
108  | 
input_shasum: SHA1.Shasum,  | 
| 
78237
 
c2c59de57df9
clarified static Build_Process.Context vs. dynamic Build_Process.State;
 
wenzelm 
parents: 
78213 
diff
changeset
 | 
109  | 
node_info: Host.Node_Info,  | 
| 
 
c2c59de57df9
clarified static Build_Process.Context vs. dynamic Build_Process.State;
 
wenzelm 
parents: 
78213 
diff
changeset
 | 
110  | 
store_heap: Boolean  | 
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
111  | 
  ) extends Build_Job {
 | 
| 
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
112  | 
def session_name: String = session_background.session_name  | 
| 
77349
 
5a84de89170d
more operations to support management of jobs, e.g. from external database;
 
wenzelm 
parents: 
77298 
diff
changeset
 | 
113  | 
|
| 
77486
 
032c76e04475
clarified execution context: main work happens within Future.thread;
 
wenzelm 
parents: 
77485 
diff
changeset
 | 
114  | 
private val future_result: Future[(Process_Result, SHA1.Shasum)] =  | 
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
115  | 
      Future.thread("build", uninterruptible = true) {
 | 
| 
78359
 
cb0a90df4871
prefer asynchronous operations: reduce time spent within synchronized_database("Build_Process.start_job");
 
wenzelm 
parents: 
78280 
diff
changeset
 | 
116  | 
val info = session_background.sessions_structure(session_name)  | 
| 
 
cb0a90df4871
prefer asynchronous operations: reduce time spent within synchronized_database("Build_Process.start_job");
 
wenzelm 
parents: 
78280 
diff
changeset
 | 
117  | 
val options = node_info.process_policy(info.options)  | 
| 
 
cb0a90df4871
prefer asynchronous operations: reduce time spent within synchronized_database("Build_Process.start_job");
 
wenzelm 
parents: 
78280 
diff
changeset
 | 
118  | 
|
| 
 
cb0a90df4871
prefer asynchronous operations: reduce time spent within synchronized_database("Build_Process.start_job");
 
wenzelm 
parents: 
78280 
diff
changeset
 | 
119  | 
val store = build_context.store  | 
| 
 
cb0a90df4871
prefer asynchronous operations: reduce time spent within synchronized_database("Build_Process.start_job");
 
wenzelm 
parents: 
78280 
diff
changeset
 | 
120  | 
|
| 78372 | 121  | 
        using_optional(store.maybe_open_database_server(server = server)) { database_server =>
 | 
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
122  | 
|
| 78372 | 123  | 
store.clean_output(database_server, session_name, session_init = true)  | 
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
124  | 
|
| 78372 | 125  | 
val session_sources =  | 
126  | 
Store.Sources.load(session_background.base, cache = store.cache.compress)  | 
|
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
127  | 
|
| 78372 | 128  | 
val env =  | 
129  | 
Isabelle_System.settings(  | 
|
130  | 
              List("ISABELLE_ML_DEBUGGER" -> options.bool("ML_debugger").toString))
 | 
|
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
131  | 
|
| 78372 | 132  | 
val session_heaps =  | 
133  | 
            session_background.info.parent match {
 | 
|
134  | 
case None => Nil  | 
|
135  | 
case Some(logic) => ML_Process.session_heaps(store, session_background, logic = logic)  | 
|
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
136  | 
}  | 
| 
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
137  | 
|
| 78372 | 138  | 
val use_prelude = if (session_heaps.isEmpty) Thy_Header.ml_roots.map(_._1) else Nil  | 
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
139  | 
|
| 78372 | 140  | 
val eval_store =  | 
141  | 
            if (store_heap) {
 | 
|
142  | 
              (if (info.theories.nonEmpty) List("ML_Heap.share_common_data ()") else Nil) :::
 | 
|
143  | 
              List("ML_Heap.save_child " +
 | 
|
144  | 
ML_Syntax.print_string_bytes(File.platform_path(store.output_heap(session_name))))  | 
|
145  | 
}  | 
|
146  | 
else Nil  | 
|
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
147  | 
|
| 78372 | 148  | 
def session_blobs(node_name: Document.Node.Name): List[(Command.Blob, Document.Blobs.Item)] =  | 
149  | 
            session_background.base.theory_load_commands.get(node_name.theory) match {
 | 
|
150  | 
case None => Nil  | 
|
151  | 
case Some(spans) =>  | 
|
152  | 
val syntax = session_background.base.theory_syntax(node_name)  | 
|
153  | 
val master_dir = Path.explode(node_name.master_dir)  | 
|
154  | 
for (span <- spans; file <- span.loaded_files(syntax).files)  | 
|
155  | 
                  yield {
 | 
|
156  | 
val src_path = Path.explode(file)  | 
|
157  | 
val blob_name = Document.Node.Name(File.symbolic_path(master_dir + src_path))  | 
|
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
158  | 
|
| 78372 | 159  | 
val bytes = session_sources(blob_name.node).bytes  | 
160  | 
val text = bytes.text  | 
|
161  | 
val chunk = Symbol.Text_Chunk(text)  | 
|
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
162  | 
|
| 78372 | 163  | 
Command.Blob(blob_name, src_path, Some((SHA1.digest(bytes), chunk))) ->  | 
164  | 
Document.Blobs.Item(bytes, text, chunk, changed = false)  | 
|
165  | 
}  | 
|
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
166  | 
}  | 
| 
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
167  | 
|
| 
77486
 
032c76e04475
clarified execution context: main work happens within Future.thread;
 
wenzelm 
parents: 
77485 
diff
changeset
 | 
168  | 
|
| 78372 | 169  | 
/* session */  | 
| 
77486
 
032c76e04475
clarified execution context: main work happens within Future.thread;
 
wenzelm 
parents: 
77485 
diff
changeset
 | 
170  | 
|
| 78372 | 171  | 
val resources =  | 
172  | 
new Resources(session_background, log = log,  | 
|
173  | 
command_timings =  | 
|
174  | 
Properties.uncompress(session_context.old_command_timings_blob, cache = store.cache))  | 
|
| 77485 | 175  | 
|
| 78372 | 176  | 
val session =  | 
177  | 
            new Session(options, resources) {
 | 
|
178  | 
override val cache: Term.Cache = store.cache  | 
|
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
179  | 
|
| 78372 | 180  | 
override def build_blobs_info(node_name: Document.Node.Name): Command.Blobs_Info =  | 
181  | 
Command.Blobs_Info.make(session_blobs(node_name))  | 
|
182  | 
||
183  | 
override def build_blobs(node_name: Document.Node.Name): Document.Blobs =  | 
|
184  | 
Document.Blobs.make(session_blobs(node_name))  | 
|
185  | 
}  | 
|
186  | 
||
187  | 
          object Build_Session_Errors {
 | 
|
188  | 
private val promise: Promise[List[String]] = Future.promise  | 
|
189  | 
||
190  | 
def result: Exn.Result[List[String]] = promise.join_result  | 
|
191  | 
def cancel(): Unit = promise.cancel()  | 
|
192  | 
            def apply(errs: List[String]): Unit = {
 | 
|
193  | 
              try { promise.fulfill(errs) }
 | 
|
194  | 
              catch { case _: IllegalStateException => }
 | 
|
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
195  | 
}  | 
| 
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
196  | 
}  | 
| 
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
197  | 
|
| 78372 | 198  | 
val export_consumer =  | 
199  | 
Export.consumer(store.open_database(session_name, output = true, server = server),  | 
|
200  | 
store.cache, progress = progress)  | 
|
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
201  | 
|
| 78372 | 202  | 
val stdout = new StringBuilder(1000)  | 
203  | 
val stderr = new StringBuilder(1000)  | 
|
204  | 
val command_timings = new mutable.ListBuffer[Properties.T]  | 
|
205  | 
val theory_timings = new mutable.ListBuffer[Properties.T]  | 
|
206  | 
val session_timings = new mutable.ListBuffer[Properties.T]  | 
|
207  | 
val runtime_statistics = new mutable.ListBuffer[Properties.T]  | 
|
208  | 
val task_statistics = new mutable.ListBuffer[Properties.T]  | 
|
209  | 
||
210  | 
def fun(  | 
|
211  | 
name: String,  | 
|
212  | 
acc: mutable.ListBuffer[Properties.T],  | 
|
213  | 
unapply: Properties.T => Option[Properties.T]  | 
|
214  | 
          ): (String, Session.Protocol_Function) = {
 | 
|
215  | 
name -> ((msg: Prover.Protocol_Output) =>  | 
|
216  | 
              unapply(msg.properties) match {
 | 
|
217  | 
case Some(props) => acc += props; true  | 
|
218  | 
case _ => false  | 
|
219  | 
})  | 
|
| 77485 | 220  | 
}  | 
221  | 
||
| 78372 | 222  | 
          session.init_protocol_handler(new Session.Protocol_Handler {
 | 
223  | 
override def exit(): Unit = Build_Session_Errors.cancel()  | 
|
224  | 
||
225  | 
              private def build_session_finished(msg: Prover.Protocol_Output): Boolean = {
 | 
|
226  | 
val (rc, errors) =  | 
|
227  | 
                  try {
 | 
|
228  | 
                    val (rc, errs) = {
 | 
|
229  | 
import XML.Decode._  | 
|
230  | 
pair(int, list(x => x))(Symbol.decode_yxml(msg.text))  | 
|
231  | 
}  | 
|
232  | 
val errors =  | 
|
233  | 
                      for (err <- errs) yield {
 | 
|
234  | 
val prt = Protocol_Message.expose_no_reports(err)  | 
|
235  | 
Pretty.string_of(prt, metric = Symbol.Metric)  | 
|
236  | 
}  | 
|
237  | 
(rc, errors)  | 
|
238  | 
}  | 
|
239  | 
                  catch { case ERROR(err) => (Process_Result.RC.failure, List(err)) }
 | 
|
240  | 
||
241  | 
                session.protocol_command("Prover.stop", rc.toString)
 | 
|
242  | 
Build_Session_Errors(errors)  | 
|
243  | 
true  | 
|
244  | 
}  | 
|
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
245  | 
|
| 78372 | 246  | 
private def loading_theory(msg: Prover.Protocol_Output): Boolean =  | 
247  | 
                msg.properties match {
 | 
|
248  | 
case Markup.Loading_Theory(Markup.Name(name)) =>  | 
|
249  | 
progress.theory(Progress.Theory(name, session = session_name))  | 
|
250  | 
false  | 
|
251  | 
case _ => false  | 
|
252  | 
}  | 
|
253  | 
||
254  | 
private def export_(msg: Prover.Protocol_Output): Boolean =  | 
|
255  | 
                msg.properties match {
 | 
|
256  | 
case Protocol.Export(args) =>  | 
|
257  | 
export_consumer.make_entry(session_name, args, msg.chunk)  | 
|
258  | 
true  | 
|
259  | 
case _ => false  | 
|
260  | 
}  | 
|
261  | 
||
262  | 
override val functions: Session.Protocol_Functions =  | 
|
263  | 
List(  | 
|
264  | 
Markup.Build_Session_Finished.name -> build_session_finished,  | 
|
265  | 
Markup.Loading_Theory.name -> loading_theory,  | 
|
266  | 
Markup.EXPORT -> export_,  | 
|
267  | 
fun(Markup.Theory_Timing.name, theory_timings, Markup.Theory_Timing.unapply),  | 
|
268  | 
fun(Markup.Session_Timing.name, session_timings, Markup.Session_Timing.unapply),  | 
|
269  | 
fun(Markup.Task_Statistics.name, task_statistics, Markup.Task_Statistics.unapply))  | 
|
270  | 
})  | 
|
271  | 
||
272  | 
          session.command_timings += Session.Consumer("command_timings") {
 | 
|
273  | 
case Session.Command_Timing(props) =>  | 
|
274  | 
              for {
 | 
|
275  | 
elapsed <- Markup.Elapsed.unapply(props)  | 
|
276  | 
elapsed_time = Time.seconds(elapsed)  | 
|
277  | 
                if elapsed_time.is_relevant && elapsed_time >= options.seconds("command_timing_threshold")
 | 
|
278  | 
} command_timings += props.filter(Markup.command_timing_property)  | 
|
279  | 
}  | 
|
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
280  | 
|
| 78372 | 281  | 
          session.runtime_statistics += Session.Consumer("ML_statistics") {
 | 
282  | 
case Session.Runtime_Statistics(props) => runtime_statistics += props  | 
|
283  | 
}  | 
|
284  | 
||
285  | 
          session.finished_theories += Session.Consumer[Document.Snapshot]("finished_theories") {
 | 
|
286  | 
case snapshot =>  | 
|
287  | 
              if (!progress.stopped) {
 | 
|
288  | 
                def export_(name: String, xml: XML.Body, compress: Boolean = true): Unit = {
 | 
|
289  | 
                  if (!progress.stopped) {
 | 
|
290  | 
val theory_name = snapshot.node_name.theory  | 
|
291  | 
val args =  | 
|
292  | 
Protocol.Export.Args(theory_name = theory_name, name = name, compress = compress)  | 
|
293  | 
val body = Bytes(Symbol.encode(YXML.string_of_body(xml)))  | 
|
294  | 
export_consumer.make_entry(session_name, args, body)  | 
|
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
295  | 
}  | 
| 78372 | 296  | 
}  | 
297  | 
def export_text(name: String, text: String, compress: Boolean = true): Unit =  | 
|
298  | 
export_(name, List(XML.Text(text)), compress = compress)  | 
|
299  | 
||
300  | 
                for (command <- snapshot.snippet_command) {
 | 
|
301  | 
export_text(Export.DOCUMENT_ID, command.id.toString, compress = false)  | 
|
302  | 
}  | 
|
303  | 
||
304  | 
export_text(Export.FILES,  | 
|
305  | 
cat_lines(snapshot.node_files.map(name => File.symbolic_path(name.path))),  | 
|
306  | 
compress = false)  | 
|
307  | 
||
308  | 
                for ((blob_name, i) <- snapshot.node_files.tail.zipWithIndex) {
 | 
|
309  | 
val xml = snapshot.switch(blob_name).xml_markup()  | 
|
310  | 
export_(Export.MARKUP + (i + 1), xml)  | 
|
311  | 
}  | 
|
312  | 
export_(Export.MARKUP, snapshot.xml_markup())  | 
|
313  | 
export_(Export.MESSAGES, snapshot.messages.map(_._1))  | 
|
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
314  | 
}  | 
| 
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
315  | 
}  | 
| 78372 | 316  | 
|
317  | 
          session.all_messages += Session.Consumer[Any]("build_session_output") {
 | 
|
318  | 
case msg: Prover.Output =>  | 
|
319  | 
val message = msg.message  | 
|
320  | 
if (msg.is_system) resources.log(Protocol.message_text(message))  | 
|
321  | 
||
322  | 
              if (msg.is_stdout) {
 | 
|
323  | 
stdout ++= Symbol.encode(XML.content(message))  | 
|
324  | 
}  | 
|
325  | 
              else if (msg.is_stderr) {
 | 
|
326  | 
stderr ++= Symbol.encode(XML.content(message))  | 
|
327  | 
}  | 
|
328  | 
              else if (msg.is_exit) {
 | 
|
329  | 
val err =  | 
|
330  | 
"Prover terminated" +  | 
|
331  | 
                    (msg.properties match {
 | 
|
332  | 
case Markup.Process_Result(result) => ": " + result.print_rc  | 
|
333  | 
case _ => ""  | 
|
334  | 
})  | 
|
335  | 
Build_Session_Errors(List(err))  | 
|
336  | 
}  | 
|
337  | 
case _ =>  | 
|
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
338  | 
}  | 
| 
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
339  | 
|
| 78372 | 340  | 
build_context.session_setup(session_name, session)  | 
341  | 
||
342  | 
          val eval_main = Command_Line.ML_tool("Isabelle_Process.init_build ()" :: eval_store)
 | 
|
343  | 
||
| 
77486
 
032c76e04475
clarified execution context: main work happens within Future.thread;
 
wenzelm 
parents: 
77485 
diff
changeset
 | 
344  | 
|
| 78372 | 345  | 
/* process */  | 
346  | 
||
347  | 
val process =  | 
|
348  | 
Isabelle_Process.start(options, session, session_background, session_heaps,  | 
|
349  | 
use_prelude = use_prelude, eval_main = eval_main, cwd = info.dir.file, env = env)  | 
|
350  | 
||
351  | 
val timeout_request: Option[Event_Timer.Request] =  | 
|
352  | 
if (info.timeout_ignored) None  | 
|
353  | 
            else Some(Event_Timer.request(Time.now() + info.timeout) { process.terminate() })
 | 
|
| 
77486
 
032c76e04475
clarified execution context: main work happens within Future.thread;
 
wenzelm 
parents: 
77485 
diff
changeset
 | 
354  | 
|
| 78372 | 355  | 
val build_errors =  | 
356  | 
            Isabelle_Thread.interrupt_handler(_ => process.terminate()) {
 | 
|
357  | 
              Exn.capture { process.await_startup() } match {
 | 
|
358  | 
case Exn.Res(_) =>  | 
|
359  | 
val resources_yxml = resources.init_session_yxml  | 
|
360  | 
val encode_options: XML.Encode.T[Options] =  | 
|
361  | 
options => session.prover_options(options).encode  | 
|
362  | 
val args_yxml =  | 
|
363  | 
YXML.string_of_body(  | 
|
364  | 
                      {
 | 
|
365  | 
import XML.Encode._  | 
|
366  | 
pair(string, list(pair(encode_options, list(pair(string, properties)))))(  | 
|
367  | 
(session_name, info.theories))  | 
|
368  | 
})  | 
|
369  | 
                  session.protocol_command("build_session", resources_yxml, args_yxml)
 | 
|
370  | 
Build_Session_Errors.result  | 
|
371  | 
case Exn.Exn(exn) => Exn.Res(List(Exn.message(exn)))  | 
|
372  | 
}  | 
|
373  | 
}  | 
|
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
374  | 
|
| 78372 | 375  | 
val result0 =  | 
376  | 
            Isabelle_Thread.interrupt_handler(_ => process.terminate()) { process.await_shutdown() }
 | 
|
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
377  | 
|
| 78372 | 378  | 
val was_timeout =  | 
379  | 
            timeout_request match {
 | 
|
380  | 
case None => false  | 
|
381  | 
case Some(request) => !request.cancel()  | 
|
382  | 
}  | 
|
383  | 
||
384  | 
session.stop()  | 
|
385  | 
||
386  | 
val export_errors =  | 
|
387  | 
export_consumer.shutdown(close = true).map(Output.error_message_text)  | 
|
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
388  | 
|
| 78372 | 389  | 
val (document_output, document_errors) =  | 
390  | 
            try {
 | 
|
391  | 
              if (build_errors.isInstanceOf[Exn.Res[_]] && result0.ok && info.documents.nonEmpty) {
 | 
|
| 78379 | 392  | 
                using(Export.open_database_context(store, server = server)) { database_context =>
 | 
| 78372 | 393  | 
val documents =  | 
394  | 
                    using(database_context.open_session(session_background)) {
 | 
|
395  | 
session_context =>  | 
|
396  | 
Document_Build.build_documents(  | 
|
397  | 
Document_Build.context(session_context, progress = progress),  | 
|
398  | 
output_sources = info.document_output,  | 
|
399  | 
output_pdf = info.document_output)  | 
|
400  | 
}  | 
|
401  | 
using(database_context.open_database(session_name, output = true))(session_database =>  | 
|
402  | 
documents.foreach(_.write(session_database.db, session_name)))  | 
|
403  | 
(documents.flatMap(_.log_lines), Nil)  | 
|
404  | 
}  | 
|
| 77484 | 405  | 
}  | 
| 78372 | 406  | 
else (Nil, Nil)  | 
407  | 
}  | 
|
408  | 
            catch {
 | 
|
409  | 
case exn: Document_Build.Build_Error => (exn.log_lines, exn.log_errors)  | 
|
410  | 
case Exn.Interrupt.ERROR(msg) => (Nil, List(msg))  | 
|
411  | 
}  | 
|
| 77485 | 412  | 
|
| 
77486
 
032c76e04475
clarified execution context: main work happens within Future.thread;
 
wenzelm 
parents: 
77485 
diff
changeset
 | 
413  | 
|
| 78372 | 414  | 
/* process result */  | 
| 
77486
 
032c76e04475
clarified execution context: main work happens within Future.thread;
 
wenzelm 
parents: 
77485 
diff
changeset
 | 
415  | 
|
| 78372 | 416  | 
          val result1 = {
 | 
417  | 
val theory_timing =  | 
|
418  | 
theory_timings.iterator.flatMap(  | 
|
419  | 
                {
 | 
|
420  | 
case props @ Markup.Name(name) => Some(name -> props)  | 
|
421  | 
case _ => None  | 
|
422  | 
}).toMap  | 
|
423  | 
val used_theory_timings =  | 
|
424  | 
              for { (name, _) <- session_background.base.used_theories }
 | 
|
425  | 
yield theory_timing.getOrElse(name.theory, Markup.Name(name.theory))  | 
|
| 
77486
 
032c76e04475
clarified execution context: main work happens within Future.thread;
 
wenzelm 
parents: 
77485 
diff
changeset
 | 
426  | 
|
| 78372 | 427  | 
val more_output =  | 
428  | 
Library.trim_line(stdout.toString) ::  | 
|
429  | 
command_timings.toList.map(Protocol.Command_Timing_Marker.apply) :::  | 
|
430  | 
used_theory_timings.map(Protocol.Theory_Timing_Marker.apply) :::  | 
|
431  | 
session_timings.toList.map(Protocol.Session_Timing_Marker.apply) :::  | 
|
432  | 
runtime_statistics.toList.map(Protocol.ML_Statistics_Marker.apply) :::  | 
|
433  | 
task_statistics.toList.map(Protocol.Task_Statistics_Marker.apply) :::  | 
|
434  | 
document_output  | 
|
| 
77486
 
032c76e04475
clarified execution context: main work happens within Future.thread;
 
wenzelm 
parents: 
77485 
diff
changeset
 | 
435  | 
|
| 78372 | 436  | 
result0.output(more_output)  | 
437  | 
.error(Library.trim_line(stderr.toString))  | 
|
438  | 
.errors_rc(export_errors ::: document_errors)  | 
|
439  | 
}  | 
|
| 
77486
 
032c76e04475
clarified execution context: main work happens within Future.thread;
 
wenzelm 
parents: 
77485 
diff
changeset
 | 
440  | 
|
| 78372 | 441  | 
val result2 =  | 
442  | 
            build_errors match {
 | 
|
443  | 
case Exn.Res(build_errs) =>  | 
|
444  | 
val errs = build_errs ::: document_errors  | 
|
445  | 
                if (errs.nonEmpty) {
 | 
|
446  | 
result1.error_rc.output(  | 
|
447  | 
errs.flatMap(s => split_lines(Output.error_message_text(s))) :::  | 
|
448  | 
errs.map(Protocol.Error_Message_Marker.apply))  | 
|
449  | 
}  | 
|
450  | 
else if (progress.stopped && result1.ok) result1.copy(rc = Process_Result.RC.interrupt)  | 
|
451  | 
else result1  | 
|
452  | 
case Exn.Exn(Exn.Interrupt()) =>  | 
|
453  | 
if (result1.ok) result1.copy(rc = Process_Result.RC.interrupt)  | 
|
454  | 
else result1  | 
|
455  | 
case Exn.Exn(exn) => throw exn  | 
|
456  | 
}  | 
|
457  | 
||
458  | 
val process_result =  | 
|
459  | 
if (result2.ok) result2  | 
|
460  | 
            else if (was_timeout) result2.error(Output.error_message_text("Timeout")).timeout_rc
 | 
|
461  | 
            else if (result2.interrupted) result2.error(Output.error_message_text("Interrupt"))
 | 
|
462  | 
else result2  | 
|
463  | 
||
464  | 
||
465  | 
/* output heap */  | 
|
| 
77486
 
032c76e04475
clarified execution context: main work happens within Future.thread;
 
wenzelm 
parents: 
77485 
diff
changeset
 | 
466  | 
|
| 78372 | 467  | 
          val output_shasum = {
 | 
468  | 
val heap = store.output_heap(session_name)  | 
|
469  | 
            if (process_result.ok && store_heap && heap.is_file) {
 | 
|
470  | 
              val slice = Space.MiB(options.real("build_database_slice")).bytes
 | 
|
471  | 
val digest = ML_Heap.store(database_server, session_name, heap, slice)  | 
|
472  | 
SHA1.shasum(digest, session_name)  | 
|
473  | 
}  | 
|
474  | 
else SHA1.no_shasum  | 
|
475  | 
}  | 
|
476  | 
||
477  | 
val log_lines = process_result.out_lines.filterNot(Protocol_Message.Marker.test)  | 
|
| 
77486
 
032c76e04475
clarified execution context: main work happens within Future.thread;
 
wenzelm 
parents: 
77485 
diff
changeset
 | 
478  | 
|
| 78372 | 479  | 
val build_log =  | 
480  | 
Build_Log.Log_File(session_name, process_result.out_lines).  | 
|
481  | 
parse_session_info(  | 
|
482  | 
command_timings = true,  | 
|
483  | 
theory_timings = true,  | 
|
484  | 
ml_statistics = true,  | 
|
485  | 
task_statistics = true)  | 
|
486  | 
||
487  | 
// write log file  | 
|
488  | 
          if (process_result.ok) {
 | 
|
489  | 
File.write_gzip(store.output_log_gz(session_name), terminate_lines(log_lines))  | 
|
490  | 
}  | 
|
491  | 
else File.write(store.output_log(session_name), terminate_lines(log_lines))  | 
|
| 
77486
 
032c76e04475
clarified execution context: main work happens within Future.thread;
 
wenzelm 
parents: 
77485 
diff
changeset
 | 
492  | 
|
| 78372 | 493  | 
// write database  | 
| 78374 | 494  | 
def write_info(db: SQL.Database): Unit =  | 
| 78372 | 495  | 
store.write_session_info(db, session_name, session_sources,  | 
496  | 
build_log =  | 
|
497  | 
                if (process_result.timeout) build_log.error("Timeout") else build_log,
 | 
|
498  | 
build =  | 
|
499  | 
Store.Build_Info(  | 
|
500  | 
sources = sources_shasum,  | 
|
501  | 
input_heaps = input_shasum,  | 
|
502  | 
output_heap = output_shasum,  | 
|
503  | 
process_result.rc,  | 
|
| 78374 | 504  | 
build_context.build_uuid))  | 
505  | 
          database_server match {
 | 
|
506  | 
case Some(db) => write_info(db)  | 
|
507  | 
case None => using(store.open_database(session_name, output = true))(write_info)  | 
|
508  | 
}  | 
|
| 78372 | 509  | 
|
510  | 
// messages  | 
|
511  | 
process_result.err_lines.foreach(progress.echo(_))  | 
|
512  | 
||
513  | 
          if (process_result.ok) {
 | 
|
514  | 
val props = build_log.session_timing  | 
|
515  | 
val threads = Markup.Session_Timing.Threads.unapply(props) getOrElse 1  | 
|
516  | 
val timing = Markup.Timing_Properties.get(props)  | 
|
517  | 
progress.echo(  | 
|
518  | 
              "Timing " + session_name + " (" + threads + " threads, " + timing.message_factor + ")",
 | 
|
519  | 
verbose = true)  | 
|
520  | 
progress.echo(  | 
|
521  | 
              "Finished " + session_name + " (" + process_result.timing.message_resources + ")")
 | 
|
522  | 
}  | 
|
523  | 
          else {
 | 
|
524  | 
progress.echo(  | 
|
525  | 
session_name + " FAILED (see also \"isabelle build_log -H Error " + session_name + "\")")  | 
|
526  | 
            if (!process_result.interrupted) {
 | 
|
527  | 
              val tail = info.options.int("process_output_tail")
 | 
|
528  | 
val suffix = if (tail == 0) log_lines else log_lines.drop(log_lines.length - tail max 0)  | 
|
529  | 
              val prefix = if (log_lines.length == suffix.length) Nil else List("...")
 | 
|
530  | 
progress.echo(Library.trim_line(cat_lines(prefix ::: suffix)))  | 
|
531  | 
}  | 
|
532  | 
}  | 
|
533  | 
||
534  | 
(process_result.copy(out_lines = log_lines), output_shasum)  | 
|
| 
77486
 
032c76e04475
clarified execution context: main work happens within Future.thread;
 
wenzelm 
parents: 
77485 
diff
changeset
 | 
535  | 
}  | 
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
536  | 
}  | 
| 
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
537  | 
|
| 
77486
 
032c76e04475
clarified execution context: main work happens within Future.thread;
 
wenzelm 
parents: 
77485 
diff
changeset
 | 
538  | 
override def cancel(): Unit = future_result.cancel()  | 
| 77298 | 539  | 
override def is_finished: Boolean = future_result.is_finished  | 
| 
77486
 
032c76e04475
clarified execution context: main work happens within Future.thread;
 
wenzelm 
parents: 
77485 
diff
changeset
 | 
540  | 
override def join: (Process_Result, SHA1.Shasum) = future_result.join  | 
| 
77297
 
226a880d0423
clarified types: support a variety of Build_Job instances;
 
wenzelm 
parents: 
77294 
diff
changeset
 | 
541  | 
}  | 
| 
72848
 
d5d0e36eda16
read theory with PIDE markup from session database;
 
wenzelm 
parents: 
72844 
diff
changeset
 | 
542  | 
}  |