src/Pure/Tools/server.scala
changeset 67839 0c2ed45ece20
parent 67838 3a6ab890832f
child 67840 a9d450fc5a49
equal deleted inserted replaced
67838:3a6ab890832f 67839:0c2ed45ece20
    58 
    58 
    59   /* input command */
    59   /* input command */
    60 
    60 
    61   object Command
    61   object Command
    62   {
    62   {
    63     type T = PartialFunction[(Server, Any), Any]
    63     type T = PartialFunction[(Context, Any), Any]
    64 
    64 
    65     private val table: Map[String, T] =
    65     private val table: Map[String, T] =
    66       Map(
    66       Map(
    67         "echo" -> { case (_, t) => t },
    67         "echo" -> { case (_, t) => t },
    68         "help" -> { case (_, ()) => table.keySet.toList.sorted },
    68         "help" -> { case (_, ()) => table.keySet.toList.sorted },
    69         "shutdown" -> { case (server, ()) => server.close(); () })
    69         "shutdown" -> { case (context, ()) => context.shutdown(); () })
    70 
    70 
    71     def unapply(name: String): Option[T] = table.get(name)
    71     def unapply(name: String): Option[T] = table.get(name)
    72   }
    72   }
    73 
    73 
    74 
    74 
   110 
   110 
   111     def set_timeout(t: Time) { socket.setSoTimeout(t.ms.toInt) }
   111     def set_timeout(t: Time) { socket.setSoTimeout(t.ms.toInt) }
   112 
   112 
   113     private val in = new BufferedInputStream(socket.getInputStream)
   113     private val in = new BufferedInputStream(socket.getInputStream)
   114     private val out = new BufferedOutputStream(socket.getOutputStream)
   114     private val out = new BufferedOutputStream(socket.getOutputStream)
       
   115     private val out_lock: AnyRef = new Object
   115 
   116 
   116     def tty_loop(interrupt: Option[() => Unit] = None): TTY_Loop =
   117     def tty_loop(interrupt: Option[() => Unit] = None): TTY_Loop =
   117       new TTY_Loop(new OutputStreamWriter(out), new InputStreamReader(in), interrupt = interrupt)
   118       new TTY_Loop(
       
   119         new OutputStreamWriter(out),
       
   120         new InputStreamReader(in),
       
   121         writer_lock = out_lock,
       
   122         interrupt = interrupt)
   118 
   123 
   119     def read_message(): Option[String] =
   124     def read_message(): Option[String] =
   120       try {
   125       try {
   121         Bytes.read_line(in).map(_.text) match {
   126         Bytes.read_line(in).map(_.text) match {
   122           case Some(Value.Int(n)) =>
   127           case Some(Value.Int(n)) =>
   124           case res => res
   129           case res => res
   125         }
   130         }
   126       }
   131       }
   127       catch { case _: SocketException => None }
   132       catch { case _: SocketException => None }
   128 
   133 
   129     def write_message(msg: String)
   134     def write_message(msg: String): Unit = out_lock.synchronized
   130     {
   135     {
   131       val b = UTF8.bytes(msg)
   136       val b = UTF8.bytes(msg)
   132       if (b.length > 100 || b.contains(10)) {
   137       if (b.length > 100 || b.contains(10)) {
   133         out.write(UTF8.bytes((b.length + 1).toString))
   138         out.write(UTF8.bytes((b.length + 1).toString))
   134         out.write(10)
   139         out.write(10)
   148     def reply_error(arg: Any) { reply(Server.Reply.ERROR, arg) }
   153     def reply_error(arg: Any) { reply(Server.Reply.ERROR, arg) }
   149     def reply_error_message(message: String, more: (String, JSON.T)*): Unit =
   154     def reply_error_message(message: String, more: (String, JSON.T)*): Unit =
   150       reply_error(Map("message" -> message) ++ more)
   155       reply_error(Map("message" -> message) ++ more)
   151 
   156 
   152     def notify(arg: Any) { reply(Server.Reply.NOTE, arg) }
   157     def notify(arg: Any) { reply(Server.Reply.NOTE, arg) }
   153     def notify_message(message: String, more: (String, JSON.T)*): Unit =
   158     def notify_message(kind: String, msg: String, more: (String, JSON.T)*): Unit =
   154       notify(Map("message" -> message) ++ more)
   159       notify(Map(Markup.KIND -> kind, "message" -> msg) ++ more)
       
   160   }
       
   161 
       
   162 
       
   163   /* context with output channels */
       
   164 
       
   165   class Context private[Server](server: Server, connection: Connection)
       
   166   {
       
   167     context =>
       
   168 
       
   169     def shutdown() { server.close() }
       
   170 
       
   171     def message(kind: String, msg: String, more: (String, JSON.T)*): Unit =
       
   172       connection.notify_message(kind, msg, more:_*)
       
   173     def writeln(msg: String, more: (String, JSON.T)*): Unit = message(Markup.WRITELN, msg, more:_*)
       
   174     def warning(msg: String, more: (String, JSON.T)*): Unit = message(Markup.WARNING, msg, more:_*)
       
   175     def error_message(msg: String, more: (String, JSON.T)*): Unit =
       
   176       message(Markup.ERROR_MESSAGE, msg, more:_*)
       
   177 
       
   178     val logger: Connection_Logger = new Connection_Logger(context)
       
   179     def progress(): Connection_Progress = new Connection_Progress(context)
       
   180 
       
   181     override def toString: String = connection.toString
       
   182   }
       
   183 
       
   184   class Connection_Logger private[Server](context: Context) extends Logger
       
   185   {
       
   186     def apply(msg: => String): Unit = context.message(Markup.LOGGER, msg)
       
   187 
       
   188     override def toString: String = context.toString
       
   189   }
       
   190 
       
   191   class Connection_Progress private[Server](context: Context) extends Progress
       
   192   {
       
   193     override def echo(msg: String): Unit = context.writeln(msg)
       
   194     override def echo_warning(msg: String): Unit = context.warning(msg)
       
   195     override def echo_error_message(msg: String): Unit = context.error_message(msg)
       
   196     override def theory(session: String, theory: String): Unit =
       
   197       context.writeln(session + ": theory " + theory, "session" -> session, "theory" -> theory)
       
   198 
       
   199     @volatile private var is_stopped = false
       
   200     override def stopped: Boolean = is_stopped
       
   201     def stop { is_stopped = true }
       
   202 
       
   203     override def toString: String = context.toString
   155   }
   204   }
   156 
   205 
   157 
   206 
   158   /* server info */
   207   /* server info */
   159 
   208 
   338 
   387 
   339   override def toString: String = Server.print(port, password)
   388   override def toString: String = Server.print(port, password)
   340 
   389 
   341   private def handle(connection: Server.Connection)
   390   private def handle(connection: Server.Connection)
   342   {
   391   {
       
   392     val context = new Server.Context(server, connection)
       
   393 
   343     connection.read_message() match {
   394     connection.read_message() match {
   344       case Some(msg) if msg == password =>
   395       case Some(msg) if msg == password =>
   345         connection.reply_ok(())
   396         connection.reply_ok(())
   346         var finished = false
   397         var finished = false
   347         while (!finished) {
   398         while (!finished) {
   348           connection.read_message() match {
   399           connection.read_message() match {
   349             case None => finished = true
   400             case None => finished = true
   350             case Some("") =>
   401             case Some("") => context.writeln("Command 'help' provides list of commands")
   351               connection.notify_message("Command 'help' provides list of commands")
       
   352             case Some(msg) =>
   402             case Some(msg) =>
   353               val (name, argument) = Server.Argument.split(msg)
   403               val (name, argument) = Server.Argument.split(msg)
   354               name match {
   404               name match {
   355                 case Server.Command(cmd) =>
   405                 case Server.Command(cmd) =>
   356                   argument match {
   406                   argument match {
   357                     case Server.Argument(arg) =>
   407                     case Server.Argument(arg) =>
   358                       if (cmd.isDefinedAt((server, arg))) {
   408                       if (cmd.isDefinedAt((context, arg))) {
   359                         Exn.capture { cmd((server, arg)) } match {
   409                         Exn.capture { cmd((context, arg)) } match {
   360                           case Exn.Res(res) => connection.reply_ok(res)
   410                           case Exn.Res(res) => connection.reply_ok(res)
   361                           case Exn.Exn(ERROR(msg)) => connection.reply_error(msg)
   411                           case Exn.Exn(ERROR(msg)) => connection.reply_error(msg)
   362                           case Exn.Exn(exn) => throw exn
   412                           case Exn.Exn(exn) => throw exn
   363                         }
   413                         }
   364                       }
   414                       }