Documentation

Std.Data.HashMap.Raw

Hash maps with unbundled well-formedness invariant #

This module develops the type Std.HashMap.Raw of dependent hash maps with unbundled well-formedness invariant.

This version is safe to use in nested inductive types. The well-formedness predicate is available as Std.HashMap.Raw.WF and we prove in this file that all operations preserve well-formedness. When in doubt, prefer HashMap over HashMap.Raw.

Lemmas about the operations on Std.HashMap.Raw are available in the module Std.Data.HashMap.RawLemmas.

structure Std.HashMap.Raw (α : Type u) (β : Type v) :
Type (max u v)

Hash maps without a bundled well-formedness invariant, suitable for use in nested inductive types. The well-formedness invariant is called Raw.WF. When in doubt, prefer HashMap over HashMap.Raw. Lemmas about the operations on Std.Data.HashMap.Raw are available in the module Std.Data.HashMap.RawLemmas.

This is a simple separate-chaining hash table. The data of the hash map consists of a cached size and an array of buckets, where each bucket is a linked list of key-value pairs. The number of buckets is always a power of two. The hash map doubles its size upon inserting an element such that the number of elements is more than 75% of the number of buckets.

The hash table is backed by an Array. Users should make sure that the hash map is used linearly to avoid expensive copies.

The hash map uses == (provided by the BEq typeclass) to compare keys and hash (provided by the Hashable typeclass) to hash them. To ensure that the operations behave as expected, == should be an equivalence relation and a == b should imply hash a = hash b (see also the EquivBEq and LawfulHashable typeclasses). Both of these conditions are automatic if the BEq instance is lawful, i.e., if a == b implies a = b.

Dependent hash maps, in which keys may occur in their values' types, are available as Std.Data.Raw.DHashMap.

  • inner : DHashMap.Raw α fun (x : α) => β

    Internal implementation detail of the hash map

Instances For
    @[inline]
    def Std.HashMap.Raw.emptyWithCapacity {α : Type u} {β : Type v} (capacity : Nat := 8) :
    Raw α β

    Creates a new empty hash map. The optional parameter capacity can be supplied to presize the map so that it can hold the given number of mappings without reallocating. It is also possible to use the empty collection notations and {} to create an empty hash map with the default capacity.

    Equations
    Instances For
      @[reducible, inline, deprecated Std.HashMap.Raw.emptyWithCapacity (since := "2025-03-12")]
      abbrev Std.HashMap.Raw.empty {α : Type u_1} {β : Type u_2} (capacity : Nat := 8) :
      Raw α β

      Creates a new empty hash map. The optional parameter capacity can be supplied to presize the map so that it can hold the given number of mappings without reallocating. It is also possible to use the empty collection notations and {} to create an empty hash map with the default capacity.

      Equations
      Instances For
        instance Std.HashMap.Raw.instInhabited {α : Type u} {β : Type v} :
        Inhabited (Raw α β)
        Equations
        structure Std.HashMap.Raw.Equiv {α : Type u} {β : Type v} (m₁ m₂ : Raw α β) :

        Two hash maps are equivalent in the sense of Equiv iff all the keys and values are equal.

        • inner : m₁.inner.Equiv m₂.inner

          Internal implementation detail of the hash map

        Instances For

          Two hash maps are equivalent in the sense of Equiv iff all the keys and values are equal.

          Equations
          Instances For
            @[inline]
            def Std.HashMap.Raw.insert {α : Type u} {β : Type v} [beq : BEq α] [Hashable α] (m : Raw α β) (a : α) (b : β) :
            Raw α β

            Inserts the given mapping into the map. If there is already a mapping for the given key, then both key and value will be replaced.

            Note: this replacement behavior is true for HashMap, DHashMap, HashMap.Raw and DHashMap.Raw. The insert function on HashSet and HashSet.Raw behaves differently: it will return the set unchanged if a matching key is already present.

            Equations
            Instances For
              instance Std.HashMap.Raw.instSingletonProdOfBEqOfHashable {α : Type u} {β : Type v} [BEq α] [Hashable α] :
              Singleton (α × β) (Raw α β)
              Equations
              instance Std.HashMap.Raw.instInsertProdOfBEqOfHashable {α : Type u} {β : Type v} [BEq α] [Hashable α] :
              Insert (α × β) (Raw α β)
              Equations
              instance Std.HashMap.Raw.instLawfulSingletonProd {α : Type u} {β : Type v} [BEq α] [Hashable α] :
              LawfulSingleton (α × β) (Raw α β)
              @[inline]
              def Std.HashMap.Raw.insertIfNew {α : Type u} {β : Type v} [BEq α] [Hashable α] (m : Raw α β) (a : α) (b : β) :
              Raw α β

              If there is no mapping for the given key, inserts the given mapping into the map. Otherwise, returns the map unaltered.

              Equations
              Instances For
                @[inline]
                def Std.HashMap.Raw.containsThenInsert {α : Type u} {β : Type v} [BEq α] [Hashable α] (m : Raw α β) (a : α) (b : β) :
                Bool × Raw α β

                Checks whether a key is present in a map, and unconditionally inserts a value for the key.

                Equivalent to (but potentially faster than) calling contains followed by insert.

                Equations
                Instances For
                  @[inline]
                  def Std.HashMap.Raw.containsThenInsertIfNew {α : Type u} {β : Type v} [BEq α] [Hashable α] (m : Raw α β) (a : α) (b : β) :
                  Bool × Raw α β

                  Checks whether a key is present in a map and inserts a value for the key if it was not found.

                  If the returned Bool is true, then the returned map is unaltered. If the Bool is false, then the returned map has a new value inserted.

                  Equivalent to (but potentially faster than) calling contains followed by insertIfNew.

                  Equations
                  Instances For
                    @[inline]
                    def Std.HashMap.Raw.getThenInsertIfNew? {α : Type u} {β : Type v} [BEq α] [Hashable α] (m : Raw α β) (a : α) (b : β) :
                    Option β × Raw α β

                    Equivalent to (but potentially faster than) calling get? followed by insertIfNew.

                    Equations
                    Instances For
                      @[inline]
                      def Std.HashMap.Raw.get? {α : Type u} {β : Type v} [beq : BEq α] [Hashable α] (m : Raw α β) (a : α) :

                      The notation m[a]? is preferred over calling this function directly.

                      Tries to retrieve the mapping for the given key, returning none if no such mapping is present.

                      Equations
                      Instances For
                        @[inline]
                        def Std.HashMap.Raw.contains {α : Type u} {β : Type v} [BEq α] [Hashable α] (m : Raw α β) (a : α) :

                        Returns true if there is a mapping for the given key. There is also a Prop-valued version of this: a ∈ m is equivalent to m.contains a = true.

                        Observe that this is different behavior than for lists: for lists, uses = and contains uses == for comparisons, while for hash maps, both use ==.

                        Equations
                        Instances For
                          instance Std.HashMap.Raw.instMembershipOfBEqOfHashable {α : Type u} {β : Type v} [BEq α] [Hashable α] :
                          Membership α (Raw α β)
                          Equations
                          instance Std.HashMap.Raw.instDecidableMem {α : Type u} {β : Type v} [BEq α] [Hashable α] {m : Raw α β} {a : α} :
                          Equations
                          @[inline]
                          def Std.HashMap.Raw.get {α : Type u} {β : Type v} [BEq α] [Hashable α] (m : Raw α β) (a : α) (h : a m) :
                          β

                          The notation m[a] or m[a]'h is preferred over calling this function directly.

                          Retrieves the mapping for the given key. Ensures that such a mapping exists by requiring a proof of a ∈ m.

                          Equations
                          Instances For
                            @[inline]
                            def Std.HashMap.Raw.getD {α : Type u} {β : Type v} [BEq α] [Hashable α] (m : Raw α β) (a : α) (fallback : β) :
                            β

                            Tries to retrieve the mapping for the given key, returning fallback if no such mapping is present.

                            Equations
                            Instances For
                              @[inline]
                              def Std.HashMap.Raw.get! {α : Type u} {β : Type v} [BEq α] [Hashable α] [Inhabited β] (m : Raw α β) (a : α) :
                              β

                              The notation m[a]! is preferred over calling this function directly.

                              Tries to retrieve the mapping for the given key, panicking if no such mapping is present.

                              Equations
                              Instances For
                                instance Std.HashMap.Raw.instGetElem?Mem {α : Type u} {β : Type v} [BEq α] [Hashable α] :
                                GetElem? (Raw α β) α β fun (m : Raw α β) (a : α) => a m
                                Equations
                                • One or more equations did not get rendered due to their size.
                                @[inline]
                                def Std.HashMap.Raw.getKey? {α : Type u} {β : Type v} [BEq α] [Hashable α] (m : Raw α β) (a : α) :

                                Checks if a mapping for the given key exists and returns the key if it does, otherwise none. The result in the some case is guaranteed to be pointer equal to the key in the map.

                                Equations
                                Instances For
                                  @[inline]
                                  def Std.HashMap.Raw.getKey {α : Type u} {β : Type v} [BEq α] [Hashable α] (m : Raw α β) (a : α) (h : a m) :
                                  α

                                  Retrieves the key from the mapping that matches a. Ensures that such a mapping exists by requiring a proof of a ∈ m. The result is guaranteed to be pointer equal to the key in the map.

                                  Equations
                                  Instances For
                                    @[inline]
                                    def Std.HashMap.Raw.getKeyD {α : Type u} {β : Type v} [BEq α] [Hashable α] (m : Raw α β) (a fallback : α) :
                                    α

                                    Checks if a mapping for the given key exists and returns the key if it does, otherwise fallback. If a mapping exists the result is guaranteed to be pointer equal to the key in the map.

                                    Equations
                                    Instances For
                                      @[inline]
                                      def Std.HashMap.Raw.getKey! {α : Type u} {β : Type v} [BEq α] [Hashable α] [Inhabited α] (m : Raw α β) (a : α) :
                                      α

                                      Checks if a mapping for the given key exists and returns the key if it does, otherwise panics. If no panic occurs the result is guaranteed to be pointer equal to the key in the map.

                                      Equations
                                      Instances For
                                        @[inline]
                                        def Std.HashMap.Raw.erase {α : Type u} {β : Type v} [BEq α] [Hashable α] (m : Raw α β) (a : α) :
                                        Raw α β

                                        Removes the mapping for the given key if it exists.

                                        Equations
                                        Instances For
                                          @[inline]
                                          def Std.HashMap.Raw.size {α : Type u} {β : Type v} (m : Raw α β) :

                                          The number of mappings present in the hash map

                                          Equations
                                          Instances For
                                            @[inline]
                                            def Std.HashMap.Raw.isEmpty {α : Type u} {β : Type v} (m : Raw α β) :

                                            Returns true if the hash map contains no mappings.

                                            Note that if your BEq instance is not reflexive or your Hashable instance is not lawful, then it is possible that this function returns false even though is not possible to get anything out of the hash map.

                                            Equations
                                            Instances For
                                              @[inline]
                                              def Std.HashMap.Raw.keys {α : Type u} {β : Type v} (m : Raw α β) :
                                              List α

                                              Returns a list of all keys present in the hash map in some order.

                                              Equations
                                              Instances For
                                                @[inline]
                                                def Std.HashMap.Raw.ofList {α : Type u} {β : Type v} [BEq α] [Hashable α] (l : List (α × β)) :
                                                Raw α β

                                                Creates a hash map from a list of mappings. If the same key appears multiple times, the last occurrence takes precedence.

                                                Equations
                                                Instances For
                                                  @[inline]
                                                  def Std.HashMap.Raw.unitOfList {α : Type u} [BEq α] [Hashable α] (l : List α) :

                                                  Creates a hash map from a list of keys, associating the value () with each key.

                                                  This is mainly useful to implement HashSet.ofList, so if you are considering using this, HashSet or HashSet.Raw might be a better fit for you.

                                                  Equations
                                                  Instances For
                                                    @[inline]
                                                    def Std.HashMap.Raw.alter {α : Type u} {β : Type v} [BEq α] [EquivBEq α] [Hashable α] (m : Raw α β) (a : α) (f : Option βOption β) :
                                                    Raw α β

                                                    Modifies in place the value associated with a given key, allowing creating new values and deleting values via an Option valued replacement function.

                                                    This function ensures that the value is used linearly.

                                                    Equations
                                                    Instances For
                                                      @[inline]
                                                      def Std.HashMap.Raw.modify {α : Type u} {β : Type v} [BEq α] [EquivBEq α] [Hashable α] (m : Raw α β) (a : α) (f : ββ) :
                                                      Raw α β

                                                      Modifies in place the value associated with a given key.

                                                      This function ensures that the value is used linearly.

                                                      Equations
                                                      Instances For
                                                        @[inline]
                                                        def Std.HashMap.Raw.toList {α : Type u} {β : Type v} (m : Raw α β) :
                                                        List (α × β)

                                                        Transforms the hash map into a list of mappings in some order.

                                                        Equations
                                                        Instances For
                                                          @[inline]
                                                          def Std.HashMap.Raw.foldM {α : Type u} {β : Type v} {m : Type w → Type w'} [Monad m] {γ : Type w} (f : γαβm γ) (init : γ) (b : Raw α β) :
                                                          m γ

                                                          Monadically computes a value by folding the given function over the mappings in the hash map in some order.

                                                          Equations
                                                          Instances For
                                                            @[inline]
                                                            def Std.HashMap.Raw.fold {α : Type u} {β : Type v} {γ : Type w} (f : γαβγ) (init : γ) (b : Raw α β) :
                                                            γ

                                                            Folds the given function over the mappings in the hash map in some order.

                                                            Equations
                                                            Instances For
                                                              @[inline]
                                                              def Std.HashMap.Raw.forM {α : Type u} {β : Type v} {m : Type w → Type w'} [Monad m] (f : αβm PUnit) (b : Raw α β) :

                                                              Carries out a monadic action on each mapping in the hash map in some order.

                                                              Equations
                                                              Instances For
                                                                @[inline]
                                                                def Std.HashMap.Raw.forIn {α : Type u} {β : Type v} {m : Type w → Type w'} [Monad m] {γ : Type w} (f : αβγm (ForInStep γ)) (init : γ) (b : Raw α β) :
                                                                m γ

                                                                Support for the for loop construct in do blocks.

                                                                Equations
                                                                Instances For
                                                                  instance Std.HashMap.Raw.instForMProd {α : Type u} {β : Type v} {m : Type w → Type w'} :
                                                                  ForM m (Raw α β) (α × β)
                                                                  Equations
                                                                  instance Std.HashMap.Raw.instForInProd {α : Type u} {β : Type v} {m : Type w → Type w'} :
                                                                  ForIn m (Raw α β) (α × β)
                                                                  Equations
                                                                  • One or more equations did not get rendered due to their size.
                                                                  @[inline]
                                                                  def Std.HashMap.Raw.union {α : Type u} {β : Type v} [BEq α] [Hashable α] (m₁ m₂ : Raw α β) :
                                                                  Raw α β

                                                                  Computes the union of the given hash maps, inserting smaller hash map into a bigger hash map. In the case of clashes of keys, entries from the left argument, are replaced with entries from the right argument.

                                                                  Equations
                                                                  Instances For
                                                                    @[inline]
                                                                    def Std.HashMap.Raw.filterMap {α : Type u} {β : Type v} {γ : Type w} (f : αβOption γ) (m : Raw α β) :
                                                                    Raw α γ

                                                                    Updates the values of the hash map by applying the given function to all mappings, keeping only those mappings where the function returns some value.

                                                                    Equations
                                                                    Instances For
                                                                      @[inline]
                                                                      def Std.HashMap.Raw.map {α : Type u} {β : Type v} {γ : Type w} (f : αβγ) (m : Raw α β) :
                                                                      Raw α γ

                                                                      Updates the values of the hash map by applying the given function to all mappings.

                                                                      Equations
                                                                      Instances For
                                                                        @[inline]
                                                                        def Std.HashMap.Raw.filter {α : Type u} {β : Type v} (f : αβBool) (m : Raw α β) :
                                                                        Raw α β

                                                                        Removes all mappings of the hash map for which the given function returns false.

                                                                        Equations
                                                                        Instances For
                                                                          @[inline]
                                                                          def Std.HashMap.Raw.toArray {α : Type u} {β : Type v} (m : Raw α β) :
                                                                          Array (α × β)

                                                                          Transforms the hash map into an array of mappings in some order.

                                                                          Equations
                                                                          Instances For
                                                                            @[inline]
                                                                            def Std.HashMap.Raw.keysArray {α : Type u} {β : Type v} (m : Raw α β) :

                                                                            Returns an array of all keys present in the hash map in some order.

                                                                            Equations
                                                                            Instances For

                                                                              We currently do not provide lemmas for the functions below.

                                                                              @[inline]
                                                                              def Std.HashMap.Raw.values {α : Type u} {β : Type v} (m : Raw α β) :
                                                                              List β

                                                                              Returns a list of all values present in the hash map in some order.

                                                                              Equations
                                                                              Instances For
                                                                                @[inline]
                                                                                def Std.HashMap.Raw.valuesArray {α : Type u} {β : Type v} (m : Raw α β) :

                                                                                Returns an array of all values present in the hash map in some order.

                                                                                Equations
                                                                                Instances For
                                                                                  @[inline]
                                                                                  def Std.HashMap.Raw.insertMany {α : Type u} {β : Type v} [BEq α] [Hashable α] {ρ : Type w} [ForIn Id ρ (α × β)] (m : Raw α β) (l : ρ) :
                                                                                  Raw α β

                                                                                  Inserts multiple mappings into the hash map by iterating over the given collection and calling insert. If the same key appears multiple times, the last occurrence takes precedence.

                                                                                  Note: this precedence behavior is true for HashMap, DHashMap, HashMap.Raw and DHashMap.Raw. The insertMany function on HashSet and HashSet.Raw behaves differently: it will prefer the first appearance.

                                                                                  Equations
                                                                                  Instances For
                                                                                    @[inline]
                                                                                    def Std.HashMap.Raw.insertManyIfNewUnit {α : Type u} [BEq α] [Hashable α] {ρ : Type w} [ForIn Id ρ α] (m : Raw α Unit) (l : ρ) :

                                                                                    Inserts multiple keys with the value () into the hash map by iterating over the given collection and calling insertIfNew. If the same key appears multiple times, the first occurrence takes precedence.

                                                                                    This is mainly useful to implement HashSet.insertMany, so if you are considering using this, HashSet or HashSet.Raw might be a better fit for you.

                                                                                    Equations
                                                                                    Instances For
                                                                                      @[inline]
                                                                                      def Std.HashMap.Raw.unitOfArray {α : Type u} [BEq α] [Hashable α] (l : Array α) :

                                                                                      Creates a hash map from an array of keys, associating the value () with each key.

                                                                                      This is mainly useful to implement HashSet.ofArray, so if you are considering using this, HashSet or HashSet.Raw might be a better fit for you.

                                                                                      Equations
                                                                                      Instances For
                                                                                        def Std.HashMap.Raw.Internal.numBuckets {α : Type u} {β : Type v} (m : Raw α β) :

                                                                                        Returns the number of buckets in the internal representation of the hash map. This function may be useful for things like monitoring system health, but it should be considered an internal implementation detail.

                                                                                        Equations
                                                                                        Instances For
                                                                                          instance Std.HashMap.Raw.instRepr {α : Type u} {β : Type v} [Repr α] [Repr β] :
                                                                                          Repr (Raw α β)
                                                                                          Equations
                                                                                          structure Std.HashMap.Raw.WF {α : Type u} {β : Type v} [BEq α] [Hashable α] (m : Raw α β) :

                                                                                          Well-formedness predicate for hash maps. Users of HashMap will not need to interact with this. Users of HashMap.Raw will need to provide proofs of WF to lemmas and should use lemmas WF.empty and WF.insert (which are always named exactly like the operations they are about) to show that map operations preserve well-formedness.

                                                                                          • out : m.inner.WF

                                                                                            Internal implementation detail of the hash map

                                                                                          Instances For
                                                                                            theorem Std.HashMap.Raw.WF.empty {α : Type u} {β : Type v} [BEq α] [Hashable α] :
                                                                                            @[reducible, inline, deprecated Std.HashMap.Raw.WF.empty (since := "2025-03-12")]
                                                                                            abbrev Std.HashMap.Raw.WF.emptyc {α : Type u_1} {β : Type u_2} [BEq α] [Hashable α] :
                                                                                            Equations
                                                                                            Instances For
                                                                                              theorem Std.HashMap.Raw.WF.insert {α : Type u} {β : Type v} [BEq α] [Hashable α] {m : Raw α β} {a : α} {b : β} (h : m.WF) :
                                                                                              (m.insert a b).WF
                                                                                              theorem Std.HashMap.Raw.WF.insertIfNew {α : Type u} {β : Type v} [BEq α] [Hashable α] {m : Raw α β} {a : α} {b : β} (h : m.WF) :
                                                                                              (m.insertIfNew a b).WF
                                                                                              theorem Std.HashMap.Raw.WF.containsThenInsert {α : Type u} {β : Type v} [BEq α] [Hashable α] {m : Raw α β} {a : α} {b : β} (h : m.WF) :
                                                                                              theorem Std.HashMap.Raw.WF.containsThenInsertIfNew {α : Type u} {β : Type v} [BEq α] [Hashable α] {m : Raw α β} {a : α} {b : β} (h : m.WF) :
                                                                                              theorem Std.HashMap.Raw.WF.getThenInsertIfNew? {α : Type u} {β : Type v} [BEq α] [Hashable α] {m : Raw α β} {a : α} {b : β} (h : m.WF) :
                                                                                              theorem Std.HashMap.Raw.WF.erase {α : Type u} {β : Type v} [BEq α] [Hashable α] {m : Raw α β} {a : α} (h : m.WF) :
                                                                                              (m.erase a).WF
                                                                                              theorem Std.HashMap.Raw.WF.filter {α : Type u} {β : Type v} [BEq α] [Hashable α] {m : Raw α β} {f : αβBool} (h : m.WF) :
                                                                                              theorem Std.HashMap.Raw.WF.insertMany {α : Type u} {β : Type v} [BEq α] [Hashable α] {ρ : Type w} [ForIn Id ρ (α × β)] {m : Raw α β} {l : ρ} (h : m.WF) :
                                                                                              theorem Std.HashMap.Raw.WF.insertManyIfNewUnit {α : Type u} [BEq α] [Hashable α] {ρ : Type w} [ForIn Id ρ α] {m : Raw α Unit} {l : ρ} (h : m.WF) :
                                                                                              theorem Std.HashMap.Raw.WF.ofList {α : Type u} {β : Type v} [BEq α] [Hashable α] {l : List (α × β)} :
                                                                                              theorem Std.HashMap.Raw.WF.unitOfList {α : Type u} [BEq α] [Hashable α] {l : List α} :
                                                                                              theorem Std.HashMap.Raw.WF.union {α : Type u} {β : Type v} [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) :
                                                                                              (m₁.union m₂).WF