| author | wenzelm | 
| Tue, 10 Aug 2010 12:09:53 +0200 | |
| changeset 38258 | dd7dcb9b2637 | 
| parent 38253 | 3d4e521014f7 | 
| child 38259 | 2b61c5e27399 | 
| permissions | -rw-r--r-- | 
| 
30173
 
eabece26b89b
moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
 
wenzelm 
parents: 
29648 
diff
changeset
 | 
1  | 
/* Title: Pure/System/isabelle_process.ML  | 
| 
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  | 
|
| 
27955
 
4c32c5e75eca
use java.util.concurrent.LinkedBlockingQueue, which blocks as required;
 
wenzelm 
parents: 
27949 
diff
changeset
 | 
10  | 
import java.util.concurrent.LinkedBlockingQueue  | 
| 28045 | 11  | 
import java.io.{BufferedReader, BufferedWriter, InputStreamReader, OutputStreamWriter,
 | 
| 28056 | 12  | 
InputStream, OutputStream, IOException}  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
13  | 
|
| 
32474
 
0818e6b1c8a6
Isabelle_Process: receiver as Actor, not EventBus;
 
wenzelm 
parents: 
32448 
diff
changeset
 | 
14  | 
import scala.actors.Actor  | 
| 
 
0818e6b1c8a6
Isabelle_Process: receiver as Actor, not EventBus;
 
wenzelm 
parents: 
32448 
diff
changeset
 | 
15  | 
import Actor._  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
16  | 
|
| 
27973
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
17  | 
|
| 
32474
 
0818e6b1c8a6
Isabelle_Process: receiver as Actor, not EventBus;
 
wenzelm 
parents: 
32448 
diff
changeset
 | 
18  | 
object Isabelle_Process  | 
| 
 
0818e6b1c8a6
Isabelle_Process: receiver as Actor, not EventBus;
 
wenzelm 
parents: 
32448 
diff
changeset
 | 
19  | 
{
 | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
20  | 
/* results */  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
21  | 
|
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
22  | 
  object Kind {
 | 
| 
29522
 
793766d4c1b5
moved message markup into Scala layer -- reduced redundancy;
 
wenzelm 
parents: 
29506 
diff
changeset
 | 
23  | 
// message markup  | 
| 
 
793766d4c1b5
moved message markup into Scala layer -- reduced redundancy;
 
wenzelm 
parents: 
29506 
diff
changeset
 | 
24  | 
val markup = Map(  | 
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
25  | 
      ('A' : Int) -> Markup.INIT,
 | 
| 
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
26  | 
      ('B' : Int) -> Markup.STATUS,
 | 
| 
38236
 
d8c7be27e01d
explicitly distinguish Output.status (essential feedback) vs. Output.report (useful markup);
 
wenzelm 
parents: 
38231 
diff
changeset
 | 
27  | 
      ('C' : Int) -> Markup.REPORT,
 | 
| 
 
d8c7be27e01d
explicitly distinguish Output.status (essential feedback) vs. Output.report (useful markup);
 
wenzelm 
parents: 
38231 
diff
changeset
 | 
28  | 
      ('D' : Int) -> Markup.WRITELN,
 | 
| 
 
d8c7be27e01d
explicitly distinguish Output.status (essential feedback) vs. Output.report (useful markup);
 
wenzelm 
parents: 
38231 
diff
changeset
 | 
29  | 
      ('E' : Int) -> Markup.TRACING,
 | 
| 
 
d8c7be27e01d
explicitly distinguish Output.status (essential feedback) vs. Output.report (useful markup);
 
wenzelm 
parents: 
38231 
diff
changeset
 | 
30  | 
      ('F' : Int) -> Markup.WARNING,
 | 
| 
 
d8c7be27e01d
explicitly distinguish Output.status (essential feedback) vs. Output.report (useful markup);
 
wenzelm 
parents: 
38231 
diff
changeset
 | 
31  | 
      ('G' : Int) -> Markup.ERROR,
 | 
| 
 
d8c7be27e01d
explicitly distinguish Output.status (essential feedback) vs. Output.report (useful markup);
 
wenzelm 
parents: 
38231 
diff
changeset
 | 
32  | 
      ('H' : Int) -> Markup.DEBUG)
 | 
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
33  | 
def is_raw(kind: String) =  | 
| 
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
34  | 
kind == Markup.STDOUT  | 
| 
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
35  | 
def is_control(kind: String) =  | 
| 
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
36  | 
kind == Markup.SYSTEM ||  | 
| 
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
37  | 
kind == Markup.SIGNAL ||  | 
| 
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
38  | 
kind == Markup.EXIT  | 
| 
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
39  | 
def is_system(kind: String) =  | 
| 
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
40  | 
kind == Markup.SYSTEM ||  | 
| 
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
41  | 
kind == Markup.STDIN ||  | 
| 
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
42  | 
kind == Markup.SIGNAL ||  | 
| 
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
43  | 
kind == Markup.EXIT ||  | 
| 
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
44  | 
kind == Markup.STATUS  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
45  | 
}  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
46  | 
|
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
47  | 
class Result(val message: XML.Elem)  | 
| 34100 | 48  | 
  {
 | 
| 
38230
 
ed147003de4b
simplified type XML.Tree: embed Markup directly, avoid slightly odd triple;
 
wenzelm 
parents: 
37712 
diff
changeset
 | 
49  | 
def kind = message.markup.name  | 
| 
 
ed147003de4b
simplified type XML.Tree: embed Markup directly, avoid slightly odd triple;
 
wenzelm 
parents: 
37712 
diff
changeset
 | 
50  | 
def properties = message.markup.properties  | 
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
51  | 
def body = message.body  | 
| 
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
52  | 
|
| 
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
53  | 
def is_raw = Kind.is_raw(kind)  | 
| 
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
54  | 
def is_control = Kind.is_control(kind)  | 
| 
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
55  | 
def is_system = Kind.is_system(kind)  | 
| 
 
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  | 
| 38231 | 58  | 
def is_ready = is_status && body == List(XML.Elem(Markup.Ready, Nil))  | 
| 34100 | 59  | 
|
60  | 
override def toString: String =  | 
|
61  | 
    {
 | 
|
| 
29522
 
793766d4c1b5
moved message markup into Scala layer -- reduced redundancy;
 
wenzelm 
parents: 
29506 
diff
changeset
 | 
62  | 
val res =  | 
| 
38236
 
d8c7be27e01d
explicitly distinguish Output.status (essential feedback) vs. Output.report (useful markup);
 
wenzelm 
parents: 
38231 
diff
changeset
 | 
63  | 
if (is_status || is_report) message.body.map(_.toString).mkString  | 
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
64  | 
else Pretty.string_of(message.body)  | 
| 
38230
 
ed147003de4b
simplified type XML.Tree: embed Markup directly, avoid slightly odd triple;
 
wenzelm 
parents: 
37712 
diff
changeset
 | 
65  | 
if (properties.isEmpty)  | 
| 
29572
 
e3a99d957392
replaced java.util.Properties by plain association list;
 
wenzelm 
parents: 
29522 
diff
changeset
 | 
66  | 
kind.toString + " [[" + res + "]]"  | 
| 
 
e3a99d957392
replaced java.util.Properties by plain association list;
 
wenzelm 
parents: 
29522 
diff
changeset
 | 
67  | 
else  | 
| 
 
e3a99d957392
replaced java.util.Properties by plain association list;
 
wenzelm 
parents: 
29522 
diff
changeset
 | 
68  | 
kind.toString + " " +  | 
| 
38230
 
ed147003de4b
simplified type XML.Tree: embed Markup directly, avoid slightly odd triple;
 
wenzelm 
parents: 
37712 
diff
changeset
 | 
69  | 
          (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
 | 
70  | 
}  | 
| 34109 | 71  | 
|
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
72  | 
def cache(c: XML.Cache): Result = new Result(c.cache_tree(message).asInstanceOf[XML.Elem])  | 
| 
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  | 
|
| 34100 | 77  | 
class Isabelle_Process(system: Isabelle_System, 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*) =  | 
| 
32474
 
0818e6b1c8a6
Isabelle_Process: receiver as Actor, not EventBus;
 
wenzelm 
parents: 
32448 
diff
changeset
 | 
85  | 
this(new Isabelle_System,  | 
| 
34213
 
9e86c1ca6e51
removed obsolete version check -- sanity delegated to Isabelle_System;
 
wenzelm 
parents: 
34201 
diff
changeset
 | 
86  | 
      actor { loop { react { case res => Console.println(res) } } }, args: _*)
 | 
| 29174 | 87  | 
|
88  | 
||
| 
27973
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
89  | 
/* process information */  | 
| 
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
90  | 
|
| 29648 | 91  | 
@volatile private var proc: Process = null  | 
92  | 
@volatile private var closing = false  | 
|
93  | 
@volatile private var pid: String = null  | 
|
| 
27973
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
94  | 
|
| 
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
95  | 
|
| 
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
96  | 
/* results */  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
97  | 
|
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
98  | 
private def put_result(kind: String, props: List[(String, String)], body: List[XML.Tree])  | 
| 
29572
 
e3a99d957392
replaced java.util.Properties by plain association list;
 
wenzelm 
parents: 
29522 
diff
changeset
 | 
99  | 
  {
 | 
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
100  | 
    if (kind == Markup.INIT) {
 | 
| 
34214
 
99eefb83a35d
simplified init message -- removed redundant session property;
 
wenzelm 
parents: 
34213 
diff
changeset
 | 
101  | 
for ((Markup.PID, p) <- props) pid = p  | 
| 27963 | 102  | 
}  | 
| 
38230
 
ed147003de4b
simplified type XML.Tree: embed Markup directly, avoid slightly odd triple;
 
wenzelm 
parents: 
37712 
diff
changeset
 | 
103  | 
receiver ! new Result(XML.Elem(Markup(kind, props), body))  | 
| 34100 | 104  | 
}  | 
105  | 
||
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
106  | 
private def put_result(kind: String, text: String)  | 
| 34100 | 107  | 
  {
 | 
108  | 
put_result(kind, Nil, List(XML.Text(system.symbols.decode(text))))  | 
|
| 27992 | 109  | 
}  | 
110  | 
||
| 
27973
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
111  | 
|
| 
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
112  | 
/* signals */  | 
| 
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
113  | 
|
| 
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
114  | 
  def interrupt() = synchronized {
 | 
| 
27993
 
6dd90ef9f927
simplified exceptions: use plain error function / RuntimeException;
 
wenzelm 
parents: 
27992 
diff
changeset
 | 
115  | 
    if (proc == null) error("Cannot interrupt Isabelle: no process")
 | 
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
116  | 
if (pid == null) put_result(Markup.SYSTEM, "Cannot interrupt: unknown pid")  | 
| 
27973
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
117  | 
    else {
 | 
| 
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
118  | 
      try {
 | 
| 34100 | 119  | 
if (system.execute(true, "kill", "-INT", pid).waitFor == 0)  | 
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
120  | 
put_result(Markup.SIGNAL, "INT")  | 
| 27990 | 121  | 
else  | 
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
122  | 
put_result(Markup.SYSTEM, "Cannot interrupt: kill command failed")  | 
| 
27973
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
123  | 
}  | 
| 
27993
 
6dd90ef9f927
simplified exceptions: use plain error function / RuntimeException;
 
wenzelm 
parents: 
27992 
diff
changeset
 | 
124  | 
      catch { case e: IOException => error("Cannot interrupt Isabelle: " + e.getMessage) }
 | 
| 
27973
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
125  | 
}  | 
| 
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
126  | 
}  | 
| 
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
127  | 
|
| 
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
128  | 
  def kill() = synchronized {
 | 
| 
27993
 
6dd90ef9f927
simplified exceptions: use plain error function / RuntimeException;
 
wenzelm 
parents: 
27992 
diff
changeset
 | 
129  | 
    if (proc == 0) error("Cannot kill Isabelle: no process")
 | 
| 
27973
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
130  | 
    else {
 | 
| 
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
131  | 
try_close()  | 
| 37132 | 132  | 
Thread.sleep(500) // FIXME property!?  | 
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
133  | 
put_result(Markup.SIGNAL, "KILL")  | 
| 
27973
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
134  | 
proc.destroy  | 
| 
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
135  | 
proc = null  | 
| 27990 | 136  | 
pid = null  | 
| 
27973
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
137  | 
}  | 
| 
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
138  | 
}  | 
| 
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
139  | 
|
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
140  | 
|
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
141  | 
/* output being piped into the process */  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
142  | 
|
| 27990 | 143  | 
private val output = new LinkedBlockingQueue[String]  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
144  | 
|
| 28045 | 145  | 
  private def output_raw(text: String) = synchronized {
 | 
| 
27993
 
6dd90ef9f927
simplified exceptions: use plain error function / RuntimeException;
 
wenzelm 
parents: 
27992 
diff
changeset
 | 
146  | 
    if (proc == null) error("Cannot output to Isabelle: no process")
 | 
| 
 
6dd90ef9f927
simplified exceptions: use plain error function / RuntimeException;
 
wenzelm 
parents: 
27992 
diff
changeset
 | 
147  | 
    if (closing) error("Cannot output to Isabelle: already closing")
 | 
| 
27955
 
4c32c5e75eca
use java.util.concurrent.LinkedBlockingQueue, which blocks as required;
 
wenzelm 
parents: 
27949 
diff
changeset
 | 
148  | 
output.put(text)  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
149  | 
}  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
150  | 
|
| 28303 | 151  | 
def output_sync(text: String) =  | 
| 28045 | 152  | 
    output_raw(" \\<^sync>\n; " + text + " \\<^sync>;\n")
 | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
153  | 
|
| 27963 | 154  | 
|
155  | 
def command(text: String) =  | 
|
| 
32448
 
a89f876731c5
moved Pure/Tools/isabelle_syntax.scala to Pure/System/isabelle_syntax.scala;
 
wenzelm 
parents: 
31797 
diff
changeset
 | 
156  | 
    output_sync("Isabelle.command " + Isabelle_Syntax.encode_string(text))
 | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
157  | 
|
| 
29572
 
e3a99d957392
replaced java.util.Properties by plain association list;
 
wenzelm 
parents: 
29522 
diff
changeset
 | 
158  | 
def command(props: List[(String, String)], text: String) =  | 
| 
32448
 
a89f876731c5
moved Pure/Tools/isabelle_syntax.scala to Pure/System/isabelle_syntax.scala;
 
wenzelm 
parents: 
31797 
diff
changeset
 | 
159  | 
    output_sync("Isabelle.command " + Isabelle_Syntax.encode_properties(props) + " " +
 | 
| 
 
a89f876731c5
moved Pure/Tools/isabelle_syntax.scala to Pure/System/isabelle_syntax.scala;
 
wenzelm 
parents: 
31797 
diff
changeset
 | 
160  | 
Isabelle_Syntax.encode_string(text))  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
161  | 
|
| 37712 | 162  | 
def ML_val(text: String) =  | 
| 
32448
 
a89f876731c5
moved Pure/Tools/isabelle_syntax.scala to Pure/System/isabelle_syntax.scala;
 
wenzelm 
parents: 
31797 
diff
changeset
 | 
163  | 
    output_sync("ML_val " + Isabelle_Syntax.encode_string(text))
 | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
164  | 
|
| 37712 | 165  | 
def ML_command(text: String) =  | 
166  | 
    output_sync("ML_command " + Isabelle_Syntax.encode_string(text))
 | 
|
167  | 
||
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
168  | 
  def close() = synchronized {    // FIXME watchdog/timeout
 | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
169  | 
    output_raw("\u0000")
 | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
170  | 
closing = true  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
171  | 
}  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
172  | 
|
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
173  | 
  def try_close() = synchronized {
 | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
174  | 
    if (proc != null && !closing) {
 | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
175  | 
      try { close() }
 | 
| 
27993
 
6dd90ef9f927
simplified exceptions: use plain error function / RuntimeException;
 
wenzelm 
parents: 
27992 
diff
changeset
 | 
176  | 
      catch { case _: RuntimeException => }
 | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
177  | 
}  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
178  | 
}  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
179  | 
|
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
180  | 
|
| 
38253
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
181  | 
/* commands */  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
182  | 
|
| 
38253
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
183  | 
  private class Command_Thread(fifo: String) extends Thread("isabelle: commands")
 | 
| 
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
184  | 
  {
 | 
| 
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
185  | 
override def run()  | 
| 
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
186  | 
    {
 | 
| 
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
187  | 
val stream = system.fifo_output_stream(fifo)  | 
| 
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
188  | 
val writer = new BufferedWriter(new OutputStreamWriter(stream, Standard_System.charset))  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
189  | 
var finished = false  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
190  | 
      while (!finished) {
 | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
191  | 
        try {
 | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
192  | 
          //{{{
 | 
| 
27955
 
4c32c5e75eca
use java.util.concurrent.LinkedBlockingQueue, which blocks as required;
 
wenzelm 
parents: 
27949 
diff
changeset
 | 
193  | 
val s = output.take  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
194  | 
          if (s == "\u0000") {
 | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
195  | 
writer.close  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
196  | 
finished = true  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
197  | 
}  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
198  | 
          else {
 | 
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
199  | 
put_result(Markup.STDIN, s)  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
200  | 
writer.write(s)  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
201  | 
writer.flush  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
202  | 
}  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
203  | 
//}}}  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
204  | 
}  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
205  | 
        catch {
 | 
| 
38253
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
206  | 
case e: IOException => put_result(Markup.SYSTEM, "Command thread: " + e.getMessage)  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
207  | 
}  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
208  | 
}  | 
| 
38253
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
209  | 
put_result(Markup.SYSTEM, "Command thread terminated")  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
210  | 
}  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
211  | 
}  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
212  | 
|
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
213  | 
|
| 
38253
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
214  | 
/* raw stdout */  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
215  | 
|
| 
38253
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
216  | 
  private class Stdout_Thread(stream: InputStream) extends Thread("isabelle: stdout")
 | 
| 
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
217  | 
  {
 | 
| 
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
218  | 
override def run() =  | 
| 
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
219  | 
    {
 | 
| 
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
220  | 
val reader = new BufferedReader(new InputStreamReader(stream, Standard_System.charset))  | 
| 28045 | 221  | 
var result = new StringBuilder(100)  | 
| 
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  | 
var finished = false  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
224  | 
      while (!finished) {
 | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
225  | 
        try {
 | 
| 28045 | 226  | 
          //{{{
 | 
227  | 
var c = -1  | 
|
228  | 
var done = false  | 
|
229  | 
          while (!done && (result.length == 0 || reader.ready)) {
 | 
|
230  | 
c = reader.read  | 
|
231  | 
if (c >= 0) result.append(c.asInstanceOf[Char])  | 
|
232  | 
else done = true  | 
|
233  | 
}  | 
|
234  | 
          if (result.length > 0) {
 | 
|
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
235  | 
put_result(Markup.STDOUT, result.toString)  | 
| 28045 | 236  | 
result.length = 0  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
237  | 
}  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
238  | 
          else {
 | 
| 28045 | 239  | 
reader.close  | 
240  | 
finished = true  | 
|
241  | 
try_close()  | 
|
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
242  | 
}  | 
| 28045 | 243  | 
//}}}  | 
| 27963 | 244  | 
}  | 
245  | 
        catch {
 | 
|
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
246  | 
case e: IOException => put_result(Markup.SYSTEM, "Stdout thread: " + e.getMessage)  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
247  | 
}  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
248  | 
}  | 
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
249  | 
put_result(Markup.SYSTEM, "Stdout thread terminated")  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
250  | 
}  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
251  | 
}  | 
| 
 
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  | 
|
| 28045 | 254  | 
/* messages */  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
255  | 
|
| 
34213
 
9e86c1ca6e51
removed obsolete version check -- sanity delegated to Isabelle_System;
 
wenzelm 
parents: 
34201 
diff
changeset
 | 
256  | 
  private class Message_Thread(fifo: String) extends Thread("isabelle: messages")
 | 
| 34100 | 257  | 
  {
 | 
258  | 
private class Protocol_Error(msg: String) extends Exception(msg)  | 
|
259  | 
override def run()  | 
|
260  | 
    {
 | 
|
| 
38253
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
261  | 
val stream = system.fifo_input_stream(fifo)  | 
| 34100 | 262  | 
val default_buffer = new Array[Byte](65536)  | 
263  | 
var c = -1  | 
|
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
264  | 
|
| 34100 | 265  | 
def read_chunk(): List[XML.Tree] =  | 
266  | 
      {
 | 
|
267  | 
        //{{{
 | 
|
268  | 
// chunk size  | 
|
269  | 
var n = 0  | 
|
270  | 
c = stream.read  | 
|
271  | 
        while (48 <= c && c <= 57) {
 | 
|
272  | 
n = 10 * n + (c - 48)  | 
|
273  | 
c = stream.read  | 
|
274  | 
}  | 
|
275  | 
        if (c != 10) throw new Protocol_Error("bad message chunk header")
 | 
|
276  | 
||
277  | 
// chunk content  | 
|
278  | 
val buf =  | 
|
279  | 
if (n <= default_buffer.size) default_buffer  | 
|
280  | 
else new Array[Byte](n)  | 
|
281  | 
||
282  | 
var i = 0  | 
|
283  | 
var m = 0  | 
|
284  | 
        do {
 | 
|
285  | 
m = stream.read(buf, i, n - i)  | 
|
286  | 
i += m  | 
|
287  | 
} while (m > 0 && n > i)  | 
|
288  | 
||
289  | 
        if (i != n) throw new Protocol_Error("bad message chunk content")
 | 
|
290  | 
||
291  | 
YXML.parse_body_failsafe(YXML.decode_chars(system.symbols.decode, buf, 0, n))  | 
|
292  | 
//}}}  | 
|
293  | 
}  | 
|
294  | 
||
295  | 
      do {
 | 
|
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
296  | 
        try {
 | 
| 34100 | 297  | 
          //{{{
 | 
298  | 
c = stream.read  | 
|
299  | 
var non_sync = 0  | 
|
300  | 
          while (c >= 0 && c != 2) {
 | 
|
301  | 
non_sync += 1  | 
|
302  | 
c = stream.read  | 
|
| 28063 | 303  | 
}  | 
| 34100 | 304  | 
if (non_sync > 0)  | 
305  | 
            throw new Protocol_Error("lost synchronization -- skipping " + non_sync + " bytes")
 | 
|
306  | 
          if (c == 2) {
 | 
|
307  | 
val header = read_chunk()  | 
|
308  | 
val body = read_chunk()  | 
|
309  | 
            header match {
 | 
|
| 
38230
 
ed147003de4b
simplified type XML.Tree: embed Markup directly, avoid slightly odd triple;
 
wenzelm 
parents: 
37712 
diff
changeset
 | 
310  | 
case List(XML.Elem(Markup(name, props), Nil))  | 
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
311  | 
if name.size == 1 && Kind.markup.isDefinedAt(name(0)) =>  | 
| 
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
312  | 
put_result(Kind.markup(name(0)), props, body)  | 
| 34100 | 313  | 
              case _ => throw new Protocol_Error("bad header: " + header.toString)
 | 
| 28063 | 314  | 
}  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
315  | 
}  | 
| 34100 | 316  | 
//}}}  | 
| 27963 | 317  | 
}  | 
| 28063 | 318  | 
        catch {
 | 
| 34100 | 319  | 
case e: IOException =>  | 
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
320  | 
put_result(Markup.SYSTEM, "Cannot read message:\n" + e.getMessage)  | 
| 34100 | 321  | 
case e: Protocol_Error =>  | 
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
322  | 
put_result(Markup.SYSTEM, "Malformed message:\n" + e.getMessage)  | 
| 28063 | 323  | 
}  | 
| 34100 | 324  | 
} while (c != -1)  | 
325  | 
stream.close  | 
|
326  | 
try_close()  | 
|
327  | 
||
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
328  | 
put_result(Markup.SYSTEM, "Message thread terminated")  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
329  | 
}  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
330  | 
}  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
331  | 
|
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
332  | 
|
| 29192 | 333  | 
|
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
334  | 
/** main **/  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
335  | 
|
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
336  | 
  {
 | 
| 
38253
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
337  | 
/* private communication channels */  | 
| 28045 | 338  | 
|
| 
38253
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
339  | 
val in_fifo = system.mk_fifo()  | 
| 
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
340  | 
val out_fifo = system.mk_fifo()  | 
| 
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
341  | 
    def rm_fifos() = { system.rm_fifo(in_fifo); system.rm_fifo(out_fifo) }
 | 
| 28045 | 342  | 
|
| 
38253
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
343  | 
val command_thread = new Command_Thread(in_fifo)  | 
| 
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
344  | 
val message_thread = new Message_Thread(out_fifo)  | 
| 
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
345  | 
|
| 
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
346  | 
command_thread.start  | 
| 28063 | 347  | 
message_thread.start  | 
| 28045 | 348  | 
|
349  | 
||
| 
27973
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
350  | 
/* exec process */  | 
| 27963 | 351  | 
|
| 
27973
 
18d02c0b90b6
moved class Result into static object, removed dynamic tree method;
 
wenzelm 
parents: 
27963 
diff
changeset
 | 
352  | 
    try {
 | 
| 
38253
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
353  | 
val cmdline =  | 
| 
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
354  | 
        List(system.getenv_strict("ISABELLE_PROCESS"), "-W", in_fifo + ":" + out_fifo) ++ args
 | 
| 34100 | 355  | 
proc = system.execute(true, cmdline: _*)  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
356  | 
}  | 
| 
27993
 
6dd90ef9f927
simplified exceptions: use plain error function / RuntimeException;
 
wenzelm 
parents: 
27992 
diff
changeset
 | 
357  | 
    catch {
 | 
| 28063 | 358  | 
case e: IOException =>  | 
| 
38253
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
359  | 
rm_fifos()  | 
| 28063 | 360  | 
        error("Failed to execute Isabelle process: " + e.getMessage)
 | 
| 
27993
 
6dd90ef9f927
simplified exceptions: use plain error function / RuntimeException;
 
wenzelm 
parents: 
27992 
diff
changeset
 | 
361  | 
}  | 
| 
34213
 
9e86c1ca6e51
removed obsolete version check -- sanity delegated to Isabelle_System;
 
wenzelm 
parents: 
34201 
diff
changeset
 | 
362  | 
new Stdout_Thread(proc.getInputStream).start  | 
| 28045 | 363  | 
|
364  | 
||
365  | 
/* exit */  | 
|
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
366  | 
|
| 28063 | 367  | 
    new Thread("isabelle: exit") {
 | 
| 
38253
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
368  | 
override def run()  | 
| 
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
369  | 
      {
 | 
| 28045 | 370  | 
val rc = proc.waitFor()  | 
| 37132 | 371  | 
Thread.sleep(300) // FIXME property!?  | 
| 
37689
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
372  | 
put_result(Markup.SYSTEM, "Exit thread terminated")  | 
| 
 
628eabe2213a
simplified Isabelle_Process.Result: use markup directly;
 
wenzelm 
parents: 
37132 
diff
changeset
 | 
373  | 
put_result(Markup.EXIT, rc.toString)  | 
| 
38253
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
374  | 
rm_fifos()  | 
| 28045 | 375  | 
}  | 
| 28063 | 376  | 
}.start  | 
| 
27949
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
377  | 
}  | 
| 
 
6eb0327c0b9b
Isabelle process management -- always reactive due to multi-threaded I/O.
 
wenzelm 
parents:  
diff
changeset
 | 
378  | 
}  |