| author | wenzelm | 
| Mon, 11 Jul 2011 11:13:33 +0200 | |
| changeset 43746 | a41f618c641d | 
| parent 43745 | 562e35bc351e | 
| child 43747 | 74a9e9c8d5e8 | 
| 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  | 
| 27963 | 3  | 
Options: :folding=explicit:collapseFolds=1:  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
4  | 
|
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
5  | 
Isabelle process management -- always reactive due to multi-threaded I/O.  | 
| 
 
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  | 
|
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
package isabelle  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
9  | 
|
| 
43520
 
cec9b95fa35d
explicit import java.lang.System to prevent odd scope problems;
 
wenzelm 
parents: 
43514 
diff
changeset
 | 
10  | 
import java.lang.System  | 
| 
27955
 
4c32c5e75eca
use java.util.concurrent.LinkedBlockingQueue, which blocks as required;
 
wenzelm 
parents: 
27949 
diff
changeset
 | 
11  | 
import java.util.concurrent.LinkedBlockingQueue  | 
| 28045 | 12  | 
import java.io.{BufferedReader, BufferedWriter, InputStreamReader, OutputStreamWriter,
 | 
| 
38270
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
13  | 
InputStream, OutputStream, BufferedOutputStream, IOException}  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
14  | 
|
| 
32474
 
0818e6b1c8a6
Isabelle_Process: receiver as Actor, not EventBus;
 
wenzelm 
parents: 
32448 
diff
changeset
 | 
15  | 
import scala.actors.Actor  | 
| 
 
0818e6b1c8a6
Isabelle_Process: receiver as Actor, not EventBus;
 
wenzelm 
parents: 
32448 
diff
changeset
 | 
16  | 
import Actor._  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
17  | 
|
| 
27973
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
18  | 
|
| 
32474
 
0818e6b1c8a6
Isabelle_Process: receiver as Actor, not EventBus;
 
wenzelm 
parents: 
32448 
diff
changeset
 | 
19  | 
object Isabelle_Process  | 
| 
 
0818e6b1c8a6
Isabelle_Process: receiver as Actor, not EventBus;
 
wenzelm 
parents: 
32448 
diff
changeset
 | 
20  | 
{
 | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
21  | 
/* results */  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
22  | 
|
| 39525 | 23  | 
object Kind  | 
24  | 
  {
 | 
|
25  | 
val message_markup = Map(  | 
|
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
26  | 
      ('A' : Int) -> Markup.INIT,
 | 
| 
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
27  | 
      ('B' : Int) -> Markup.STATUS,
 | 
| 
38236
 
d8c7be27e01d
explicitly distinguish Output.status (essential feedback) vs. Output.report (useful markup);
 
wenzelm 
parents: 
38231 
diff
changeset
 | 
28  | 
      ('C' : Int) -> Markup.REPORT,
 | 
| 
 
d8c7be27e01d
explicitly distinguish Output.status (essential feedback) vs. Output.report (useful markup);
 
wenzelm 
parents: 
38231 
diff
changeset
 | 
29  | 
      ('D' : Int) -> Markup.WRITELN,
 | 
| 
 
d8c7be27e01d
explicitly distinguish Output.status (essential feedback) vs. Output.report (useful markup);
 
wenzelm 
parents: 
38231 
diff
changeset
 | 
30  | 
      ('E' : Int) -> Markup.TRACING,
 | 
| 
 
d8c7be27e01d
explicitly distinguish Output.status (essential feedback) vs. Output.report (useful markup);
 
wenzelm 
parents: 
38231 
diff
changeset
 | 
31  | 
      ('F' : Int) -> Markup.WARNING,
 | 
| 
43746
 
a41f618c641d
some support for raw messages, which bypass standard Symbol/YXML decoding;
 
wenzelm 
parents: 
43745 
diff
changeset
 | 
32  | 
      ('G' : Int) -> Markup.ERROR,
 | 
| 
 
a41f618c641d
some support for raw messages, which bypass standard Symbol/YXML decoding;
 
wenzelm 
parents: 
43745 
diff
changeset
 | 
33  | 
      ('H' : Int) -> Markup.RAW)
 | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
34  | 
}  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
35  | 
|
| 
43721
 
fad8634cee62
echo prover input via raw_messages, for improved protocol tracing;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
36  | 
abstract class Message  | 
| 
 
fad8634cee62
echo prover input via raw_messages, for improved protocol tracing;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
37  | 
|
| 
 
fad8634cee62
echo prover input via raw_messages, for improved protocol tracing;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
38  | 
class Input(name: String, args: List[String]) extends Message  | 
| 
 
fad8634cee62
echo prover input via raw_messages, for improved protocol tracing;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
39  | 
  {
 | 
| 
 
fad8634cee62
echo prover input via raw_messages, for improved protocol tracing;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
40  | 
override def toString: String =  | 
| 
 
fad8634cee62
echo prover input via raw_messages, for improved protocol tracing;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
41  | 
XML.Elem(Markup(Markup.PROVER_COMMAND, List((Markup.NAME, name))),  | 
| 
 
fad8634cee62
echo prover input via raw_messages, for improved protocol tracing;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
42  | 
args.map(s =>  | 
| 
 
fad8634cee62
echo prover input via raw_messages, for improved protocol tracing;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
43  | 
          List(XML.Text("\n"), XML.elem(Markup.PROVER_ARG, YXML.parse_body(s)))).flatten).toString
 | 
| 
 
fad8634cee62
echo prover input via raw_messages, for improved protocol tracing;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
44  | 
}  | 
| 
 
fad8634cee62
echo prover input via raw_messages, for improved protocol tracing;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
45  | 
|
| 
 
fad8634cee62
echo prover input via raw_messages, for improved protocol tracing;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
46  | 
class Result(val message: XML.Elem) extends Message  | 
| 34100 | 47  | 
  {
 | 
| 
38230
 
ed147003de4b
simplified type XML.Tree: embed Markup directly, avoid slightly odd triple;
 
wenzelm 
parents: 
37712 
diff
changeset
 | 
48  | 
def kind = message.markup.name  | 
| 
 
ed147003de4b
simplified type XML.Tree: embed Markup directly, avoid slightly odd triple;
 
wenzelm 
parents: 
37712 
diff
changeset
 | 
49  | 
def properties = message.markup.properties  | 
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
50  | 
def body = message.body  | 
| 
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
51  | 
|
| 39525 | 52  | 
def is_init = kind == Markup.INIT  | 
53  | 
def is_exit = kind == Markup.EXIT  | 
|
54  | 
def is_stdout = kind == Markup.STDOUT  | 
|
55  | 
def is_system = kind == Markup.SYSTEM  | 
|
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
56  | 
def is_status = kind == Markup.STATUS  | 
| 
38236
 
d8c7be27e01d
explicitly distinguish Output.status (essential feedback) vs. Output.report (useful markup);
 
wenzelm 
parents: 
38231 
diff
changeset
 | 
57  | 
def is_report = kind == Markup.REPORT  | 
| 
43746
 
a41f618c641d
some support for raw messages, which bypass standard Symbol/YXML decoding;
 
wenzelm 
parents: 
43745 
diff
changeset
 | 
58  | 
def is_raw = kind == Markup.RAW  | 
| 39625 | 59  | 
def is_ready = Isar_Document.is_ready(message)  | 
60  | 
def is_syslog = is_init || is_exit || is_system || is_ready  | 
|
| 34100 | 61  | 
|
62  | 
override def toString: String =  | 
|
63  | 
    {
 | 
|
| 
29522
 
793766d4c1b5
moved message markup into Scala layer -- reduced redundancy;
 
wenzelm 
parents: 
29506 
diff
changeset
 | 
64  | 
val res =  | 
| 
38236
 
d8c7be27e01d
explicitly distinguish Output.status (essential feedback) vs. Output.report (useful markup);
 
wenzelm 
parents: 
38231 
diff
changeset
 | 
65  | 
if (is_status || is_report) message.body.map(_.toString).mkString  | 
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
66  | 
else Pretty.string_of(message.body)  | 
| 
38230
 
ed147003de4b
simplified type XML.Tree: embed Markup directly, avoid slightly odd triple;
 
wenzelm 
parents: 
37712 
diff
changeset
 | 
67  | 
if (properties.isEmpty)  | 
| 
29572
 
e3a99d957392
replaced java.util.Properties by plain association list;
 
wenzelm 
parents: 
29522 
diff
changeset
 | 
68  | 
kind.toString + " [[" + res + "]]"  | 
| 
 
e3a99d957392
replaced java.util.Properties by plain association list;
 
wenzelm 
parents: 
29522 
diff
changeset
 | 
69  | 
else  | 
| 
 
e3a99d957392
replaced java.util.Properties by plain association list;
 
wenzelm 
parents: 
29522 
diff
changeset
 | 
70  | 
kind.toString + " " +  | 
| 
38230
 
ed147003de4b
simplified type XML.Tree: embed Markup directly, avoid slightly odd triple;
 
wenzelm 
parents: 
37712 
diff
changeset
 | 
71  | 
          (for ((x, y) <- properties) yield x + "=" + y).mkString("{", ",", "}") + " [[" + res + "]]"
 | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
72  | 
}  | 
| 
27973
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
73  | 
}  | 
| 
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
74  | 
}  | 
| 
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
75  | 
|
| 
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
76  | 
|
| 
43661
 
39fdbd814c7f
quasi-static Isabelle_System -- reduced tendency towards "functorial style";
 
wenzelm 
parents: 
43520 
diff
changeset
 | 
77  | 
class Isabelle_Process(timeout: Time, receiver: Actor, args: String*)  | 
| 29192 | 78  | 
{
 | 
| 31797 | 79  | 
import Isabelle_Process._  | 
| 29194 | 80  | 
|
| 
27973
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
81  | 
|
| 29192 | 82  | 
/* demo constructor */  | 
| 
27973
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
83  | 
|
| 29192 | 84  | 
def this(args: String*) =  | 
| 
43661
 
39fdbd814c7f
quasi-static Isabelle_System -- reduced tendency towards "functorial style";
 
wenzelm 
parents: 
43520 
diff
changeset
 | 
85  | 
    this(Time.seconds(10), actor { loop { react { case res => Console.println(res) } } }, args: _*)
 | 
| 29174 | 86  | 
|
87  | 
||
| 
39626
 
a5d0bcfb95a3
manage persistent syslog via Session, not Isabelle_Process;
 
wenzelm 
parents: 
39625 
diff
changeset
 | 
88  | 
/* results */  | 
| 39573 | 89  | 
|
90  | 
private def system_result(text: String)  | 
|
91  | 
  {
 | 
|
92  | 
receiver ! new Result(XML.Elem(Markup(Markup.SYSTEM, Nil), List(XML.Text(text))))  | 
|
93  | 
}  | 
|
94  | 
||
| 43745 | 95  | 
private val xml_cache = new XML.Cache()  | 
| 39573 | 96  | 
|
97  | 
private def put_result(kind: String, props: List[(String, String)], body: XML.Body)  | 
|
98  | 
  {
 | 
|
| 
39585
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
99  | 
if (kind == Markup.INIT) rm_fifos()  | 
| 
43746
 
a41f618c641d
some support for raw messages, which bypass standard Symbol/YXML decoding;
 
wenzelm 
parents: 
43745 
diff
changeset
 | 
100  | 
if (kind == Markup.RAW)  | 
| 
 
a41f618c641d
some support for raw messages, which bypass standard Symbol/YXML decoding;
 
wenzelm 
parents: 
43745 
diff
changeset
 | 
101  | 
receiver ! new Result(XML.Elem(Markup(kind, props), body))  | 
| 
 
a41f618c641d
some support for raw messages, which bypass standard Symbol/YXML decoding;
 
wenzelm 
parents: 
43745 
diff
changeset
 | 
102  | 
    else {
 | 
| 
 
a41f618c641d
some support for raw messages, which bypass standard Symbol/YXML decoding;
 
wenzelm 
parents: 
43745 
diff
changeset
 | 
103  | 
val msg = XML.Elem(Markup(kind, props), Isar_Document.clean_message(body))  | 
| 
 
a41f618c641d
some support for raw messages, which bypass standard Symbol/YXML decoding;
 
wenzelm 
parents: 
43745 
diff
changeset
 | 
104  | 
xml_cache.cache_tree(msg)((message: XML.Tree) =>  | 
| 
 
a41f618c641d
some support for raw messages, which bypass standard Symbol/YXML decoding;
 
wenzelm 
parents: 
43745 
diff
changeset
 | 
105  | 
receiver ! new Result(message.asInstanceOf[XML.Elem]))  | 
| 
 
a41f618c641d
some support for raw messages, which bypass standard Symbol/YXML decoding;
 
wenzelm 
parents: 
43745 
diff
changeset
 | 
106  | 
}  | 
| 39573 | 107  | 
}  | 
108  | 
||
109  | 
private def put_result(kind: String, text: String)  | 
|
110  | 
  {
 | 
|
| 
43695
 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 
wenzelm 
parents: 
43661 
diff
changeset
 | 
111  | 
put_result(kind, Nil, List(XML.Text(Symbol.decode(text))))  | 
| 39573 | 112  | 
}  | 
113  | 
||
114  | 
||
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
115  | 
/* input actors */  | 
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
116  | 
|
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
117  | 
private case class Input_Text(text: String)  | 
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
118  | 
private case class Input_Chunks(chunks: List[Array[Byte]])  | 
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
119  | 
|
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
120  | 
private case object Close  | 
| 
39585
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
121  | 
private def close(p: (Thread, Actor))  | 
| 
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
122  | 
  {
 | 
| 
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
123  | 
    if (p != null && p._1.isAlive) {
 | 
| 
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
124  | 
p._2 ! Close  | 
| 
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
125  | 
p._1.join  | 
| 
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
126  | 
}  | 
| 
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
127  | 
}  | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
128  | 
|
| 
39585
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
129  | 
@volatile private var standard_input: (Thread, Actor) = null  | 
| 
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
130  | 
@volatile private var command_input: (Thread, Actor) = null  | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
131  | 
|
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
132  | 
|
| 
39585
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
133  | 
/** process manager **/  | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
134  | 
|
| 
43661
 
39fdbd814c7f
quasi-static Isabelle_System -- reduced tendency towards "functorial style";
 
wenzelm 
parents: 
43520 
diff
changeset
 | 
135  | 
private val in_fifo = Isabelle_System.mk_fifo()  | 
| 
 
39fdbd814c7f
quasi-static Isabelle_System -- reduced tendency towards "functorial style";
 
wenzelm 
parents: 
43520 
diff
changeset
 | 
136  | 
private val out_fifo = Isabelle_System.mk_fifo()  | 
| 
 
39fdbd814c7f
quasi-static Isabelle_System -- reduced tendency towards "functorial style";
 
wenzelm 
parents: 
43520 
diff
changeset
 | 
137  | 
  private def rm_fifos() { Isabelle_System.rm_fifo(in_fifo); Isabelle_System.rm_fifo(out_fifo) }
 | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
138  | 
|
| 
39585
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
139  | 
private val process =  | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
140  | 
    try {
 | 
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
141  | 
val cmdline =  | 
| 
43661
 
39fdbd814c7f
quasi-static Isabelle_System -- reduced tendency towards "functorial style";
 
wenzelm 
parents: 
43520 
diff
changeset
 | 
142  | 
        List(Isabelle_System.getenv_strict("ISABELLE_PROCESS"), "-W",
 | 
| 
 
39fdbd814c7f
quasi-static Isabelle_System -- reduced tendency towards "functorial style";
 
wenzelm 
parents: 
43520 
diff
changeset
 | 
143  | 
in_fifo + ":" + out_fifo) ++ args  | 
| 
 
39fdbd814c7f
quasi-static Isabelle_System -- reduced tendency towards "functorial style";
 
wenzelm 
parents: 
43520 
diff
changeset
 | 
144  | 
new Isabelle_System.Managed_Process(true, cmdline: _*)  | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
145  | 
}  | 
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
146  | 
    catch { case e: IOException => rm_fifos(); throw(e) }
 | 
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
147  | 
|
| 
39587
 
f84b70e3bb9c
more reactive handling of Isabelle_Process startup errors;
 
wenzelm 
parents: 
39585 
diff
changeset
 | 
148  | 
val process_result =  | 
| 
 
f84b70e3bb9c
more reactive handling of Isabelle_Process startup errors;
 
wenzelm 
parents: 
39585 
diff
changeset
 | 
149  | 
    Simple_Thread.future("process_result") { process.join }
 | 
| 
 
f84b70e3bb9c
more reactive handling of Isabelle_Process startup errors;
 
wenzelm 
parents: 
39585 
diff
changeset
 | 
150  | 
|
| 
39585
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
151  | 
private def terminate_process()  | 
| 
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
152  | 
  {
 | 
| 
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
153  | 
    try { process.terminate }
 | 
| 
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
154  | 
    catch { case e: IOException => system_result("Failed to terminate Isabelle: " + e.getMessage) }
 | 
| 
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
155  | 
}  | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
156  | 
|
| 
39585
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
157  | 
  private val process_manager = Simple_Thread.fork("process_manager")
 | 
| 39572 | 158  | 
  {
 | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
159  | 
val (startup_failed, startup_output) =  | 
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
160  | 
    {
 | 
| 
40848
 
8662b9b1f123
more abstract/uniform handling of time, preferring seconds as Double;
 
wenzelm 
parents: 
39731 
diff
changeset
 | 
161  | 
val expired = System.currentTimeMillis() + timeout.ms  | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
162  | 
val result = new StringBuilder(100)  | 
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
163  | 
|
| 
39587
 
f84b70e3bb9c
more reactive handling of Isabelle_Process startup errors;
 
wenzelm 
parents: 
39585 
diff
changeset
 | 
164  | 
var finished: Option[Boolean] = None  | 
| 
 
f84b70e3bb9c
more reactive handling of Isabelle_Process startup errors;
 
wenzelm 
parents: 
39585 
diff
changeset
 | 
165  | 
      while (finished.isEmpty && System.currentTimeMillis() <= expired) {
 | 
| 
 
f84b70e3bb9c
more reactive handling of Isabelle_Process startup errors;
 
wenzelm 
parents: 
39585 
diff
changeset
 | 
166  | 
        while (finished.isEmpty && process.stdout.ready) {
 | 
| 
39585
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
167  | 
val c = process.stdout.read  | 
| 
39587
 
f84b70e3bb9c
more reactive handling of Isabelle_Process startup errors;
 
wenzelm 
parents: 
39585 
diff
changeset
 | 
168  | 
if (c == 2) finished = Some(true)  | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
169  | 
else result += c.toChar  | 
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
170  | 
}  | 
| 
39587
 
f84b70e3bb9c
more reactive handling of Isabelle_Process startup errors;
 
wenzelm 
parents: 
39585 
diff
changeset
 | 
171  | 
if (process_result.is_finished) finished = Some(false)  | 
| 
 
f84b70e3bb9c
more reactive handling of Isabelle_Process startup errors;
 
wenzelm 
parents: 
39585 
diff
changeset
 | 
172  | 
else Thread.sleep(10)  | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
173  | 
}  | 
| 39623 | 174  | 
(finished.isEmpty || !finished.get, result.toString.trim)  | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
175  | 
}  | 
| 39572 | 176  | 
system_result(startup_output)  | 
177  | 
||
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
178  | 
    if (startup_failed) {
 | 
| 39632 | 179  | 
put_result(Markup.EXIT, "Return code: 127")  | 
| 
39585
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
180  | 
process.stdin.close  | 
| 
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
181  | 
Thread.sleep(300)  | 
| 
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
182  | 
terminate_process()  | 
| 
39587
 
f84b70e3bb9c
more reactive handling of Isabelle_Process startup errors;
 
wenzelm 
parents: 
39585 
diff
changeset
 | 
183  | 
process_result.join  | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
184  | 
}  | 
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
185  | 
    else {
 | 
| 
39530
 
16adc476348f
Isabelle_Process: more robust rendezvous, even without proper blocking on open (Cygwin);
 
wenzelm 
parents: 
39528 
diff
changeset
 | 
186  | 
// rendezvous  | 
| 
43661
 
39fdbd814c7f
quasi-static Isabelle_System -- reduced tendency towards "functorial style";
 
wenzelm 
parents: 
43520 
diff
changeset
 | 
187  | 
val command_stream = Isabelle_System.fifo_output_stream(in_fifo)  | 
| 
 
39fdbd814c7f
quasi-static Isabelle_System -- reduced tendency towards "functorial style";
 
wenzelm 
parents: 
43520 
diff
changeset
 | 
188  | 
val message_stream = Isabelle_System.fifo_input_stream(out_fifo)  | 
| 
39530
 
16adc476348f
Isabelle_Process: more robust rendezvous, even without proper blocking on open (Cygwin);
 
wenzelm 
parents: 
39528 
diff
changeset
 | 
189  | 
|
| 
39585
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
190  | 
standard_input = stdin_actor()  | 
| 39572 | 191  | 
val stdout = stdout_actor()  | 
| 
39585
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
192  | 
command_input = input_actor(command_stream)  | 
| 39572 | 193  | 
val message = message_actor(message_stream)  | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
194  | 
|
| 
39587
 
f84b70e3bb9c
more reactive handling of Isabelle_Process startup errors;
 
wenzelm 
parents: 
39585 
diff
changeset
 | 
195  | 
val rc = process_result.join  | 
| 39618 | 196  | 
      system_result("process terminated")
 | 
| 
39585
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
197  | 
for ((thread, _) <- List(standard_input, stdout, command_input, message)) thread.join  | 
| 39572 | 198  | 
      system_result("process_manager terminated")
 | 
| 39632 | 199  | 
put_result(Markup.EXIT, "Return code: " + rc.toString)  | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
200  | 
}  | 
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
201  | 
rm_fifos()  | 
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
202  | 
}  | 
| 
27973
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
203  | 
|
| 39572 | 204  | 
|
| 
39585
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
205  | 
/* management methods */  | 
| 
27973
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
206  | 
|
| 
39585
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
207  | 
  def join() { process_manager.join() }
 | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
208  | 
|
| 
38270
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
209  | 
def interrupt()  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
210  | 
  {
 | 
| 
39585
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
211  | 
    try { process.interrupt }
 | 
| 
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
212  | 
    catch { case e: IOException => system_result("Failed to interrupt Isabelle: " + e.getMessage) }
 | 
| 
27973
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
213  | 
}  | 
| 
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
214  | 
|
| 
39585
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
215  | 
def terminate()  | 
| 
38270
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
216  | 
  {
 | 
| 
39585
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
217  | 
close()  | 
| 
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
218  | 
    system_result("Terminating Isabelle process")
 | 
| 
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
219  | 
terminate_process()  | 
| 
27973
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
220  | 
}  | 
| 
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
221  | 
|
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
222  | 
|
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
223  | 
|
| 
38259
 
2b61c5e27399
distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
 
wenzelm 
parents: 
38253 
diff
changeset
 | 
224  | 
/** stream actors **/  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
225  | 
|
| 
38270
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
226  | 
/* raw stdin */  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
227  | 
|
| 39572 | 228  | 
private def stdin_actor(): (Thread, Actor) =  | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
229  | 
  {
 | 
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
230  | 
val name = "standard_input"  | 
| 
38636
 
b7647ca7de5a
module for simplified thread operations (Scala version);
 
wenzelm 
parents: 
38573 
diff
changeset
 | 
231  | 
    Simple_Thread.actor(name) {
 | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
232  | 
var finished = false  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
233  | 
      while (!finished) {
 | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
234  | 
        try {
 | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
235  | 
          //{{{
 | 
| 
38259
 
2b61c5e27399
distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
 
wenzelm 
parents: 
38253 
diff
changeset
 | 
236  | 
          receive {
 | 
| 
38270
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
237  | 
case Input_Text(text) =>  | 
| 
39585
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
238  | 
process.stdin.write(text)  | 
| 
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
239  | 
process.stdin.flush  | 
| 
38259
 
2b61c5e27399
distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
 
wenzelm 
parents: 
38253 
diff
changeset
 | 
240  | 
case Close =>  | 
| 
39585
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
241  | 
process.stdin.close  | 
| 
38259
 
2b61c5e27399
distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
 
wenzelm 
parents: 
38253 
diff
changeset
 | 
242  | 
finished = true  | 
| 
 
2b61c5e27399
distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
 
wenzelm 
parents: 
38253 
diff
changeset
 | 
243  | 
case bad => System.err.println(name + ": ignoring bad message " + bad)  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
244  | 
}  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
245  | 
//}}}  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
246  | 
}  | 
| 39525 | 247  | 
        catch { case e: IOException => system_result(name + ": " + e.getMessage) }
 | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
248  | 
}  | 
| 39525 | 249  | 
system_result(name + " terminated")  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
250  | 
}  | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
251  | 
}  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
252  | 
|
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
253  | 
|
| 
38270
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
254  | 
/* raw stdout */  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
255  | 
|
| 39572 | 256  | 
private def stdout_actor(): (Thread, Actor) =  | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
257  | 
  {
 | 
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
258  | 
val name = "standard_output"  | 
| 
38636
 
b7647ca7de5a
module for simplified thread operations (Scala version);
 
wenzelm 
parents: 
38573 
diff
changeset
 | 
259  | 
    Simple_Thread.actor(name) {
 | 
| 28045 | 260  | 
var result = new StringBuilder(100)  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
261  | 
|
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
262  | 
var finished = false  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
263  | 
      while (!finished) {
 | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
264  | 
        try {
 | 
| 28045 | 265  | 
          //{{{
 | 
266  | 
var c = -1  | 
|
267  | 
var done = false  | 
|
| 
39585
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
268  | 
          while (!done && (result.length == 0 || process.stdout.ready)) {
 | 
| 
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
269  | 
c = process.stdout.read  | 
| 28045 | 270  | 
if (c >= 0) result.append(c.asInstanceOf[Char])  | 
271  | 
else done = true  | 
|
272  | 
}  | 
|
273  | 
          if (result.length > 0) {
 | 
|
| 
38270
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
274  | 
put_result(Markup.STDOUT, result.toString)  | 
| 28045 | 275  | 
result.length = 0  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
276  | 
}  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
277  | 
          else {
 | 
| 
39585
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
278  | 
process.stdout.close  | 
| 28045 | 279  | 
finished = true  | 
| 
38270
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
280  | 
}  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
281  | 
//}}}  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
282  | 
}  | 
| 39525 | 283  | 
        catch { case e: IOException => system_result(name + ": " + e.getMessage) }
 | 
| 
38270
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
284  | 
}  | 
| 39525 | 285  | 
system_result(name + " terminated")  | 
| 
38270
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
286  | 
}  | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
287  | 
}  | 
| 
38270
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
288  | 
|
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
289  | 
|
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
290  | 
/* command input */  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
291  | 
|
| 39572 | 292  | 
private def input_actor(raw_stream: OutputStream): (Thread, Actor) =  | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
293  | 
  {
 | 
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
294  | 
val name = "command_input"  | 
| 
38636
 
b7647ca7de5a
module for simplified thread operations (Scala version);
 
wenzelm 
parents: 
38573 
diff
changeset
 | 
295  | 
    Simple_Thread.actor(name) {
 | 
| 
39530
 
16adc476348f
Isabelle_Process: more robust rendezvous, even without proper blocking on open (Cygwin);
 
wenzelm 
parents: 
39528 
diff
changeset
 | 
296  | 
val stream = new BufferedOutputStream(raw_stream)  | 
| 
38270
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
297  | 
var finished = false  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
298  | 
      while (!finished) {
 | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
299  | 
        try {
 | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
300  | 
          //{{{
 | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
301  | 
          receive {
 | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
302  | 
case Input_Chunks(chunks) =>  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
303  | 
stream.write(Standard_System.string_bytes(  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
304  | 
                  chunks.map(_.length).mkString("", ",", "\n")));
 | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
305  | 
chunks.foreach(stream.write(_));  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
306  | 
stream.flush  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
307  | 
case Close =>  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
308  | 
stream.close  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
309  | 
finished = true  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
310  | 
case bad => System.err.println(name + ": ignoring bad message " + bad)  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
311  | 
}  | 
| 28045 | 312  | 
//}}}  | 
| 27963 | 313  | 
}  | 
| 39525 | 314  | 
        catch { case e: IOException => system_result(name + ": " + e.getMessage) }
 | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
315  | 
}  | 
| 39525 | 316  | 
system_result(name + " terminated")  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
317  | 
}  | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
318  | 
}  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
319  | 
|
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
320  | 
|
| 
38259
 
2b61c5e27399
distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
 
wenzelm 
parents: 
38253 
diff
changeset
 | 
321  | 
/* message output */  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
322  | 
|
| 39572 | 323  | 
private def message_actor(stream: InputStream): (Thread, Actor) =  | 
| 39526 | 324  | 
  {
 | 
325  | 
class EOF extends Exception  | 
|
326  | 
class Protocol_Error(msg: String) extends Exception(msg)  | 
|
| 
38259
 
2b61c5e27399
distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
 
wenzelm 
parents: 
38253 
diff
changeset
 | 
327  | 
|
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
328  | 
val name = "message_output"  | 
| 
38636
 
b7647ca7de5a
module for simplified thread operations (Scala version);
 
wenzelm 
parents: 
38573 
diff
changeset
 | 
329  | 
    Simple_Thread.actor(name) {
 | 
| 34100 | 330  | 
val default_buffer = new Array[Byte](65536)  | 
331  | 
var c = -1  | 
|
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
332  | 
|
| 
43746
 
a41f618c641d
some support for raw messages, which bypass standard Symbol/YXML decoding;
 
wenzelm 
parents: 
43745 
diff
changeset
 | 
333  | 
def read_chunk(decode: Boolean): XML.Body =  | 
| 34100 | 334  | 
      {
 | 
335  | 
        //{{{
 | 
|
336  | 
// chunk size  | 
|
337  | 
var n = 0  | 
|
338  | 
c = stream.read  | 
|
| 39526 | 339  | 
if (c == -1) throw new EOF  | 
| 34100 | 340  | 
        while (48 <= c && c <= 57) {
 | 
341  | 
n = 10 * n + (c - 48)  | 
|
342  | 
c = stream.read  | 
|
343  | 
}  | 
|
344  | 
        if (c != 10) throw new Protocol_Error("bad message chunk header")
 | 
|
345  | 
||
346  | 
// chunk content  | 
|
347  | 
val buf =  | 
|
348  | 
if (n <= default_buffer.size) default_buffer  | 
|
349  | 
else new Array[Byte](n)  | 
|
350  | 
||
351  | 
var i = 0  | 
|
352  | 
var m = 0  | 
|
353  | 
        do {
 | 
|
354  | 
m = stream.read(buf, i, n - i)  | 
|
| 
39731
 
5cb0d7b0d601
bulk read: observe EOF protocol more carefully -- 0 counts as successful read;
 
wenzelm 
parents: 
39632 
diff
changeset
 | 
355  | 
if (m != -1) i += m  | 
| 
 
5cb0d7b0d601
bulk read: observe EOF protocol more carefully -- 0 counts as successful read;
 
wenzelm 
parents: 
39632 
diff
changeset
 | 
356  | 
} while (m != -1 && n > i)  | 
| 34100 | 357  | 
|
358  | 
        if (i != n) throw new Protocol_Error("bad message chunk content")
 | 
|
359  | 
||
| 
43746
 
a41f618c641d
some support for raw messages, which bypass standard Symbol/YXML decoding;
 
wenzelm 
parents: 
43745 
diff
changeset
 | 
360  | 
if (decode)  | 
| 
 
a41f618c641d
some support for raw messages, which bypass standard Symbol/YXML decoding;
 
wenzelm 
parents: 
43745 
diff
changeset
 | 
361  | 
YXML.parse_body_failsafe(Standard_System.decode_chars(Symbol.decode, buf, 0, n))  | 
| 
 
a41f618c641d
some support for raw messages, which bypass standard Symbol/YXML decoding;
 
wenzelm 
parents: 
43745 
diff
changeset
 | 
362  | 
else List(XML.Text(Standard_System.decode_chars(s => s, buf, 0, n).toString))  | 
| 34100 | 363  | 
//}}}  | 
364  | 
}  | 
|
365  | 
||
366  | 
      do {
 | 
|
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
367  | 
        try {
 | 
| 
43746
 
a41f618c641d
some support for raw messages, which bypass standard Symbol/YXML decoding;
 
wenzelm 
parents: 
43745 
diff
changeset
 | 
368  | 
val header = read_chunk(true)  | 
| 
38445
 
ba9ea6b9b75c
simplified internal message format: dropped special Symbol.STX header;
 
wenzelm 
parents: 
38372 
diff
changeset
 | 
369  | 
          header match {
 | 
| 
 
ba9ea6b9b75c
simplified internal message format: dropped special Symbol.STX header;
 
wenzelm 
parents: 
38372 
diff
changeset
 | 
370  | 
case List(XML.Elem(Markup(name, props), Nil))  | 
| 39525 | 371  | 
if name.size == 1 && Kind.message_markup.isDefinedAt(name(0)) =>  | 
| 
43746
 
a41f618c641d
some support for raw messages, which bypass standard Symbol/YXML decoding;
 
wenzelm 
parents: 
43745 
diff
changeset
 | 
372  | 
val kind = Kind.message_markup(name(0))  | 
| 
 
a41f618c641d
some support for raw messages, which bypass standard Symbol/YXML decoding;
 
wenzelm 
parents: 
43745 
diff
changeset
 | 
373  | 
val body = read_chunk(kind != Markup.RAW)  | 
| 
 
a41f618c641d
some support for raw messages, which bypass standard Symbol/YXML decoding;
 
wenzelm 
parents: 
43745 
diff
changeset
 | 
374  | 
put_result(kind, props, body)  | 
| 
 
a41f618c641d
some support for raw messages, which bypass standard Symbol/YXML decoding;
 
wenzelm 
parents: 
43745 
diff
changeset
 | 
375  | 
case _ =>  | 
| 
 
a41f618c641d
some support for raw messages, which bypass standard Symbol/YXML decoding;
 
wenzelm 
parents: 
43745 
diff
changeset
 | 
376  | 
read_chunk(false)  | 
| 
 
a41f618c641d
some support for raw messages, which bypass standard Symbol/YXML decoding;
 
wenzelm 
parents: 
43745 
diff
changeset
 | 
377  | 
              throw new Protocol_Error("bad header: " + header.toString)
 | 
| 28063 | 378  | 
}  | 
| 27963 | 379  | 
}  | 
| 28063 | 380  | 
        catch {
 | 
| 39525 | 381  | 
          case e: IOException => system_result("Cannot read message:\n" + e.getMessage)
 | 
382  | 
          case e: Protocol_Error => system_result("Malformed message:\n" + e.getMessage)
 | 
|
| 39526 | 383  | 
case _: EOF =>  | 
| 28063 | 384  | 
}  | 
| 34100 | 385  | 
} while (c != -1)  | 
386  | 
stream.close  | 
|
387  | 
||
| 39525 | 388  | 
system_result(name + " terminated")  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
389  | 
}  | 
| 39526 | 390  | 
}  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
391  | 
|
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
392  | 
|
| 
38259
 
2b61c5e27399
distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
 
wenzelm 
parents: 
38253 
diff
changeset
 | 
393  | 
/** main methods **/  | 
| 
 
2b61c5e27399
distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
 
wenzelm 
parents: 
38253 
diff
changeset
 | 
394  | 
|
| 
39585
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
395  | 
def input_raw(text: String): Unit = standard_input._2 ! Input_Text(text)  | 
| 
38259
 
2b61c5e27399
distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
 
wenzelm 
parents: 
38253 
diff
changeset
 | 
396  | 
|
| 
38372
 
e753f71b6b34
added Isabelle_Process.input_bytes, which avoids the somewhat slow Standard_System.string_bytes (just in case someone wants to stream raw data at 250MB/s);
 
wenzelm 
parents: 
38270 
diff
changeset
 | 
397  | 
def input_bytes(name: String, args: Array[Byte]*): Unit =  | 
| 
39585
 
00be8711082f
main Isabelle_Process via Isabelle_System.Managed_Process;
 
wenzelm 
parents: 
39575 
diff
changeset
 | 
398  | 
command_input._2 ! Input_Chunks(Standard_System.string_bytes(name) :: args.toList)  | 
| 
38372
 
e753f71b6b34
added Isabelle_Process.input_bytes, which avoids the somewhat slow Standard_System.string_bytes (just in case someone wants to stream raw data at 250MB/s);
 
wenzelm 
parents: 
38270 
diff
changeset
 | 
399  | 
|
| 
38270
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38262 
diff
changeset
 | 
400  | 
def input(name: String, args: String*): Unit =  | 
| 
43721
 
fad8634cee62
echo prover input via raw_messages, for improved protocol tracing;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
401  | 
  {
 | 
| 
 
fad8634cee62
echo prover input via raw_messages, for improved protocol tracing;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
402  | 
receiver ! new Input(name, args.toList)  | 
| 
38372
 
e753f71b6b34
added Isabelle_Process.input_bytes, which avoids the somewhat slow Standard_System.string_bytes (just in case someone wants to stream raw data at 250MB/s);
 
wenzelm 
parents: 
38270 
diff
changeset
 | 
403  | 
input_bytes(name, args.map(Standard_System.string_bytes): _*)  | 
| 
43721
 
fad8634cee62
echo prover input via raw_messages, for improved protocol tracing;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
404  | 
}  | 
| 
38259
 
2b61c5e27399
distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
 
wenzelm 
parents: 
38253 
diff
changeset
 | 
405  | 
|
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39526 
diff
changeset
 | 
406  | 
  def close(): Unit = { close(command_input); close(standard_input) }
 | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
407  | 
}  |