| author | wenzelm | 
| Sat, 11 Feb 2023 23:02:51 +0100 | |
| changeset 77255 | b810e99b5afb | 
| parent 75393 | 87ebf5a50283 | 
| child 82142 | 508a673c87ac | 
| permissions | -rw-r--r-- | 
| 52975 | 1  | 
/* Title: Pure/General/multi_map.scala  | 
| 56744 | 2  | 
Author: Makarius  | 
| 52975 | 3  | 
|
4  | 
Maps with multiple entries per key.  | 
|
5  | 
*/  | 
|
6  | 
||
7  | 
package isabelle  | 
|
8  | 
||
| 73136 | 9  | 
import scala.collection.mutable  | 
10  | 
import scala.collection.{IterableFactory, MapFactory, MapFactoryDefaults}
 | 
|
11  | 
import scala.collection.immutable.{Iterable, MapOps}
 | 
|
| 52975 | 12  | 
|
13  | 
||
| 75393 | 14  | 
object Multi_Map extends MapFactory[Multi_Map] {
 | 
| 52975 | 15  | 
private val empty_val: Multi_Map[Any, Nothing] = new Multi_Map[Any, Nothing](Map.empty)  | 
| 73136 | 16  | 
def empty[A, B]: Multi_Map[A, B] = empty_val.asInstanceOf[Multi_Map[A, B]]  | 
17  | 
||
18  | 
def from[A, B](entries: IterableOnce[(A, B)]): Multi_Map[A, B] =  | 
|
| 73362 | 19  | 
    entries.iterator.foldLeft(empty[A, B]) { case (m, (a, b)) => m.insert(a, b) }
 | 
| 52975 | 20  | 
|
| 73136 | 21  | 
override def newBuilder[A, B]: mutable.Builder[(A, B), Multi_Map[A, B]] = new Builder[A, B]  | 
| 75393 | 22  | 
  private class Builder[A, B] extends mutable.Builder[(A, B), Multi_Map[A, B]] {
 | 
| 73136 | 23  | 
private var res = empty[A, B]  | 
| 73340 | 24  | 
    override def clear(): Unit = { res = empty[A, B] }
 | 
| 73136 | 25  | 
    override def addOne(p: (A, B)): this.type = { res = res.insert(p._1, p._2); this }
 | 
26  | 
override def result(): Multi_Map[A, B] = res  | 
|
27  | 
}  | 
|
| 52975 | 28  | 
}  | 
29  | 
||
| 59073 | 30  | 
final class Multi_Map[A, +B] private(protected val rep: Map[A, List[B]])  | 
| 73136 | 31  | 
extends Iterable[(A, B)]  | 
32  | 
with MapOps[A, B, Multi_Map, Multi_Map[A, B]]  | 
|
| 75393 | 33  | 
    with MapFactoryDefaults[A, B, Multi_Map, Iterable] {
 | 
| 52975 | 34  | 
/* Multi_Map operations */  | 
35  | 
||
| 
55488
 
60c159d490a2
more integrity checks of theory names vs. full node names;
 
wenzelm 
parents: 
52975 
diff
changeset
 | 
36  | 
def iterator_list: Iterator[(A, List[B])] = rep.iterator  | 
| 
 
60c159d490a2
more integrity checks of theory names vs. full node names;
 
wenzelm 
parents: 
52975 
diff
changeset
 | 
37  | 
|
| 52975 | 38  | 
def get_list(a: A): List[B] = rep.getOrElse(a, Nil)  | 
39  | 
||
| 75393 | 40  | 
  def insert[B1 >: B](a: A, b: B1): Multi_Map[A, B1] = {
 | 
| 52975 | 41  | 
val bs = get_list(a)  | 
42  | 
if (bs.contains(b)) this  | 
|
43  | 
else new Multi_Map(rep + (a -> (b :: bs)))  | 
|
44  | 
}  | 
|
45  | 
||
| 75393 | 46  | 
  def remove[B1 >: B](a: A, b: B1): Multi_Map[A, B1] = {
 | 
| 52975 | 47  | 
val bs = get_list(a)  | 
48  | 
    if (bs.contains(b)) {
 | 
|
49  | 
      bs.filterNot(_ == b) match {
 | 
|
50  | 
case Nil => new Multi_Map(rep - a)  | 
|
51  | 
case bs1 => new Multi_Map(rep + (a -> bs1))  | 
|
52  | 
}  | 
|
53  | 
}  | 
|
54  | 
else this  | 
|
55  | 
}  | 
|
56  | 
||
| 59073 | 57  | 
def ++[B1 >: B] (other: Multi_Map[A, B1]): Multi_Map[A, B1] =  | 
58  | 
if (this eq other) this  | 
|
59  | 
else if (isEmpty) other  | 
|
60  | 
else  | 
|
| 73359 | 61  | 
      other.rep.iterator.foldLeft(this.asInstanceOf[Multi_Map[A, B1]]) {
 | 
| 73360 | 62  | 
        case (m1, (a, bs)) => bs.foldRight(m1) { case (b, m2) => m2.insert(a, b) }
 | 
| 59073 | 63  | 
}  | 
64  | 
||
| 52975 | 65  | 
|
66  | 
/* Map operations */  | 
|
67  | 
||
| 71601 | 68  | 
override def empty: Multi_Map[A, Nothing] = Multi_Map.empty  | 
| 52975 | 69  | 
override def isEmpty: Boolean = rep.isEmpty  | 
70  | 
||
71  | 
override def keySet: Set[A] = rep.keySet  | 
|
72  | 
||
73  | 
override def iterator: Iterator[(A, B)] =  | 
|
74  | 
for ((a, bs) <- rep.iterator; b <- bs.iterator) yield (a, b)  | 
|
75  | 
||
76  | 
def get(a: A): Option[B] = get_list(a).headOption  | 
|
77  | 
||
| 73136 | 78  | 
override def updated[B1 >: B](a: A, b: B1): Multi_Map[A, B1] = insert(a, b)  | 
79  | 
||
80  | 
override def removed(a: A): Multi_Map[A, B] =  | 
|
81  | 
if (rep.isDefinedAt(a)) new Multi_Map(rep - a) else this  | 
|
| 52975 | 82  | 
|
| 73136 | 83  | 
override def mapFactory: MapFactory[Multi_Map] = Multi_Map  | 
| 63579 | 84  | 
|
| 73136 | 85  | 
  override def toString: String = mkString("Multi_Map(", ", ", ")")
 | 
| 52975 | 86  | 
}  |