16 object Scala |
16 object Scala |
17 { |
17 { |
18 /** registered functions **/ |
18 /** registered functions **/ |
19 |
19 |
20 abstract class Fun(val name: String, val thread: Boolean = false) |
20 abstract class Fun(val name: String, val thread: Boolean = false) |
21 extends Function[String, String] |
|
22 { |
21 { |
23 override def toString: String = name |
22 override def toString: String = name |
|
23 def multi: Boolean = true |
24 def position: Properties.T = here.position |
24 def position: Properties.T = here.position |
25 def here: Scala_Project.Here |
25 def here: Scala_Project.Here |
|
26 def invoke(args: List[Bytes]): List[Bytes] |
|
27 } |
|
28 |
|
29 abstract class Fun_Strings(name: String, thread: Boolean = false) |
|
30 extends Fun(name, thread = thread) |
|
31 { |
|
32 override def invoke(args: List[Bytes]): List[Bytes] = |
|
33 apply(args.map(_.text)).map(Bytes.apply) |
|
34 def apply(args: List[String]): List[String] |
|
35 } |
|
36 |
|
37 abstract class Fun_String(name: String, thread: Boolean = false) |
|
38 extends Fun_Strings(name, thread = thread) |
|
39 { |
|
40 override def multi: Boolean = false |
|
41 override def apply(args: List[String]): List[String] = |
|
42 args match { |
|
43 case List(arg) => List(apply(arg)) |
|
44 case _ => error("Expected single argument for Scala function " + quote(name)) |
|
45 } |
26 def apply(arg: String): String |
46 def apply(arg: String): String |
27 } |
47 } |
28 |
48 |
29 class Functions(val functions: Fun*) extends Isabelle_System.Service |
49 class Functions(val functions: Fun*) extends Isabelle_System.Service |
30 |
50 |
33 |
53 |
34 |
54 |
35 |
55 |
36 /** demo functions **/ |
56 /** demo functions **/ |
37 |
57 |
38 object Echo extends Fun("echo") |
58 object Echo extends Fun_String("echo") |
39 { |
59 { |
40 val here = Scala_Project.here |
60 val here = Scala_Project.here |
41 def apply(arg: String): String = arg |
61 def apply(arg: String): String = arg |
42 } |
62 } |
43 |
63 |
44 object Sleep extends Fun("sleep") |
64 object Sleep extends Fun_String("sleep") |
45 { |
65 { |
46 val here = Scala_Project.here |
66 val here = Scala_Project.here |
47 def apply(seconds: String): String = |
67 def apply(seconds: String): String = |
48 { |
68 { |
49 val t = |
69 val t = |
125 if (!ok && errors.isEmpty) List("Error") else errors |
145 if (!ok && errors.isEmpty) List("Error") else errors |
126 } |
146 } |
127 } |
147 } |
128 } |
148 } |
129 |
149 |
130 object Toplevel extends Fun("scala_toplevel") |
150 object Toplevel extends Fun_String("scala_toplevel") |
131 { |
151 { |
132 val here = Scala_Project.here |
152 val here = Scala_Project.here |
133 def apply(arg: String): String = |
153 def apply(arg: String): String = |
134 { |
154 { |
135 val (interpret, source) = |
155 val (interpret, source) = |
160 functions.find(fun => fun.name == name) match { |
180 functions.find(fun => fun.name == name) match { |
161 case Some(fun) => fun.thread |
181 case Some(fun) => fun.thread |
162 case None => false |
182 case None => false |
163 } |
183 } |
164 |
184 |
165 def function_body(name: String, arg: String): (Tag.Value, String) = |
185 def function_body(name: String, args: List[Bytes]): (Tag.Value, List[Bytes]) = |
166 functions.find(fun => fun.name == name) match { |
186 functions.find(fun => fun.name == name) match { |
167 case Some(fun) => |
187 case Some(fun) => |
168 Exn.capture { fun(arg) } match { |
188 Exn.capture { fun.invoke(args) } match { |
169 case Exn.Res(null) => (Tag.NULL, "") |
189 case Exn.Res(null) => (Tag.NULL, Nil) |
170 case Exn.Res(res) => (Tag.OK, res) |
190 case Exn.Res(res) => (Tag.OK, res) |
171 case Exn.Exn(Exn.Interrupt()) => (Tag.INTERRUPT, "") |
191 case Exn.Exn(Exn.Interrupt()) => (Tag.INTERRUPT, Nil) |
172 case Exn.Exn(e) => (Tag.ERROR, Exn.message(e)) |
192 case Exn.Exn(e) => (Tag.ERROR, List(Bytes(Exn.message(e)))) |
173 } |
193 } |
174 case None => (Tag.FAIL, "Unknown Isabelle/Scala function: " + quote(name)) |
194 case None => (Tag.FAIL, List(Bytes("Unknown Isabelle/Scala function: " + quote(name)))) |
175 } |
195 } |
176 |
196 |
177 |
197 |
178 /* protocol handler */ |
198 /* protocol handler */ |
179 |
199 |
189 { |
209 { |
190 for ((id, future) <- futures) cancel(id, future) |
210 for ((id, future) <- futures) cancel(id, future) |
191 futures = Map.empty |
211 futures = Map.empty |
192 } |
212 } |
193 |
213 |
194 private def result(id: String, tag: Scala.Tag.Value, res: String): Unit = |
214 private def result(id: String, tag: Scala.Tag.Value, res: List[Bytes]): Unit = |
195 synchronized |
215 synchronized |
196 { |
216 { |
197 if (futures.isDefinedAt(id)) { |
217 if (futures.isDefinedAt(id)) { |
198 session.protocol_command("Scala.result", id, tag.id.toString, res) |
218 session.protocol_command_raw("Scala.result", Bytes(id) :: Bytes(tag.id.toString) :: res) |
199 futures -= id |
219 futures -= id |
200 } |
220 } |
201 } |
221 } |
202 |
222 |
203 private def cancel(id: String, future: Future[Unit]): Unit = |
223 private def cancel(id: String, future: Future[Unit]): Unit = |
204 { |
224 { |
205 future.cancel() |
225 future.cancel() |
206 result(id, Scala.Tag.INTERRUPT, "") |
226 result(id, Scala.Tag.INTERRUPT, Nil) |
207 } |
227 } |
208 |
228 |
209 private def invoke_scala(msg: Prover.Protocol_Output): Boolean = synchronized |
229 private def invoke_scala(msg: Prover.Protocol_Output): Boolean = synchronized |
210 { |
230 { |
211 msg.properties match { |
231 msg.properties match { |
212 case Markup.Invoke_Scala(name, id) => |
232 case Markup.Invoke_Scala(name, id) => |
213 def body: Unit = |
233 def body: Unit = |
214 { |
234 { |
215 val (tag, res) = Scala.function_body(name, msg.text) |
235 val (tag, res) = Scala.function_body(name, msg.chunks) |
216 result(id, tag, res) |
236 result(id, tag, res) |
217 } |
237 } |
218 val future = |
238 val future = |
219 if (Scala.function_thread(name)) { |
239 if (Scala.function_thread(name)) { |
220 Future.thread(name = Isabelle_Thread.make_name(base = "invoke_scala"))(body) |
240 Future.thread(name = Isabelle_Thread.make_name(base = "invoke_scala"))(body) |