merged
authorhaftmann
Mon, 29 Dec 2008 17:57:18 +0100
changeset 29199 69b806be5a0a
parent 29198 418ed6411847 (current diff)
parent 29196 de62fdd4b432 (diff)
child 29201 03908107bc5b
merged
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/General/event_bus.scala	Mon Dec 29 17:57:18 2008 +0100
@@ -0,0 +1,23 @@
+/*  Title:      Pure/General/event_bus.scala
+    Author:     Makarius
+
+Generic event bus.
+*/
+
+package isabelle
+
+import scala.collection.mutable.ListBuffer
+
+
+class EventBus[Event] {
+  type Handler = Event => Unit
+  private val handlers = new ListBuffer[Handler]
+
+  def += (h: Handler) = synchronized { handlers += h }
+  def + (h: Handler) = { this += h; this }
+
+  def -= (h: Handler) = synchronized { handlers -= h }
+  def - (h: Handler) = { this -= h; this }
+
+  def event(e: Event) = synchronized { handlers.toList } foreach (h => h(e))
+}
--- a/src/Pure/General/markup.scala	Mon Dec 29 17:56:37 2008 +0100
+++ b/src/Pure/General/markup.scala	Mon Dec 29 17:57:18 2008 +0100
@@ -120,6 +120,8 @@
   val PID = "pid"
   val SESSION = "session"
 
+  val MESSAGE = "message"
+
 
   /* content */
 
--- a/src/Pure/IsaMakefile	Mon Dec 29 17:56:37 2008 +0100
+++ b/src/Pure/IsaMakefile	Mon Dec 29 17:57:18 2008 +0100
@@ -121,10 +121,11 @@
 
 ## Scala material
 
-SCALA_FILES = General/markup.scala General/position.scala		\
-  General/symbol.scala General/xml.scala General/yxml.scala		\
-  Isar/isar.scala Thy/thy_header.scala Tools/isabelle_process.scala	\
-  Tools/isabelle_syntax.scala Tools/isabelle_system.scala
+SCALA_FILES = General/event_bus.scala General/markup.scala	\
+  General/position.scala General/symbol.scala General/xml.scala	\
+  General/yxml.scala Isar/isar.scala Thy/thy_header.scala	\
+  Tools/isabelle_process.scala Tools/isabelle_syntax.scala	\
+  Tools/isabelle_system.scala
 
 
 SCALA_TARGET = $(ISABELLE_HOME)/lib/classes/Pure.jar
--- a/src/Pure/Isar/isar.scala	Mon Dec 29 17:56:37 2008 +0100
+++ b/src/Pure/Isar/isar.scala	Mon Dec 29 17:57:18 2008 +0100
@@ -9,8 +9,9 @@
 import java.util.Properties
 
 
-class Isar(isabelle_system: IsabelleSystem, args: String*) extends
-    IsabelleProcess(isabelle_system, args: _*) {
+class Isar(isabelle_system: IsabelleSystem, results: EventBus[IsabelleProcess.Result], args: String*)
+  extends IsabelleProcess(isabelle_system, results, args: _*)
+{
 
   /* basic editor commands */
 
--- a/src/Pure/Tools/isabelle_process.scala	Mon Dec 29 17:56:37 2008 +0100
+++ b/src/Pure/Tools/isabelle_process.scala	Mon Dec 29 17:57:18 2008 +0100
@@ -71,22 +71,24 @@
       if (props == null) kind.toString + " [[" + res + "]]"
       else kind.toString + " " + props.toString + " [[" + res + "]]"
     }
-    def is_raw() = Kind.is_raw(kind)
-    def is_control() = Kind.is_control(kind)
-    def is_system() = Kind.is_system(kind)
+    def is_raw = Kind.is_raw(kind)
+    def is_control = Kind.is_control(kind)
+    def is_system = Kind.is_system(kind)
   }
 
 }
 
 
-class IsabelleProcess(isabelle_system: IsabelleSystem, args: String*) {
-
+class IsabelleProcess(isabelle_system: IsabelleSystem,
+  results: EventBus[IsabelleProcess.Result], args: String*)
+{
   import IsabelleProcess._
 
 
-  /* alternative constructors */
+  /* demo constructor */
 
-  def this(args: String*) = this(new IsabelleSystem, args: _*)
+  def this(args: String*) =
+    this(new IsabelleSystem, new EventBus[IsabelleProcess.Result] + Console.println, args: _*)
 
 
   /* process information */
@@ -100,21 +102,31 @@
 
   /* results */
 
-  private val results = new LinkedBlockingQueue[Result]
+  private val result_queue = new LinkedBlockingQueue[Result]
 
   private def put_result(kind: Kind.Value, props: Properties, result: String) {
     if (kind == Kind.INIT && props != null) {
       pid = props.getProperty(Markup.PID)
       the_session = props.getProperty(Markup.SESSION)
     }
-    results.put(new Result(kind, props, result))
+    result_queue.put(new Result(kind, props, result))
   }
 
-  def get_result() = results.take
+  private class ResultThread extends Thread("isabelle: results") {
+    override def run() = {
+      var finished = false
+      while (!finished) {
+        val result =
+          try { result_queue.take }
+          catch { case _: NullPointerException => null }
 
-  def try_result() = {
-    val res = results.poll
-    if (res != null) Some(res) else None
+        if (result != null) {
+          results.event(result)  // FIXME try/catch (!??)
+          if (result.kind == Kind.EXIT) finished = true
+        }
+        else finished = true
+      }
+    }
   }
 
 
@@ -334,6 +346,7 @@
   }
 
 
+
   /** main **/
 
   {
@@ -346,7 +359,7 @@
     }
 
 
-    /* message fifo */
+    /* messages */
 
     val message_fifo = isabelle_system.mk_fifo()
     def rm_fifo() = isabelle_system.rm_fifo(message_fifo)
@@ -354,6 +367,8 @@
     val message_thread = new MessageThread(message_fifo)
     message_thread.start
 
+    new ResultThread().start
+
 
     /* exec process */
 
@@ -388,5 +403,4 @@
     }.start
 
   }
-
 }