author | wenzelm |
Fri, 09 Mar 2018 14:30:13 +0100 | |
changeset 67795 | f8037c7e4659 |
parent 67794 | 82c5ca89ac61 |
child 67796 | bb2bbabf3d37 |
permissions | -rw-r--r-- |
66347 | 1 |
/* Title: Pure/Tools/server.scala |
2 |
Author: Makarius |
|
3 |
||
4 |
Resident Isabelle servers. |
|
5 |
*/ |
|
6 |
||
7 |
package isabelle |
|
8 |
||
9 |
||
66353 | 10 |
import java.io.{BufferedReader, BufferedWriter, InputStreamReader, OutputStreamWriter, |
11 |
IOException} |
|
67793 | 12 |
import java.net.{Socket, SocketException, ServerSocket, InetAddress} |
66347 | 13 |
|
14 |
||
15 |
object Server |
|
16 |
{ |
|
66927 | 17 |
/* protocol */ |
18 |
||
67794 | 19 |
def split_line(line: String): (String, String) = |
20 |
{ |
|
21 |
val head = line.takeWhile(c => Symbol.is_ascii_letter(c) || Symbol.is_ascii_letdig(c)) |
|
22 |
val rest = line.substring(head.length).dropWhile(Symbol.is_ascii_blank(_)) |
|
23 |
(head, proper_string(rest) getOrElse "{}") |
|
24 |
} |
|
25 |
||
67785 | 26 |
val commands: Map[String, PartialFunction[(Server, JSON.T), JSON.T]] = |
66929 | 27 |
Map( |
67785 | 28 |
"echo" -> { case (_, t) => t }, |
29 |
"help" -> { case (_, JSON.empty) => commands.keySet.toList.sorted }, |
|
67791 | 30 |
"shutdown" -> { case (server, JSON.empty) => server.close(); JSON.empty }) |
66929 | 31 |
|
66927 | 32 |
object Reply extends Enumeration |
33 |
{ |
|
34 |
val OK, ERROR = Value |
|
35 |
} |
|
36 |
||
37 |
||
67786 | 38 |
/* socket connection */ |
39 |
||
40 |
object Connection |
|
41 |
{ |
|
42 |
def apply(socket: Socket): Connection = |
|
43 |
new Connection(socket) |
|
44 |
} |
|
45 |
||
67788 | 46 |
class Connection private(val socket: Socket) |
67786 | 47 |
{ |
48 |
override def toString: String = socket.toString |
|
49 |
||
50 |
def close() { socket.close } |
|
51 |
||
52 |
val reader = new BufferedReader(new InputStreamReader(socket.getInputStream, UTF8.charset)) |
|
53 |
val writer = new BufferedWriter(new OutputStreamWriter(socket.getOutputStream, UTF8.charset)) |
|
54 |
||
55 |
def read_line(): Option[String] = |
|
67793 | 56 |
Exn.capture { reader.readLine() } match { |
57 |
case Exn.Res(null) => None |
|
58 |
case Exn.Res(line) => Some(line) |
|
59 |
case Exn.Exn(_: SocketException) => None |
|
60 |
case Exn.Exn(exn) => throw exn |
|
67786 | 61 |
} |
62 |
||
63 |
def write_line(msg: String) |
|
64 |
{ |
|
65 |
require(split_lines(msg).length <= 1) |
|
66 |
writer.write(msg) |
|
67 |
writer.newLine() |
|
68 |
writer.flush() |
|
69 |
} |
|
70 |
||
71 |
def reply(r: Server.Reply.Value, t: JSON.T) |
|
72 |
{ |
|
73 |
write_line(if (t == JSON.empty) r.toString else r.toString + " " + JSON.Format(t)) |
|
74 |
} |
|
75 |
||
76 |
def reply_ok(t: JSON.T) { reply(Server.Reply.OK, t) } |
|
77 |
def reply_error(t: JSON.T) { reply(Server.Reply.ERROR, t) } |
|
78 |
def reply_error_message(message: String, more: (String, JSON.T)*): Unit = |
|
79 |
reply_error(Map("message" -> message) ++ more) |
|
80 |
} |
|
81 |
||
82 |
||
66347 | 83 |
/* per-user servers */ |
84 |
||
67787 | 85 |
def print(port: Int, password: String): String = |
86 |
"127.0.0.1:" + port + " (password " + quote(password) + ")" |
|
87 |
||
66347 | 88 |
object Data |
89 |
{ |
|
66349 | 90 |
val database = Path.explode("$ISABELLE_HOME_USER/servers.db") |
66347 | 91 |
|
66857 | 92 |
val name = SQL.Column.string("name").make_primary_key |
66349 | 93 |
val port = SQL.Column.int("port") |
66347 | 94 |
val password = SQL.Column.string("password") |
66349 | 95 |
val table = SQL.Table("isabelle_servers", List(name, port, password)) |
66347 | 96 |
|
66349 | 97 |
sealed case class Entry(name: String, port: Int, password: String) |
66347 | 98 |
{ |
67787 | 99 |
override def toString: String = |
100 |
"server " + quote(name) + " = " + print(port, password) |
|
66353 | 101 |
|
67786 | 102 |
def connection(): Connection = |
103 |
Connection(new Socket(InetAddress.getByName("127.0.0.1"), port)) |
|
67785 | 104 |
|
105 |
def active(): Boolean = |
|
67786 | 106 |
try { connection().close; true } |
66353 | 107 |
catch { case _: IOException => false } |
66347 | 108 |
} |
109 |
} |
|
110 |
||
111 |
def list(db: SQLite.Database): List[Data.Entry] = |
|
112 |
if (db.tables.contains(Data.table.name)) { |
|
113 |
db.using_statement(Data.table.select())(stmt => |
|
114 |
stmt.execute_query().iterator(res => |
|
115 |
Data.Entry( |
|
66349 | 116 |
res.string(Data.name), |
117 |
res.int(Data.port), |
|
118 |
res.string(Data.password))).toList.sortBy(_.name)) |
|
66347 | 119 |
} |
120 |
else Nil |
|
121 |
||
122 |
def find(db: SQLite.Database, name: String): Option[Data.Entry] = |
|
66353 | 123 |
list(db).find(entry => entry.name == name && entry.active) |
66347 | 124 |
|
67789 | 125 |
def init(name: String = "", port: Int = 0): (Data.Entry, Option[Server]) = |
66347 | 126 |
{ |
127 |
using(SQLite.open_database(Data.database))(db => |
|
128 |
db.transaction { |
|
129 |
find(db, name) match { |
|
66348 | 130 |
case Some(entry) => (entry, None) |
66347 | 131 |
case None => |
66352 | 132 |
val server = new Server(port) |
66347 | 133 |
val entry = Data.Entry(name, server.port, server.password) |
134 |
||
135 |
Isabelle_System.bash("chmod 600 " + File.bash_path(Data.database)).check |
|
136 |
db.create_table(Data.table) |
|
66353 | 137 |
db.using_statement(Data.table.delete(Data.name.where_equal(name)))(_.execute) |
66347 | 138 |
db.using_statement(Data.table.insert())(stmt => |
139 |
{ |
|
66349 | 140 |
stmt.string(1) = entry.name |
141 |
stmt.int(2) = entry.port |
|
66347 | 142 |
stmt.string(3) = entry.password |
143 |
stmt.execute() |
|
144 |
}) |
|
66348 | 145 |
|
67790
1babcc248be0
clarified server start, notably for invocation within regular Isabelle/Scala process;
wenzelm
parents:
67789
diff
changeset
|
146 |
server.start |
67785 | 147 |
(entry, Some(server)) |
66347 | 148 |
} |
149 |
}) |
|
150 |
} |
|
151 |
||
67789 | 152 |
def exit(name: String = ""): Boolean = |
66347 | 153 |
{ |
154 |
using(SQLite.open_database(Data.database))(db => |
|
155 |
db.transaction { |
|
156 |
find(db, name) match { |
|
157 |
case Some(entry) => |
|
67786 | 158 |
using(entry.connection())(connection => |
67785 | 159 |
{ |
67786 | 160 |
connection.write_line(entry.password) |
161 |
connection.write_line("shutdown") |
|
67785 | 162 |
}) |
163 |
while(entry.active) { Thread.sleep(100) } |
|
66347 | 164 |
true |
67785 | 165 |
case None => false |
66347 | 166 |
} |
167 |
}) |
|
168 |
} |
|
169 |
||
170 |
||
171 |
/* Isabelle tool wrapper */ |
|
172 |
||
173 |
val isabelle_tool = |
|
174 |
Isabelle_Tool("server", "manage resident Isabelle servers", args => |
|
175 |
{ |
|
66348 | 176 |
var operation_list = false |
177 |
var name = "" |
|
178 |
var port = 0 |
|
66347 | 179 |
|
180 |
val getopts = |
|
181 |
Getopts(""" |
|
182 |
Usage: isabelle server [OPTIONS] |
|
183 |
||
184 |
Options are: |
|
66348 | 185 |
-L list servers |
186 |
-n NAME explicit server name |
|
187 |
-p PORT explicit server port |
|
66347 | 188 |
|
189 |
Manage resident Isabelle servers. |
|
190 |
""", |
|
66348 | 191 |
"L" -> (_ => operation_list = true), |
192 |
"n:" -> (arg => name = arg), |
|
193 |
"p:" -> (arg => port = Value.Int.parse(arg))) |
|
66347 | 194 |
|
195 |
val more_args = getopts(args) |
|
66348 | 196 |
if (more_args.nonEmpty) getopts.usage() |
66347 | 197 |
|
66353 | 198 |
if (operation_list) { |
199 |
for (entry <- using(SQLite.open_database(Data.database))(list(_)) if entry.active) |
|
67787 | 200 |
Output.writeln(entry.toString, stdout = true) |
66353 | 201 |
} |
66348 | 202 |
else { |
67789 | 203 |
val (entry, server) = init(name, port) |
67787 | 204 |
Output.writeln(entry.toString, stdout = true) |
67785 | 205 |
server.foreach(_.join) |
66348 | 206 |
} |
66347 | 207 |
}) |
208 |
} |
|
209 |
||
66352 | 210 |
class Server private(_port: Int) |
66347 | 211 |
{ |
67785 | 212 |
server => |
213 |
||
66350 | 214 |
private val server_socket = new ServerSocket(_port, 50, InetAddress.getByName("127.0.0.1")) |
67791 | 215 |
|
216 |
def close() { server_socket.close } |
|
217 |
||
66350 | 218 |
def port: Int = server_socket.getLocalPort |
67787 | 219 |
val password: String = Library.UUID() |
66350 | 220 |
|
67787 | 221 |
override def toString: String = Server.print(port, password) |
66348 | 222 |
|
67786 | 223 |
private def handle(connection: Server.Connection) |
66350 | 224 |
{ |
67786 | 225 |
connection.read_line() match { |
226 |
case None => |
|
67792 | 227 |
case Some(line) if line != password => connection.reply_error("Bad protocol") |
66350 | 228 |
case _ => |
229 |
var finished = false |
|
230 |
while (!finished) { |
|
67786 | 231 |
connection.read_line() match { |
232 |
case None => finished = true |
|
67795 | 233 |
case Some(line) if line != "" => |
67794 | 234 |
val (cmd, input) = Server.split_line(line) |
66929 | 235 |
Server.commands.get(cmd) match { |
67786 | 236 |
case None => connection.reply_error("Bad command " + quote(cmd)) |
66929 | 237 |
case Some(body) => |
67794 | 238 |
input match { |
66929 | 239 |
case JSON.Format(arg) => |
67785 | 240 |
if (body.isDefinedAt((server, arg))) { |
67786 | 241 |
try { connection.reply_ok(body(server, arg)) } |
242 |
catch { case ERROR(msg) => connection.reply_error(msg) } |
|
66929 | 243 |
} |
244 |
else { |
|
67786 | 245 |
connection.reply_error_message( |
66929 | 246 |
"Bad argument for command", "command" -> cmd, "argument" -> arg) |
247 |
} |
|
248 |
case _ => |
|
67786 | 249 |
connection.reply_error_message( |
66929 | 250 |
"Malformed command-line", "command" -> cmd, "input" -> input) |
251 |
} |
|
66921 | 252 |
} |
253 |
case _ => |
|
66350 | 254 |
} |
255 |
} |
|
256 |
} |
|
257 |
} |
|
258 |
||
67786 | 259 |
private lazy val server_thread: Thread = |
66350 | 260 |
Standard_Thread.fork("server") { |
261 |
var finished = false |
|
262 |
while (!finished) { |
|
263 |
Exn.capture(server_socket.accept) match { |
|
264 |
case Exn.Res(socket) => |
|
265 |
Standard_Thread.fork("server_connection") |
|
67786 | 266 |
{ using(Server.Connection(socket))(handle(_)) } |
66350 | 267 |
case Exn.Exn(_) => finished = true |
268 |
} |
|
269 |
} |
|
270 |
} |
|
67785 | 271 |
|
67790
1babcc248be0
clarified server start, notably for invocation within regular Isabelle/Scala process;
wenzelm
parents:
67789
diff
changeset
|
272 |
def start { server_thread } |
1babcc248be0
clarified server start, notably for invocation within regular Isabelle/Scala process;
wenzelm
parents:
67789
diff
changeset
|
273 |
|
67791 | 274 |
def join { server_thread.join; close() } |
66347 | 275 |
} |