| author | wenzelm | 
| Sun, 07 Feb 2021 16:31:43 +0100 | |
| changeset 73228 | 0575cfd2ecfc | 
| parent 73031 | f93f0597f4fb | 
| child 73340 | 0ffcad1f6130 | 
| permissions | -rw-r--r-- | 
| 43283 | 1 | /* Title: Pure/System/isabelle_process.scala | 
| 27949 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 wenzelm parents: diff
changeset | 2 | Author: Makarius | 
| 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 wenzelm parents: diff
changeset | 3 | |
| 57916 | 4 | Isabelle process wrapper. | 
| 27949 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 wenzelm parents: diff
changeset | 5 | */ | 
| 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 wenzelm parents: diff
changeset | 6 | |
| 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 wenzelm parents: diff
changeset | 7 | package isabelle | 
| 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 wenzelm parents: diff
changeset | 8 | |
| 55618 | 9 | |
| 65310 | 10 | import java.io.{File => JFile}
 | 
| 11 | ||
| 12 | ||
| 57917 | 13 | object Isabelle_Process | 
| 14 | {
 | |
| 71597 | 15 | def apply( | 
| 71594 | 16 | session: Session, | 
| 65216 | 17 | options: Options, | 
| 71594 | 18 | sessions_structure: Sessions.Structure, | 
| 71598 | 19 | store: Sessions.Store, | 
| 65216 | 20 | logic: String = "", | 
| 71639 
ec84f542e411
clarified signature of ML_Process vs. Isabelle_Process: proper support for "isabelle build -P -b";
 wenzelm parents: 
71607diff
changeset | 21 | raw_ml_system: Boolean = false, | 
| 
ec84f542e411
clarified signature of ML_Process vs. Isabelle_Process: proper support for "isabelle build -P -b";
 wenzelm parents: 
71607diff
changeset | 22 | use_prelude: List[String] = Nil, | 
| 
ec84f542e411
clarified signature of ML_Process vs. Isabelle_Process: proper support for "isabelle build -P -b";
 wenzelm parents: 
71607diff
changeset | 23 | eval_main: String = "", | 
| 65216 | 24 | modes: List[String] = Nil, | 
| 65310 | 25 | cwd: JFile = null, | 
| 71604 | 26 | env: Map[String, String] = Isabelle_System.settings()): Isabelle_Process = | 
| 65216 | 27 |   {
 | 
| 62556 
c115e69f457f
more abstract Session.start, without prover command-line;
 wenzelm parents: 
62555diff
changeset | 28 | val channel = System_Channel() | 
| 
c115e69f457f
more abstract Session.start, without prover command-line;
 wenzelm parents: 
62555diff
changeset | 29 | val process = | 
| 57916 | 30 |       try {
 | 
| 69572 
09a6a7c04b45
more robust system channel via options that are private to the user;
 wenzelm parents: 
68209diff
changeset | 31 | val channel_options = | 
| 
09a6a7c04b45
more robust system channel via options that are private to the user;
 wenzelm parents: 
68209diff
changeset | 32 |           options.string.update("system_channel_address", channel.address).
 | 
| 
09a6a7c04b45
more robust system channel via options that are private to the user;
 wenzelm parents: 
68209diff
changeset | 33 |             string.update("system_channel_password", channel.password)
 | 
| 71598 | 34 | ML_Process(channel_options, sessions_structure, store, | 
| 71639 
ec84f542e411
clarified signature of ML_Process vs. Isabelle_Process: proper support for "isabelle build -P -b";
 wenzelm parents: 
71607diff
changeset | 35 | logic = logic, raw_ml_system = raw_ml_system, | 
| 
ec84f542e411
clarified signature of ML_Process vs. Isabelle_Process: proper support for "isabelle build -P -b";
 wenzelm parents: 
71607diff
changeset | 36 | use_prelude = use_prelude, eval_main = eval_main, | 
| 
ec84f542e411
clarified signature of ML_Process vs. Isabelle_Process: proper support for "isabelle build -P -b";
 wenzelm parents: 
71607diff
changeset | 37 | modes = modes, cwd = cwd, env = env) | 
| 57916 | 38 | } | 
| 69572 
09a6a7c04b45
more robust system channel via options that are private to the user;
 wenzelm parents: 
68209diff
changeset | 39 |       catch { case exn @ ERROR(_) => channel.shutdown(); throw exn }
 | 
| 62556 
c115e69f457f
more abstract Session.start, without prover command-line;
 wenzelm parents: 
62555diff
changeset | 40 | process.stdin.close | 
| 57917 | 41 | |
| 71604 | 42 | new Isabelle_Process(session, channel, process) | 
| 57917 | 43 | } | 
| 57916 | 44 | } | 
| 71604 | 45 | |
| 46 | class Isabelle_Process private(session: Session, channel: System_Channel, process: Bash.Process) | |
| 47 | {
 | |
| 71607 | 48 | private val startup = Future.promise[String] | 
| 49 | private val terminated = Future.promise[Process_Result] | |
| 71604 | 50 | |
| 51 | session.phase_changed += | |
| 52 |     Session.Consumer(getClass.getName) {
 | |
| 53 | case Session.Ready => | |
| 71607 | 54 |         startup.fulfill("")
 | 
| 55 | case Session.Terminated(result) => | |
| 56 |         if (!result.ok && !startup.is_finished) {
 | |
| 57 | val syslog = session.syslog_content() | |
| 58 | val err = "Session startup failed" + (if (syslog.isEmpty) "" else ":\n" + syslog) | |
| 59 | startup.fulfill(err) | |
| 60 | } | |
| 61 | terminated.fulfill(result) | |
| 71604 | 62 | case _ => | 
| 63 | } | |
| 64 | ||
| 73031 
f93f0597f4fb
clarified signature: absorb XZ.Cache into XML.Cache;
 wenzelm parents: 
71718diff
changeset | 65 | session.start(receiver => new Prover(receiver, session.cache, channel, process)) | 
| 71607 | 66 | |
| 67 | def await_startup(): Isabelle_Process = | |
| 68 |     startup.join match {
 | |
| 69 | case "" => this | |
| 70 | case err => session.stop(); error(err) | |
| 71604 | 71 | } | 
| 72 | ||
| 71667 
4d2de35214c5
proper treatment of protocol exceptions and prover termination: avoid session.stop while saving image;
 wenzelm parents: 
71639diff
changeset | 73 | def await_shutdown(): Process_Result = | 
| 
4d2de35214c5
proper treatment of protocol exceptions and prover termination: avoid session.stop while saving image;
 wenzelm parents: 
71639diff
changeset | 74 |   {
 | 
| 
4d2de35214c5
proper treatment of protocol exceptions and prover termination: avoid session.stop while saving image;
 wenzelm parents: 
71639diff
changeset | 75 | val result = terminated.join | 
| 
4d2de35214c5
proper treatment of protocol exceptions and prover termination: avoid session.stop while saving image;
 wenzelm parents: 
71639diff
changeset | 76 | session.stop() | 
| 
4d2de35214c5
proper treatment of protocol exceptions and prover termination: avoid session.stop while saving image;
 wenzelm parents: 
71639diff
changeset | 77 | result | 
| 
4d2de35214c5
proper treatment of protocol exceptions and prover termination: avoid session.stop while saving image;
 wenzelm parents: 
71639diff
changeset | 78 | } | 
| 71718 | 79 | |
| 80 |   def terminate { process.terminate }
 | |
| 71667 
4d2de35214c5
proper treatment of protocol exceptions and prover termination: avoid session.stop while saving image;
 wenzelm parents: 
71639diff
changeset | 81 | } |