src/Pure/System/isabelle_process.scala
author wenzelm
Sat, 18 Sep 2010 17:14:47 +0200
changeset 39519 b376f53bcc18
parent 39513 fce2202892c4
child 39523 d8971680b0fc
permissions -rw-r--r--
slightly more robust Isabelle_Process startup -- NB: openening fifo streams synchronizes with other end, which may fail to reach that point;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
c9ea82444189 YXML.parse_failsafe;
wenzelm
parents: 27955
diff changeset
     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
1a6f273108ae join stdout/stderr, eliminated Kind.STDERR;
wenzelm
parents: 27999
diff changeset
    11
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
    12
  InputStream, OutputStream, BufferedOutputStream, 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,
39513
fce2202892c4 discontinued Output.debug, which belongs to early PGIP experiments (b6788dbd2ef9) and causes just too many problems (like spamming the message channel if it is used by more than one module);
wenzelm
parents: 39439
diff changeset
    31
      ('G' : Int) -> Markup.ERROR)
37689
628eabe2213a simplified Isabelle_Process.Result: use markup directly;
wenzelm
parents: 37132
diff changeset
    32
    def is_raw(kind: String) =
628eabe2213a simplified Isabelle_Process.Result: use markup directly;
wenzelm
parents: 37132
diff changeset
    33
      kind == Markup.STDOUT
628eabe2213a simplified Isabelle_Process.Result: use markup directly;
wenzelm
parents: 37132
diff changeset
    34
    def is_control(kind: String) =
628eabe2213a simplified Isabelle_Process.Result: use markup directly;
wenzelm
parents: 37132
diff changeset
    35
      kind == Markup.SYSTEM ||
628eabe2213a simplified Isabelle_Process.Result: use markup directly;
wenzelm
parents: 37132
diff changeset
    36
      kind == Markup.SIGNAL ||
628eabe2213a simplified Isabelle_Process.Result: use markup directly;
wenzelm
parents: 37132
diff changeset
    37
      kind == Markup.EXIT
628eabe2213a simplified Isabelle_Process.Result: use markup directly;
wenzelm
parents: 37132
diff changeset
    38
    def is_system(kind: String) =
628eabe2213a simplified Isabelle_Process.Result: use markup directly;
wenzelm
parents: 37132
diff changeset
    39
      kind == Markup.SYSTEM ||
38259
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
    40
      kind == Markup.INPUT ||
37689
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
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
    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
968844caaff9 simplified some Markup;
wenzelm
parents: 38230
diff changeset
    58
    def is_ready = is_status && body == List(XML.Elem(Markup.Ready, Nil))
34100
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
    59
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
    60
    override def toString: String =
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
    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
    }
27973
18d02c0b90b6 moved class Result into static object, removed dynamic tree method;
wenzelm
parents: 27963
diff changeset
    71
  }
18d02c0b90b6 moved class Result into static object, removed dynamic tree method;
wenzelm
parents: 27963
diff changeset
    72
}
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
34100
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
    75
class Isabelle_Process(system: Isabelle_System, receiver: Actor, args: String*)
29192
082ee2a01a6d explicit EventBus for results;
wenzelm
parents: 29178
diff changeset
    76
{
31797
203d5e61e3bc renamed IsabelleProcess to Isabelle_Process;
wenzelm
parents: 31498
diff changeset
    77
  import Isabelle_Process._
29194
wenzelm
parents: 29192
diff changeset
    78
27973
18d02c0b90b6 moved class Result into static object, removed dynamic tree method;
wenzelm
parents: 27963
diff changeset
    79
29192
082ee2a01a6d explicit EventBus for results;
wenzelm
parents: 29178
diff changeset
    80
  /* demo constructor */
27973
18d02c0b90b6 moved class Result into static object, removed dynamic tree method;
wenzelm
parents: 27963
diff changeset
    81
29192
082ee2a01a6d explicit EventBus for results;
wenzelm
parents: 29178
diff changeset
    82
  def this(args: String*) =
32474
0818e6b1c8a6 Isabelle_Process: receiver as Actor, not EventBus;
wenzelm
parents: 32448
diff changeset
    83
    this(new Isabelle_System,
34213
9e86c1ca6e51 removed obsolete version check -- sanity delegated to Isabelle_System;
wenzelm
parents: 34201
diff changeset
    84
      actor { loop { react { case res => Console.println(res) } } }, args: _*)
29174
d4058295affb proper class IsabelleSystem -- no longer static;
wenzelm
parents: 29140
diff changeset
    85
d4058295affb proper class IsabelleSystem -- no longer static;
wenzelm
parents: 29140
diff changeset
    86
27973
18d02c0b90b6 moved class Result into static object, removed dynamic tree method;
wenzelm
parents: 27963
diff changeset
    87
  /* process information */
18d02c0b90b6 moved class Result into static object, removed dynamic tree method;
wenzelm
parents: 27963
diff changeset
    88
38270
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
    89
  @volatile private var proc: Option[Process] = None
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
    90
  @volatile private var pid: Option[String] = None
27973
18d02c0b90b6 moved class Result into static object, removed dynamic tree method;
wenzelm
parents: 27963
diff changeset
    91
18d02c0b90b6 moved class Result into static object, removed dynamic tree method;
wenzelm
parents: 27963
diff changeset
    92
18d02c0b90b6 moved class Result into static object, removed dynamic tree method;
wenzelm
parents: 27963
diff changeset
    93
  /* results */
27949
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
    94
38446
9d59dab38fef XML.Cache: pipe-lined (thread-safe) version using actor;
wenzelm
parents: 38445
diff changeset
    95
  private val xml_cache = new XML.Cache(131071)
9d59dab38fef XML.Cache: pipe-lined (thread-safe) version using actor;
wenzelm
parents: 38445
diff changeset
    96
38573
d163f0f28e8c tuned signatures;
wenzelm
parents: 38446
diff changeset
    97
  private def put_result(kind: String, props: List[(String, String)], body: XML.Body)
29572
e3a99d957392 replaced java.util.Properties by plain association list;
wenzelm
parents: 29522
diff changeset
    98
  {
38446
9d59dab38fef XML.Cache: pipe-lined (thread-safe) version using actor;
wenzelm
parents: 38445
diff changeset
    99
    if (pid.isEmpty && kind == Markup.INIT)
9d59dab38fef XML.Cache: pipe-lined (thread-safe) version using actor;
wenzelm
parents: 38445
diff changeset
   100
      pid = props.find(_._1 == Markup.PID).map(_._1)
9d59dab38fef XML.Cache: pipe-lined (thread-safe) version using actor;
wenzelm
parents: 38445
diff changeset
   101
39439
1c294d150ded eliminated markup "location" in favour of more explicit "no_report", which is actually deleted from messages;
wenzelm
parents: 38636
diff changeset
   102
    val msg = XML.Elem(Markup(kind, props), Isar_Document.clean_message(body))
1c294d150ded eliminated markup "location" in favour of more explicit "no_report", which is actually deleted from messages;
wenzelm
parents: 38636
diff changeset
   103
    xml_cache.cache_tree(msg)((message: XML.Tree) =>
38446
9d59dab38fef XML.Cache: pipe-lined (thread-safe) version using actor;
wenzelm
parents: 38445
diff changeset
   104
      receiver ! new Result(message.asInstanceOf[XML.Elem]))
34100
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   105
  }
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   106
37689
628eabe2213a simplified Isabelle_Process.Result: use markup directly;
wenzelm
parents: 37132
diff changeset
   107
  private def put_result(kind: String, text: String)
34100
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   108
  {
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   109
    put_result(kind, Nil, List(XML.Text(system.symbols.decode(text))))
27992
131f7ea9e6e6 added try_result;
wenzelm
parents: 27990
diff changeset
   110
  }
131f7ea9e6e6 added try_result;
wenzelm
parents: 27990
diff changeset
   111
27973
18d02c0b90b6 moved class Result into static object, removed dynamic tree method;
wenzelm
parents: 27963
diff changeset
   112
18d02c0b90b6 moved class Result into static object, removed dynamic tree method;
wenzelm
parents: 27963
diff changeset
   113
  /* signals */
18d02c0b90b6 moved class Result into static object, removed dynamic tree method;
wenzelm
parents: 27963
diff changeset
   114
38270
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   115
  def interrupt()
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   116
  {
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   117
    if (proc.isEmpty) put_result(Markup.SYSTEM, "Cannot interrupt Isabelle: no process")
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   118
    else
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   119
      pid match {
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   120
        case None => put_result(Markup.SYSTEM, "Cannot interrupt Isabelle: unknowd pid")
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   121
        case Some(i) =>
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   122
          try {
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   123
            if (system.execute(true, "kill", "-INT", i).waitFor == 0)
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   124
              put_result(Markup.SIGNAL, "INT")
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   125
            else
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   126
              put_result(Markup.SYSTEM, "Cannot interrupt Isabelle: kill command failed")
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   127
          }
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   128
          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
   129
      }
18d02c0b90b6 moved class Result into static object, removed dynamic tree method;
wenzelm
parents: 27963
diff changeset
   130
  }
18d02c0b90b6 moved class Result into static object, removed dynamic tree method;
wenzelm
parents: 27963
diff changeset
   131
38270
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   132
  def kill()
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   133
  {
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   134
    proc match {
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   135
      case None => put_result(Markup.SYSTEM, "Cannot kill Isabelle: no process")
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   136
      case Some(p) =>
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   137
        close()
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   138
        Thread.sleep(500)  // FIXME !?
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   139
        put_result(Markup.SIGNAL, "KILL")
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   140
        p.destroy
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   141
        proc = None
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   142
        pid = None
27973
18d02c0b90b6 moved class Result into static object, removed dynamic tree method;
wenzelm
parents: 27963
diff changeset
   143
    }
18d02c0b90b6 moved class Result into static object, removed dynamic tree method;
wenzelm
parents: 27963
diff changeset
   144
  }
18d02c0b90b6 moved class Result into static object, removed dynamic tree method;
wenzelm
parents: 27963
diff changeset
   145
27949
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   146
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   147
38259
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   148
  /** stream actors **/
27949
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   149
39519
b376f53bcc18 slightly more robust Isabelle_Process startup -- NB: openening fifo streams synchronizes with other end, which may fail to reach that point;
wenzelm
parents: 39513
diff changeset
   150
  private val in_fifo = system.mk_fifo()
b376f53bcc18 slightly more robust Isabelle_Process startup -- NB: openening fifo streams synchronizes with other end, which may fail to reach that point;
wenzelm
parents: 39513
diff changeset
   151
  private val out_fifo = system.mk_fifo()
b376f53bcc18 slightly more robust Isabelle_Process startup -- NB: openening fifo streams synchronizes with other end, which may fail to reach that point;
wenzelm
parents: 39513
diff changeset
   152
  private def rm_fifos() = { system.rm_fifo(in_fifo); system.rm_fifo(out_fifo) }
b376f53bcc18 slightly more robust Isabelle_Process startup -- NB: openening fifo streams synchronizes with other end, which may fail to reach that point;
wenzelm
parents: 39513
diff changeset
   153
b376f53bcc18 slightly more robust Isabelle_Process startup -- NB: openening fifo streams synchronizes with other end, which may fail to reach that point;
wenzelm
parents: 39513
diff changeset
   154
  private case class Input_Text(text: String)
b376f53bcc18 slightly more robust Isabelle_Process startup -- NB: openening fifo streams synchronizes with other end, which may fail to reach that point;
wenzelm
parents: 39513
diff changeset
   155
  private case class Input_Chunks(chunks: List[Array[Byte]])
b376f53bcc18 slightly more robust Isabelle_Process startup -- NB: openening fifo streams synchronizes with other end, which may fail to reach that point;
wenzelm
parents: 39513
diff changeset
   156
  private case object Close
27949
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   157
38270
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   158
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   159
  /* raw stdin */
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   160
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   161
  private def stdin_actor(name: String, stream: OutputStream): Actor =
38636
b7647ca7de5a module for simplified thread operations (Scala version);
wenzelm
parents: 38573
diff changeset
   162
    Simple_Thread.actor(name) {
38253
3d4e521014f7 Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
wenzelm
parents: 38236
diff changeset
   163
      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
   164
      var finished = false
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   165
      while (!finished) {
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   166
        try {
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   167
          //{{{
38259
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   168
          receive {
38270
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   169
            case Input_Text(text) =>
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   170
              // FIXME echo input?!
38259
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   171
              writer.write(text)
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   172
              writer.flush
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   173
            case Close =>
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   174
              writer.close
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   175
              finished = true
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   176
            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
   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
        catch {
38259
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   181
          case e: IOException => put_result(Markup.SYSTEM, name + ": " + e.getMessage)
27949
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   182
        }
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   183
      }
38259
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   184
      put_result(Markup.SYSTEM, name + " terminated")
27949
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   185
    }
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   186
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   187
38270
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   188
  /* raw stdout */
27949
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   189
38270
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   190
  private def stdout_actor(name: String, stream: InputStream): Actor =
38636
b7647ca7de5a module for simplified thread operations (Scala version);
wenzelm
parents: 38573
diff changeset
   191
    Simple_Thread.actor(name) {
38253
3d4e521014f7 Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
wenzelm
parents: 38236
diff changeset
   192
      val reader = new BufferedReader(new InputStreamReader(stream, Standard_System.charset))
28045
1a6f273108ae join stdout/stderr, eliminated Kind.STDERR;
wenzelm
parents: 27999
diff changeset
   193
      var result = new StringBuilder(100)
27949
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   194
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   195
      var finished = false
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   196
      while (!finished) {
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   197
        try {
28045
1a6f273108ae join stdout/stderr, eliminated Kind.STDERR;
wenzelm
parents: 27999
diff changeset
   198
          //{{{
1a6f273108ae join stdout/stderr, eliminated Kind.STDERR;
wenzelm
parents: 27999
diff changeset
   199
          var c = -1
1a6f273108ae join stdout/stderr, eliminated Kind.STDERR;
wenzelm
parents: 27999
diff changeset
   200
          var done = false
1a6f273108ae join stdout/stderr, eliminated Kind.STDERR;
wenzelm
parents: 27999
diff changeset
   201
          while (!done && (result.length == 0 || reader.ready)) {
1a6f273108ae join stdout/stderr, eliminated Kind.STDERR;
wenzelm
parents: 27999
diff changeset
   202
            c = reader.read
1a6f273108ae join stdout/stderr, eliminated Kind.STDERR;
wenzelm
parents: 27999
diff changeset
   203
            if (c >= 0) result.append(c.asInstanceOf[Char])
1a6f273108ae join stdout/stderr, eliminated Kind.STDERR;
wenzelm
parents: 27999
diff changeset
   204
            else done = true
1a6f273108ae join stdout/stderr, eliminated Kind.STDERR;
wenzelm
parents: 27999
diff changeset
   205
          }
1a6f273108ae join stdout/stderr, eliminated Kind.STDERR;
wenzelm
parents: 27999
diff changeset
   206
          if (result.length > 0) {
38270
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   207
            put_result(Markup.STDOUT, result.toString)
28045
1a6f273108ae join stdout/stderr, eliminated Kind.STDERR;
wenzelm
parents: 27999
diff changeset
   208
            result.length = 0
27949
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   209
          }
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   210
          else {
28045
1a6f273108ae join stdout/stderr, eliminated Kind.STDERR;
wenzelm
parents: 27999
diff changeset
   211
            reader.close
1a6f273108ae join stdout/stderr, eliminated Kind.STDERR;
wenzelm
parents: 27999
diff changeset
   212
            finished = true
38270
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   213
            close()
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   214
          }
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   215
          //}}}
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   216
        }
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   217
        catch {
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   218
          case e: IOException => put_result(Markup.SYSTEM, name + ": " + e.getMessage)
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   219
        }
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   220
      }
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   221
      put_result(Markup.SYSTEM, name + " terminated")
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   222
    }
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   223
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   224
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   225
  /* command input */
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   226
39519
b376f53bcc18 slightly more robust Isabelle_Process startup -- NB: openening fifo streams synchronizes with other end, which may fail to reach that point;
wenzelm
parents: 39513
diff changeset
   227
  private def input_actor(name: String): Actor =
38636
b7647ca7de5a module for simplified thread operations (Scala version);
wenzelm
parents: 38573
diff changeset
   228
    Simple_Thread.actor(name) {
39519
b376f53bcc18 slightly more robust Isabelle_Process startup -- NB: openening fifo streams synchronizes with other end, which may fail to reach that point;
wenzelm
parents: 39513
diff changeset
   229
      val stream = new BufferedOutputStream(system.fifo_output_stream(in_fifo))  // FIXME potentially blocking forever
38270
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   230
      var finished = false
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   231
      while (!finished) {
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   232
        try {
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   233
          //{{{
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   234
          receive {
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   235
            case Input_Chunks(chunks) =>
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   236
              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
   237
                  chunks.map(_.length).mkString("", ",", "\n")));
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   238
              chunks.foreach(stream.write(_));
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   239
              stream.flush
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   240
            case Close =>
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   241
              stream.close
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   242
              finished = true
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
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
          }
28045
1a6f273108ae join stdout/stderr, eliminated Kind.STDERR;
wenzelm
parents: 27999
diff changeset
   245
          //}}}
27963
c9ea82444189 YXML.parse_failsafe;
wenzelm
parents: 27955
diff changeset
   246
        }
c9ea82444189 YXML.parse_failsafe;
wenzelm
parents: 27955
diff changeset
   247
        catch {
38259
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   248
          case e: IOException => put_result(Markup.SYSTEM, name + ": " + e.getMessage)
27949
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   249
        }
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   250
      }
38259
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   251
      put_result(Markup.SYSTEM, name + " terminated")
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
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   254
38259
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   255
  /* message output */
27949
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   256
38259
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   257
  private class Protocol_Error(msg: String) extends Exception(msg)
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   258
39519
b376f53bcc18 slightly more robust Isabelle_Process startup -- NB: openening fifo streams synchronizes with other end, which may fail to reach that point;
wenzelm
parents: 39513
diff changeset
   259
  private def message_actor(name: String): Actor =
38636
b7647ca7de5a module for simplified thread operations (Scala version);
wenzelm
parents: 38573
diff changeset
   260
    Simple_Thread.actor(name) {
39519
b376f53bcc18 slightly more robust Isabelle_Process startup -- NB: openening fifo streams synchronizes with other end, which may fail to reach that point;
wenzelm
parents: 39513
diff changeset
   261
      val stream = system.fifo_input_stream(out_fifo)  // FIXME potentially blocking forever
34100
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   262
      val default_buffer = new Array[Byte](65536)
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   263
      var c = -1
27949
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   264
38573
d163f0f28e8c tuned signatures;
wenzelm
parents: 38446
diff changeset
   265
      def read_chunk(): XML.Body =
34100
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   266
      {
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   267
        //{{{
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   268
        // chunk size
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   269
        var n = 0
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   270
        c = stream.read
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   271
        while (48 <= c && c <= 57) {
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   272
          n = 10 * n + (c - 48)
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   273
          c = stream.read
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   274
        }
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   275
        if (c != 10) throw new Protocol_Error("bad message chunk header")
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   276
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   277
        // chunk content
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   278
        val buf =
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   279
          if (n <= default_buffer.size) default_buffer
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   280
          else new Array[Byte](n)
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   281
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   282
        var i = 0
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   283
        var m = 0
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   284
        do {
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   285
          m = stream.read(buf, i, n - i)
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   286
          i += m
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   287
        } while (m > 0 && n > i)
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   288
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   289
        if (i != n) throw new Protocol_Error("bad message chunk content")
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   290
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   291
        YXML.parse_body_failsafe(YXML.decode_chars(system.symbols.decode, buf, 0, n))
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   292
        //}}}
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   293
      }
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   294
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   295
      do {
27949
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   296
        try {
38445
ba9ea6b9b75c simplified internal message format: dropped special Symbol.STX header;
wenzelm
parents: 38372
diff changeset
   297
          val header = read_chunk()
ba9ea6b9b75c simplified internal message format: dropped special Symbol.STX header;
wenzelm
parents: 38372
diff changeset
   298
          val body = read_chunk()
ba9ea6b9b75c simplified internal message format: dropped special Symbol.STX header;
wenzelm
parents: 38372
diff changeset
   299
          header match {
ba9ea6b9b75c simplified internal message format: dropped special Symbol.STX header;
wenzelm
parents: 38372
diff changeset
   300
            case List(XML.Elem(Markup(name, props), Nil))
ba9ea6b9b75c simplified internal message format: dropped special Symbol.STX header;
wenzelm
parents: 38372
diff changeset
   301
                if name.size == 1 && Kind.markup.isDefinedAt(name(0)) =>
ba9ea6b9b75c simplified internal message format: dropped special Symbol.STX header;
wenzelm
parents: 38372
diff changeset
   302
              put_result(Kind.markup(name(0)), props, body)
ba9ea6b9b75c simplified internal message format: dropped special Symbol.STX header;
wenzelm
parents: 38372
diff changeset
   303
            case _ => throw new Protocol_Error("bad header: " + header.toString)
28063
3533485fc7b8 IsabelleSystem.mk_fifo, IsabelleSystem.rm_fifo;
wenzelm
parents: 28056
diff changeset
   304
          }
27963
c9ea82444189 YXML.parse_failsafe;
wenzelm
parents: 27955
diff changeset
   305
        }
28063
3533485fc7b8 IsabelleSystem.mk_fifo, IsabelleSystem.rm_fifo;
wenzelm
parents: 28056
diff changeset
   306
        catch {
34100
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   307
          case e: IOException =>
37689
628eabe2213a simplified Isabelle_Process.Result: use markup directly;
wenzelm
parents: 37132
diff changeset
   308
            put_result(Markup.SYSTEM, "Cannot read message:\n" + e.getMessage)
34100
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   309
          case e: Protocol_Error =>
37689
628eabe2213a simplified Isabelle_Process.Result: use markup directly;
wenzelm
parents: 37132
diff changeset
   310
            put_result(Markup.SYSTEM, "Malformed message:\n" + e.getMessage)
28063
3533485fc7b8 IsabelleSystem.mk_fifo, IsabelleSystem.rm_fifo;
wenzelm
parents: 28056
diff changeset
   311
        }
34100
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   312
      } while (c != -1)
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   313
      stream.close
38270
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   314
      close()
34100
ea24958c2af5 fifo: raw byte stream;
wenzelm
parents: 32474
diff changeset
   315
38259
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   316
      put_result(Markup.SYSTEM, name + " terminated")
27949
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   317
    }
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   318
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   319
29192
082ee2a01a6d explicit EventBus for results;
wenzelm
parents: 29178
diff changeset
   320
38259
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   321
  /** init **/
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   322
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   323
  /* exec process */
27949
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   324
38259
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   325
  try {
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   326
    val cmdline =
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   327
      List(system.getenv_strict("ISABELLE_PROCESS"), "-W", in_fifo + ":" + out_fifo) ++ args
38270
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   328
    proc = Some(system.execute(true, cmdline: _*))
38259
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   329
  }
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   330
  catch {
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   331
    case e: IOException =>
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   332
      rm_fifos()
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   333
      error("Failed to execute Isabelle process: " + e.getMessage)
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   334
  }
28045
1a6f273108ae join stdout/stderr, eliminated Kind.STDERR;
wenzelm
parents: 27999
diff changeset
   335
38259
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   336
38270
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   337
  /* I/O actors */
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   338
39519
b376f53bcc18 slightly more robust Isabelle_Process startup -- NB: openening fifo streams synchronizes with other end, which may fail to reach that point;
wenzelm
parents: 39513
diff changeset
   339
  private val command_input = input_actor("command_input")
b376f53bcc18 slightly more robust Isabelle_Process startup -- NB: openening fifo streams synchronizes with other end, which may fail to reach that point;
wenzelm
parents: 39513
diff changeset
   340
  message_actor("message_output")
b376f53bcc18 slightly more robust Isabelle_Process startup -- NB: openening fifo streams synchronizes with other end, which may fail to reach that point;
wenzelm
parents: 39513
diff changeset
   341
38270
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   342
  private val standard_input = stdin_actor("standard_input", proc.get.getOutputStream)
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   343
  stdout_actor("standard_output", proc.get.getInputStream)
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   344
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   345
38259
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   346
  /* exit process */
38253
3d4e521014f7 Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
wenzelm
parents: 38236
diff changeset
   347
38636
b7647ca7de5a module for simplified thread operations (Scala version);
wenzelm
parents: 38573
diff changeset
   348
  Simple_Thread.actor("process_exit") {
38270
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   349
    proc match {
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   350
      case None =>
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   351
      case Some(p) =>
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   352
        val rc = p.waitFor()
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   353
        Thread.sleep(300)  // FIXME property!?
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   354
        put_result(Markup.SYSTEM, "process_exit terminated")
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   355
        put_result(Markup.EXIT, rc.toString)
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   356
    }
38259
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   357
    rm_fifos()
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   358
  }
28045
1a6f273108ae join stdout/stderr, eliminated Kind.STDERR;
wenzelm
parents: 27999
diff changeset
   359
1a6f273108ae join stdout/stderr, eliminated Kind.STDERR;
wenzelm
parents: 27999
diff changeset
   360
1a6f273108ae join stdout/stderr, eliminated Kind.STDERR;
wenzelm
parents: 27999
diff changeset
   361
38259
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   362
  /** main methods **/
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   363
38270
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   364
  def input_raw(text: String): Unit = standard_input ! Input_Text(text)
38259
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   365
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
   366
  def input_bytes(name: String, args: Array[Byte]*): Unit =
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
   367
    command_input ! Input_Chunks(Standard_System.string_bytes(name) :: args.toList)
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
   368
38270
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   369
  def input(name: String, args: String*): Unit =
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
   370
    input_bytes(name, args.map(Standard_System.string_bytes): _*)
38259
2b61c5e27399 distinguish proper Isabelle_Process INPUT vs. raw STDIN, tuned corresponding method names;
wenzelm
parents: 38253
diff changeset
   371
38270
71bb3c273dd1 native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
wenzelm
parents: 38262
diff changeset
   372
  def close(): Unit = command_input ! Close
27949
6eb0327c0b9b Isabelle process management -- always reactive due to multi-threaded I/O.
wenzelm
parents:
diff changeset
   373
}