| 59136 |      1 | /*  Title:      Pure/Concurrent/par_list.scala
 | 
|  |      2 |     Author:     Makarius
 | 
|  |      3 | 
 | 
|  |      4 | Parallel list combinators.
 | 
|  |      5 | */
 | 
|  |      6 | 
 | 
|  |      7 | 
 | 
|  |      8 | package isabelle
 | 
|  |      9 | 
 | 
|  |     10 | 
 | 
|  |     11 | import java.util.concurrent.{Future => JFuture, CancellationException}
 | 
|  |     12 | 
 | 
|  |     13 | 
 | 
|  |     14 | object Par_List
 | 
|  |     15 | {
 | 
|  |     16 |   def managed_results[A, B](f: A => B, xs: List[A]): List[Exn.Result[B]] =
 | 
|  |     17 |     if (xs.isEmpty || xs.tail.isEmpty) xs.map(x => Exn.capture { f(x) })
 | 
|  |     18 |     else {
 | 
|  |     19 |       val state = Synchronized((List.empty[JFuture[Exn.Result[B]]], false))
 | 
|  |     20 | 
 | 
|  |     21 |       def cancel_other(self: Int = -1): Unit =
 | 
|  |     22 |         state.change { case (tasks, canceled) =>
 | 
|  |     23 |           if (!canceled) {
 | 
|  |     24 |             for ((task, i) <- tasks.iterator.zipWithIndex if i != self)
 | 
|  |     25 |               task.cancel(true)
 | 
|  |     26 |           }
 | 
|  |     27 |           (tasks, true)
 | 
|  |     28 |         }
 | 
|  |     29 | 
 | 
|  |     30 |       try {
 | 
|  |     31 |         state.change(_ =>
 | 
|  |     32 |           (xs.iterator.zipWithIndex.map({ case (x, self) =>
 | 
|  |     33 |             Simple_Thread.submit_task {
 | 
|  |     34 |               val result = Exn.capture { f(x) }
 | 
|  |     35 |               result match { case Exn.Exn(_) => cancel_other(self) case _ => }
 | 
|  |     36 |               result
 | 
|  |     37 |             }
 | 
|  |     38 |           }).toList, false))
 | 
|  |     39 | 
 | 
|  |     40 |         state.value._1.map(future =>
 | 
|  |     41 |           try { future.get }
 | 
|  |     42 |           catch { case _: CancellationException => Exn.Exn(Exn.Interrupt()): Exn.Result[B] })
 | 
|  |     43 |       }
 | 
|  |     44 |       finally { cancel_other() }
 | 
|  |     45 |     }
 | 
|  |     46 | 
 | 
|  |     47 |   def map[A, B](f: A => B, xs: List[A]): List[B] =
 | 
|  |     48 |     Exn.release_first(managed_results(f, xs))
 | 
|  |     49 | 
 | 
|  |     50 |   def get_some[A, B](f: A => Option[B], xs: List[A]): Option[B] =
 | 
|  |     51 |   {
 | 
|  |     52 |     class Found(val res: B) extends Exception
 | 
|  |     53 |     val results =
 | 
|  |     54 |       managed_results(
 | 
|  |     55 |         (x: A) => f(x) match { case None => () case Some(y) => throw new Found(y) }, xs)
 | 
|  |     56 |     results.collectFirst({ case Exn.Exn(found: Found) => found.res }) match {
 | 
|  |     57 |       case None => Exn.release_first(results); None
 | 
|  |     58 |       case some => some
 | 
|  |     59 |     }
 | 
|  |     60 |   }
 | 
|  |     61 | 
 | 
|  |     62 |   def find_some[A](P: A => Boolean, xs: List[A]): Option[A] =
 | 
|  |     63 |     get_some((x: A) => if (P(x)) Some(x) else None, xs)
 | 
|  |     64 | 
 | 
| 59137 |     65 |   def exists[A](P: A => Boolean, xs: List[A]): Boolean = find_some(P, xs).isDefined
 | 
| 59136 |     66 |   def forall[A](P: A => Boolean, xs: List[A]): Boolean = !exists((x: A) => !P(x), xs)
 | 
|  |     67 | }
 | 
|  |     68 | 
 |