src/Pure/General/multi_map.scala
author wenzelm
Mon, 25 Mar 2019 16:45:08 +0100
changeset 69980 f2e3adfd916f
parent 64370 865b39487b5d
child 71601 97ccf48c2f0c
permissions -rw-r--r--
tuned signature;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
52975
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
     1
/*  Title:      Pure/General/multi_map.scala
56744
0b74d1df4b8e tuned headers;
wenzelm
parents: 55618
diff changeset
     2
    Author:     Makarius
52975
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
     3
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
     4
Maps with multiple entries per key.
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
     5
*/
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
     6
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
     7
package isabelle
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
     8
55618
995162143ef4 tuned imports;
wenzelm
parents: 55488
diff changeset
     9
63579
73939a9b70a3 support 'abbrevs' within theory header;
wenzelm
parents: 59073
diff changeset
    10
import scala.collection.GenTraversableOnce
52975
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    11
import scala.collection.generic.{ImmutableMapFactory, CanBuildFrom}
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    12
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    13
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    14
object Multi_Map extends ImmutableMapFactory[Multi_Map]
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    15
{
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    16
  private val empty_val: Multi_Map[Any, Nothing] = new Multi_Map[Any, Nothing](Map.empty)
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    17
  override def empty[A, B] = empty_val.asInstanceOf[Multi_Map[A, B]]
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    18
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    19
  implicit def canBuildFrom[A, B]: CanBuildFrom[Coll, (A, B), Multi_Map[A, B]] =
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    20
    new MapCanBuildFrom[A, B]
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    21
}
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    22
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    23
59073
dcecfcc56dce more merge operations;
wenzelm
parents: 56744
diff changeset
    24
final class Multi_Map[A, +B] private(protected val rep: Map[A, List[B]])
52975
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    25
  extends scala.collection.immutable.Map[A, B]
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    26
  with scala.collection.immutable.MapLike[A, B, Multi_Map[A, B]]
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    27
{
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    28
  /* Multi_Map operations */
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    29
55488
60c159d490a2 more integrity checks of theory names vs. full node names;
wenzelm
parents: 52975
diff changeset
    30
  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
    31
52975
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    32
  def get_list(a: A): List[B] = rep.getOrElse(a, Nil)
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    33
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    34
  def insert[B1 >: B](a: A, b: B1): Multi_Map[A, B1] =
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    35
  {
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    36
    val bs = get_list(a)
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    37
    if (bs.contains(b)) this
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    38
    else new Multi_Map(rep + (a -> (b :: bs)))
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    39
  }
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    40
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    41
  def remove[B1 >: B](a: A, b: B1): Multi_Map[A, B1] =
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    42
  {
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    43
    val bs = get_list(a)
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    44
    if (bs.contains(b)) {
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    45
      bs.filterNot(_ == b) match {
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    46
        case Nil => new Multi_Map(rep - a)
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    47
        case bs1 => new Multi_Map(rep + (a -> bs1))
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    48
      }
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    49
    }
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    50
    else this
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    51
  }
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    52
59073
dcecfcc56dce more merge operations;
wenzelm
parents: 56744
diff changeset
    53
  def ++[B1 >: B] (other: Multi_Map[A, B1]): Multi_Map[A, B1] =
dcecfcc56dce more merge operations;
wenzelm
parents: 56744
diff changeset
    54
    if (this eq other) this
dcecfcc56dce more merge operations;
wenzelm
parents: 56744
diff changeset
    55
    else if (isEmpty) other
dcecfcc56dce more merge operations;
wenzelm
parents: 56744
diff changeset
    56
    else
dcecfcc56dce more merge operations;
wenzelm
parents: 56744
diff changeset
    57
      (this.asInstanceOf[Multi_Map[A, B1]] /: other.rep.iterator) {
dcecfcc56dce more merge operations;
wenzelm
parents: 56744
diff changeset
    58
        case (m1, (a, bs)) => (bs :\ m1) { case (b, m2) => m2.insert(a, b) }
dcecfcc56dce more merge operations;
wenzelm
parents: 56744
diff changeset
    59
      }
dcecfcc56dce more merge operations;
wenzelm
parents: 56744
diff changeset
    60
52975
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    61
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    62
  /* Map operations */
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    63
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    64
  override def stringPrefix = "Multi_Map"
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    65
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    66
  override def empty = Multi_Map.empty
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    67
  override def isEmpty: Boolean = rep.isEmpty
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    68
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    69
  override def keySet: Set[A] = rep.keySet
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    70
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    71
  override def iterator: Iterator[(A, B)] =
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    72
    for ((a, bs) <- rep.iterator; b <- bs.iterator) yield (a, b)
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    73
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    74
  def get(a: A): Option[B] = get_list(a).headOption
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    75
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    76
  def + [B1 >: B](p: (A, B1)): Multi_Map[A, B1] = insert(p._1, p._2)
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    77
63579
73939a9b70a3 support 'abbrevs' within theory header;
wenzelm
parents: 59073
diff changeset
    78
  override def ++ [B1 >: B](entries: GenTraversableOnce[(A, B1)]): Multi_Map[A, B1] =
73939a9b70a3 support 'abbrevs' within theory header;
wenzelm
parents: 59073
diff changeset
    79
    (this.asInstanceOf[Multi_Map[A, B1]] /: entries)(_ + _)
73939a9b70a3 support 'abbrevs' within theory header;
wenzelm
parents: 59073
diff changeset
    80
52975
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    81
  def - (a: A): Multi_Map[A, B] =
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    82
    if (rep.isDefinedAt(a)) new Multi_Map(rep - a) else this
457c006f91bb support for maps with multiple entries per key;
wenzelm
parents:
diff changeset
    83
}