| author | wenzelm | 
| Fri, 05 Jun 2015 11:11:26 +0200 | |
| changeset 60369 | f393a3fe884c | 
| parent 60215 | 5fb4990dfc73 | 
| child 61556 | 0d4ee4168e41 | 
| permissions | -rw-r--r-- | 
| 34217 | 1 | /* Title: Pure/Concurrent/future.scala | 
| 45673 
cd41e3903fbf
separate compilation of PIDE vs. Pure sources, which enables independent Scala library;
 wenzelm parents: 
45667diff
changeset | 2 | Module: PIDE | 
| 34217 | 3 | Author: Makarius | 
| 4 | ||
| 57350 | 5 | Value-oriented parallel execution via futures and promises in Scala -- with | 
| 56673 
42bf8fffdf6a
modernized Future/Promise implementation, bypassing old actors;
 wenzelm parents: 
56663diff
changeset | 6 | signatures as in Isabelle/ML. | 
| 34217 | 7 | */ | 
| 8 | ||
| 9 | package isabelle | |
| 10 | ||
| 11 | ||
| 56673 
42bf8fffdf6a
modernized Future/Promise implementation, bypassing old actors;
 wenzelm parents: 
56663diff
changeset | 12 | import scala.util.{Success, Failure}
 | 
| 56731 
326e8a7ea287
just one default thread pool (which is hardwired to Runtime.availableProcessors);
 wenzelm parents: 
56674diff
changeset | 13 | import scala.concurrent.{ExecutionContext, ExecutionContextExecutor,
 | 
| 
326e8a7ea287
just one default thread pool (which is hardwired to Runtime.availableProcessors);
 wenzelm parents: 
56674diff
changeset | 14 | Future => Scala_Future, Promise => Scala_Promise, Await} | 
| 56673 
42bf8fffdf6a
modernized Future/Promise implementation, bypassing old actors;
 wenzelm parents: 
56663diff
changeset | 15 | import scala.concurrent.duration.Duration | 
| 34217 | 16 | |
| 17 | ||
| 18 | object Future | |
| 19 | {
 | |
| 56731 
326e8a7ea287
just one default thread pool (which is hardwired to Runtime.availableProcessors);
 wenzelm parents: 
56674diff
changeset | 20 | lazy val execution_context: ExecutionContextExecutor = | 
| 
326e8a7ea287
just one default thread pool (which is hardwired to Runtime.availableProcessors);
 wenzelm parents: 
56674diff
changeset | 21 | ExecutionContext.fromExecutorService(Simple_Thread.default_pool) | 
| 
326e8a7ea287
just one default thread pool (which is hardwired to Runtime.availableProcessors);
 wenzelm parents: 
56674diff
changeset | 22 | |
| 34217 | 23 | def value[A](x: A): Future[A] = new Finished_Future(x) | 
| 56731 
326e8a7ea287
just one default thread pool (which is hardwired to Runtime.availableProcessors);
 wenzelm parents: 
56674diff
changeset | 24 | |
| 
326e8a7ea287
just one default thread pool (which is hardwired to Runtime.availableProcessors);
 wenzelm parents: 
56674diff
changeset | 25 | def fork[A](body: => A): Future[A] = | 
| 
326e8a7ea287
just one default thread pool (which is hardwired to Runtime.availableProcessors);
 wenzelm parents: 
56674diff
changeset | 26 | new Pending_Future(Scala_Future[A](body)(execution_context)) | 
| 
326e8a7ea287
just one default thread pool (which is hardwired to Runtime.availableProcessors);
 wenzelm parents: 
56674diff
changeset | 27 | |
| 
326e8a7ea287
just one default thread pool (which is hardwired to Runtime.availableProcessors);
 wenzelm parents: 
56674diff
changeset | 28 | def promise[A]: Promise[A] = | 
| 60215 | 29 | new Promise_Future[A](Scala_Promise[A]()) | 
| 34217 | 30 | } | 
| 31 | ||
| 56673 
42bf8fffdf6a
modernized Future/Promise implementation, bypassing old actors;
 wenzelm parents: 
56663diff
changeset | 32 | trait Future[A] | 
| 34217 | 33 | {
 | 
| 34 | def peek: Option[Exn.Result[A]] | |
| 35 | def is_finished: Boolean = peek.isDefined | |
| 38848 
9483bb678d96
use Future.get_finished where this is the intended meaning -- prefer immediate crash over deadlock due to promises that are never fulfilled;
 wenzelm parents: 
34262diff
changeset | 36 |   def get_finished: A = { require(is_finished); Exn.release(peek.get) }
 | 
| 34217 | 37 | def join: A | 
| 38 |   def map[B](f: A => B): Future[B] = Future.fork { f(join) }
 | |
| 39 | ||
| 57912 | 40 | override def toString: String = | 
| 34217 | 41 |     peek match {
 | 
| 42 | case None => "<future>" | |
| 43 | case Some(Exn.Exn(_)) => "<failed>" | |
| 44 | case Some(Exn.Res(x)) => x.toString | |
| 45 | } | |
| 46 | } | |
| 47 | ||
| 56673 
42bf8fffdf6a
modernized Future/Promise implementation, bypassing old actors;
 wenzelm parents: 
56663diff
changeset | 48 | trait Promise[A] extends Future[A] | 
| 34240 
3274571e45c1
added Future.promise -- essentially a single-assignment variable with signalling, using the Future interface;
 wenzelm parents: 
34217diff
changeset | 49 | {
 | 
| 59365 | 50 | def cancel: Unit | 
| 34262 | 51 | def fulfill_result(res: Exn.Result[A]): Unit | 
| 56674 | 52 | def fulfill(x: A): Unit | 
| 34240 
3274571e45c1
added Future.promise -- essentially a single-assignment variable with signalling, using the Future interface;
 wenzelm parents: 
34217diff
changeset | 53 | } | 
| 
3274571e45c1
added Future.promise -- essentially a single-assignment variable with signalling, using the Future interface;
 wenzelm parents: 
34217diff
changeset | 54 | |
| 56673 
42bf8fffdf6a
modernized Future/Promise implementation, bypassing old actors;
 wenzelm parents: 
56663diff
changeset | 55 | |
| 34217 | 56 | private class Finished_Future[A](x: A) extends Future[A] | 
| 57 | {
 | |
| 58 | val peek: Option[Exn.Result[A]] = Some(Exn.Res(x)) | |
| 59 | val join: A = x | |
| 60 | } | |
| 61 | ||
| 56673 
42bf8fffdf6a
modernized Future/Promise implementation, bypassing old actors;
 wenzelm parents: 
56663diff
changeset | 62 | private class Pending_Future[A](future: Scala_Future[A]) extends Future[A] | 
| 34217 | 63 | {
 | 
| 56673 
42bf8fffdf6a
modernized Future/Promise implementation, bypassing old actors;
 wenzelm parents: 
56663diff
changeset | 64 | def peek: Option[Exn.Result[A]] = | 
| 
42bf8fffdf6a
modernized Future/Promise implementation, bypassing old actors;
 wenzelm parents: 
56663diff
changeset | 65 |     future.value match {
 | 
| 
42bf8fffdf6a
modernized Future/Promise implementation, bypassing old actors;
 wenzelm parents: 
56663diff
changeset | 66 | case Some(Success(x)) => Some(Exn.Res(x)) | 
| 
42bf8fffdf6a
modernized Future/Promise implementation, bypassing old actors;
 wenzelm parents: 
56663diff
changeset | 67 | case Some(Failure(e)) => Some(Exn.Exn(e)) | 
| 
42bf8fffdf6a
modernized Future/Promise implementation, bypassing old actors;
 wenzelm parents: 
56663diff
changeset | 68 | case None => None | 
| 
42bf8fffdf6a
modernized Future/Promise implementation, bypassing old actors;
 wenzelm parents: 
56663diff
changeset | 69 | } | 
| 
42bf8fffdf6a
modernized Future/Promise implementation, bypassing old actors;
 wenzelm parents: 
56663diff
changeset | 70 | override def is_finished: Boolean = future.isCompleted | 
| 34217 | 71 | |
| 56673 
42bf8fffdf6a
modernized Future/Promise implementation, bypassing old actors;
 wenzelm parents: 
56663diff
changeset | 72 | def join: A = Await.result(future, Duration.Inf) | 
| 56731 
326e8a7ea287
just one default thread pool (which is hardwired to Runtime.availableProcessors);
 wenzelm parents: 
56674diff
changeset | 73 | override def map[B](f: A => B): Future[B] = | 
| 
326e8a7ea287
just one default thread pool (which is hardwired to Runtime.availableProcessors);
 wenzelm parents: 
56674diff
changeset | 74 | new Pending_Future[B](future.map(f)(Future.execution_context)) | 
| 34217 | 75 | } | 
| 76 | ||
| 56673 
42bf8fffdf6a
modernized Future/Promise implementation, bypassing old actors;
 wenzelm parents: 
56663diff
changeset | 77 | private class Promise_Future[A](promise: Scala_Promise[A]) | 
| 
42bf8fffdf6a
modernized Future/Promise implementation, bypassing old actors;
 wenzelm parents: 
56663diff
changeset | 78 | extends Pending_Future(promise.future) with Promise[A] | 
| 34240 
3274571e45c1
added Future.promise -- essentially a single-assignment variable with signalling, using the Future interface;
 wenzelm parents: 
34217diff
changeset | 79 | {
 | 
| 56673 
42bf8fffdf6a
modernized Future/Promise implementation, bypassing old actors;
 wenzelm parents: 
56663diff
changeset | 80 | override def is_finished: Boolean = promise.isCompleted | 
| 34240 
3274571e45c1
added Future.promise -- essentially a single-assignment variable with signalling, using the Future interface;
 wenzelm parents: 
34217diff
changeset | 81 | |
| 59365 | 82 | def cancel: Unit = | 
| 83 |     try { fulfill_result(Exn.Exn(Exn.Interrupt())) }
 | |
| 84 |     catch { case _: IllegalStateException => }
 | |
| 85 | ||
| 56674 | 86 | def fulfill_result(res: Exn.Result[A]): Unit = | 
| 56673 
42bf8fffdf6a
modernized Future/Promise implementation, bypassing old actors;
 wenzelm parents: 
56663diff
changeset | 87 |     res match {
 | 
| 
42bf8fffdf6a
modernized Future/Promise implementation, bypassing old actors;
 wenzelm parents: 
56663diff
changeset | 88 | case Exn.Res(x) => promise.success(x) | 
| 
42bf8fffdf6a
modernized Future/Promise implementation, bypassing old actors;
 wenzelm parents: 
56663diff
changeset | 89 | case Exn.Exn(e) => promise.failure(e) | 
| 34240 
3274571e45c1
added Future.promise -- essentially a single-assignment variable with signalling, using the Future interface;
 wenzelm parents: 
34217diff
changeset | 90 | } | 
| 56674 | 91 | def fulfill(x: A): Unit = promise.success(x) | 
| 34240 
3274571e45c1
added Future.promise -- essentially a single-assignment variable with signalling, using the Future interface;
 wenzelm parents: 
34217diff
changeset | 92 | } |