The bucket array of a HashMap
is a nonempty array of AssocList
s.
(This type is an internal implementation detail of HashMap
.)
Equations
- Batteries.HashMap.Imp.Buckets α β = { b : Array (Batteries.AssocList α β) // 0 < b.size }
Instances For
Construct a new empty bucket array with the specified capacity.
Equations
- Batteries.HashMap.Imp.Buckets.mk buckets h = ⟨mkArray buckets Batteries.AssocList.nil, ⋯⟩
Instances For
Update one bucket in the bucket array with a new value.
Equations
- data.update i d h = ⟨data.val.uset i d h, ⋯⟩
Instances For
The number of elements in the bucket array.
Note: this is marked noncomputable
because it is only intended for specification.
Equations
- data.size = Nat.sum (List.map (fun (x : Batteries.AssocList α β) => x.toList.length) data.val.toList)
Instances For
Map a function over the values in the map.
Equations
- Batteries.HashMap.Imp.Buckets.mapVal f self = ⟨Array.map (Batteries.AssocList.mapVal f) self.val, ⋯⟩
Instances For
The well-formedness invariant for the bucket array says that every element hashes to its index (assuming the hash is lawful - otherwise there are no promises about where elements are located).
- distinct : ∀ [inst : Batteries.HashMap.LawfulHashable α] [inst : PartialEquivBEq α] (bucket : Batteries.AssocList α β), bucket ∈ buckets.val.toList → List.Pairwise (fun (a b : α × β) => ¬(a.fst == b.fst) = true) bucket.toList
The elements of a bucket are all distinct according to the
BEq
relation. - hash_self : ∀ (i : Nat) (h : i < buckets.val.size), Batteries.AssocList.All (fun (k : α) (x : β) => ((hash k).toUSize % buckets.val.size).toNat = i) buckets.val[i]
Every element in a bucket should hash to its location.
Instances For
The elements of a bucket are all distinct according to the BEq
relation.
Every element in a bucket should hash to its location.
HashMap.Imp α β
is the internal implementation type of HashMap α β
.
- size : Nat
- buckets : Batteries.HashMap.Imp.Buckets α β
The bucket array of the
HashMap
.
Instances For
Given a desired capacity, this returns the number of buckets we should reserve.
A "load factor" of 0.75 is the usual standard for hash maps, so we return capacity * 4 / 3
.
Equations
- Batteries.HashMap.Imp.numBucketsForCapacity capacity = capacity * 4 / 3
Instances For
Constructs an empty hash map with the specified nonzero number of buckets.
Equations
- Batteries.HashMap.Imp.empty' buckets h = { size := 0, buckets := Batteries.HashMap.Imp.Buckets.mk buckets h }
Instances For
Constructs an empty hash map with the specified target capacity.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Calculates the bucket index from a hash value u
.
Equations
- Batteries.HashMap.Imp.mkIdx h u = ⟨u % n, ⋯⟩
Instances For
Inserts a key-value pair into the bucket array. This function assumes that the data is not already in the array, which is appropriate when reinserting elements into the array after a resize.
Equations
- Batteries.HashMap.Imp.reinsertAux data a b = match Batteries.HashMap.Imp.mkIdx ⋯ (hash a).toUSize with | ⟨i, h⟩ => data.update i (Batteries.AssocList.cons a b data.val[i]) h
Instances For
Folds a monadic function over the elements in the map (in arbitrary order).
Equations
- Batteries.HashMap.Imp.foldM f d map = Array.foldlM (fun (d : δ) (b : Batteries.AssocList α β) => Batteries.AssocList.foldlM f d b) d map.buckets.val
Instances For
Folds a function over the elements in the map (in arbitrary order).
Equations
- Batteries.HashMap.Imp.fold f d m = (Batteries.HashMap.Imp.foldM f d m).run
Instances For
Runs a monadic function over the elements in the map (in arbitrary order).
Equations
- Batteries.HashMap.Imp.forM f h = Array.forM (fun (b : Batteries.AssocList α β) => Batteries.AssocList.forM f b) h.buckets.val
Instances For
Given a key a
, returns a key-value pair in the map whose key compares equal to a
.
Equations
- { size := size, buckets := buckets }.findEntry? a = match Batteries.HashMap.Imp.mkIdx ⋯ (hash a).toUSize with | ⟨i, h⟩ => Batteries.AssocList.findEntry? a buckets.val[i]
Instances For
Looks up an element in the map with key a
.
Equations
- { size := size, buckets := buckets }.find? a = match Batteries.HashMap.Imp.mkIdx ⋯ (hash a).toUSize with | ⟨i, h⟩ => Batteries.AssocList.find? a buckets.val[i]
Instances For
Returns true if the element a
is in the map.
Equations
- { size := size, buckets := buckets }.contains a = match Batteries.HashMap.Imp.mkIdx ⋯ (hash a).toUSize with | ⟨i, h⟩ => Batteries.AssocList.contains a buckets.val[i]
Instances For
Copies all the entries from buckets
into a new hash map with a larger capacity.
Equations
- Batteries.HashMap.Imp.expand size buckets = { size := size, buckets := Batteries.HashMap.Imp.expand.go 0 buckets.val (Batteries.HashMap.Imp.Buckets.mk (buckets.val.size * 2) ⋯) }
Instances For
Inner loop of expand
. Copies elements source[i:]
into target
,
destroying source
in the process.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Inserts key-value pair a, b
into the map.
If an element equal to a
is already in the map, it is replaced by b
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Removes key a
from the map. If it does not exist in the map, the map is returned unchanged.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Map a function over the values in the map.
Equations
- Batteries.HashMap.Imp.mapVal f self = { size := self.size, buckets := Batteries.HashMap.Imp.Buckets.mapVal f self.buckets }
Instances For
Performs an in-place edit of the value, ensuring that the value is used linearly.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Applies f
to each key-value pair a, b
in the map. If it returns some c
then
a, c
is pushed into the new map; else the key is removed from the map.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Inner loop of filterMap
. Note that this reverses the bucket lists,
but this is fine since bucket lists are unordered.
Equations
- One or more equations did not get rendered due to their size.
- Batteries.HashMap.Imp.filterMap.go f acc Batteries.AssocList.nil x = (acc, x)
Instances For
Constructs a map with the set of all pairs a, b
such that f
returns true.
Equations
- Batteries.HashMap.Imp.filter f m = Batteries.HashMap.Imp.filterMap (fun (a : α) (b : β) => bif f a b then some b else none) m
Instances For
The well-formedness invariant for a hash map. The first constructor is the real invariant,
and the others allow us to "cheat" in this file and define insert
and erase
,
which have more complex proofs that are delayed to Batteries.Data.HashMap.Lemmas
.
- mk: ∀ {α : Type u_1} [inst : BEq α] [inst_1 : Hashable α] {x : Type u_2} {m : Batteries.HashMap.Imp α x},
m.size = m.buckets.size → m.buckets.WF → m.WF
The real well-formedness invariant:
- The
size
field should match the actual number of elements in the map - The bucket array should be well-formed, meaning that if the hashable instance is lawful then every element hashes to its index.
- The
- empty': ∀ {α : Type u_1} [inst : BEq α] [inst_1 : Hashable α] {x : Type u_2} {n : Nat} {h : 0 < n},
(Batteries.HashMap.Imp.empty' n h).WF
The empty hash map is well formed.
- insert: ∀ {α : Type u_1} [inst : BEq α] [inst_1 : Hashable α] {x : Type u_2} {m : Batteries.HashMap.Imp α x} {a : α} {b : x},
m.WF → (m.insert a b).WF
Inserting into a well formed hash map yields a well formed hash map.
- erase: ∀ {α : Type u_1} [inst : BEq α] [inst_1 : Hashable α] {x : Type u_2} {m : Batteries.HashMap.Imp α x} {a : α},
m.WF → (m.erase a).WF
Removing an element from a well formed hash map yields a well formed hash map.
- modify: ∀ {α : Type u_1} [inst : BEq α] [inst_1 : Hashable α] {x : Type u_2} {m : Batteries.HashMap.Imp α x} {a : α}
{f : α → x → x}, m.WF → (m.modify a f).WF
Replacing an element in a well formed hash map yields a well formed hash map.
Instances For
HashMap α β
is a key-value map which stores elements in an array using a hash function
to find the values. This allows it to have very good performance for lookups
(average O(1)
for a perfectly random hash function), but it is not a persistent data structure,
meaning that one should take care to use the map linearly when performing updates.
Copies are O(n)
.
Equations
- Batteries.HashMap α β = { m : Batteries.HashMap.Imp α β // m.WF }
Instances For
Make a new hash map with the specified capacity.
Equations
- Batteries.mkHashMap capacity = ⟨Batteries.HashMap.Imp.empty capacity, ⋯⟩
Instances For
Equations
- Batteries.HashMap.instInhabited = { default := Batteries.mkHashMap }
Make a new empty hash map.
(empty : Batteries.HashMap Int Int).toList = []
Equations
- Batteries.HashMap.empty = Batteries.mkHashMap
Instances For
The number of elements in the hash map.
(ofList [("one", 1), ("two", 2)]).size = 2
Equations
- self.size = self.val.size
Instances For
Is the map empty?
(empty : Batteries.HashMap Int Int).isEmpty = true
(ofList [("one", 1), ("two", 2)]).isEmpty = false
Instances For
Inserts key-value pair a, b
into the map.
If an element equal to a
is already in the map, it is replaced by b
.
def hashMap := ofList [("one", 1), ("two", 2)]
hashMap.insert "three" 3 = {"one" => 1, "two" => 2, "three" => 3}
hashMap.insert "two" 0 = {"one" => 1, "two" => 0}
Equations
- self.insert a b = ⟨self.val.insert a b, ⋯⟩
Instances For
Similar to insert
, but also returns a boolean flag indicating whether an existing entry has been
replaced with a => b
.
def hashMap := ofList [("one", 1), ("two", 2)]
hashMap.insert' "three" 3 = ({"one" => 1, "two" => 2, "three" => 3}, false)
hashMap.insert' "two" 0 = ({"one" => 1, "two" => 0}, true)
Instances For
Removes key a
from the map. If it does not exist in the map, the map is returned unchanged.
def hashMap := ofList [("one", 1), ("two", 2)]
hashMap.erase "one" = {"two" => 2}
hashMap.erase "three" = {"one" => 1, "two" => 2}
Equations
- self.erase a = ⟨self.val.erase a, ⋯⟩
Instances For
Performs an in-place edit of the value, ensuring that the value is used linearly.
The function f
is passed the original key of the entry, along with the value in the map.
(ofList [("one", 1), ("two", 2)]).modify "one" (fun _ v => v + 1) = {"one" => 2, "two" => 2}
(ofList [("one", 1), ("two", 2)]).modify "three" (fun _ v => v + 1) = {"one" => 1, "two" => 2}
Equations
- self.modify a f = ⟨self.val.modify a f, ⋯⟩
Instances For
Given a key a
, returns a key-value pair in the map whose key compares equal to a
.
Note that the returned key may not be identical to the input, if ==
ignores some part
of the value.
def hashMap := ofList [("one", 1), ("two", 2)]
hashMap.findEntry? "one" = some ("one", 1)
hashMap.findEntry? "three" = none
Equations
- self.findEntry? a = self.val.findEntry? a
Instances For
Looks up an element in the map with key a
.
def hashMap := ofList [("one", 1), ("two", 2)]
hashMap.find? "one" = some 1
hashMap.find? "three" = none
Equations
- self.find? a = self.val.find? a
Instances For
Looks up an element in the map with key a
. Returns b₀
if the element is not found.
def hashMap := ofList [("one", 1), ("two", 2)]
hashMap.findD "one" 0 = 1
hashMap.findD "three" 0 = 0
Equations
- self.findD a b₀ = (self.find? a).getD b₀
Instances For
Looks up an element in the map with key a
. Panics if the element is not found.
def hashMap := ofList [("one", 1), ("two", 2)]
hashMap.find! "one" = 1
hashMap.find! "three" => panic!
Equations
- self.find! a = (self.find? a).getD (panicWithPosWithDecl "Batteries.Data.HashMap.Basic" "Batteries.HashMap.find!" 379 23 "key is not in the map")
Instances For
Equations
- Batteries.HashMap.instGetElemOptionTrue = { getElem := fun (m : Batteries.HashMap α β) (k : α) (x_1 : True) => m.find? k }
Returns true if the element a
is in the map.
def hashMap := ofList [("one", 1), ("two", 2)]
hashMap.contains "one" = true
hashMap.contains "three" = false
Equations
- self.contains a = self.val.contains a
Instances For
Folds a monadic function over the elements in the map (in arbitrary order).
def sumEven (sum: Nat) (k : String) (v : Nat) : Except String Nat :=
if v % 2 == 0 then pure (sum + v) else throw s!"value {v} at key {k} is not even"
foldM sumEven 0 (ofList [("one", 1), ("three", 3)]) =
Except.error "value 3 at key three is not even"
foldM sumEven 0 (ofList [("two", 2), ("four", 4)]) = Except.ok 6
Equations
- Batteries.HashMap.foldM f init self = Batteries.HashMap.Imp.foldM f init self.val
Instances For
Folds a function over the elements in the map (in arbitrary order).
fold (fun sum _ v => sum + v) 0 (ofList [("one", 1), ("two", 2)]) = 3
Equations
- Batteries.HashMap.fold f init self = Batteries.HashMap.Imp.fold f init self.val
Instances For
Combines two hashmaps using a monadic function f
to combine two values at a key.
def map1 := ofList [("one", 1), ("two", 2)]
def map2 := ofList [("two", 2), ("three", 3)]
def map3 := ofList [("two", 3), ("three", 3)]
def mergeIfNoConflict? (_ : String) (v₁ v₂ : Nat) : Option Nat :=
if v₁ != v₂ then none else some v₁
mergeWithM mergeIfNoConflict? map1 map2 = some {"one" => 1, "two" => 2, "three" => 3}
mergeWithM mergeIfNoConflict? map1 map3 = none
Equations
- One or more equations did not get rendered due to their size.
Instances For
Combines two hashmaps using function f
to combine two values at a key.
mergeWith (fun _ v₁ v₂ => v₁ + v₂ )
(ofList [("one", 1), ("two", 2)]) (ofList [("two", 2), ("three", 3)]) =
{"one" => 1, "two" => 4, "three" => 3}
Equations
- One or more equations did not get rendered due to their size.
Instances For
Runs a monadic function over the elements in the map (in arbitrary order).
def checkEven (k : String) (v : Nat) : Except String Unit :=
if v % 2 == 0 then pure () else throw s!"value {v} at key {k} is not even"
forM checkEven (ofList [("one", 1), ("three", 3)]) = Except.error "value 3 at key three is not even"
forM checkEven (ofList [("two", 2), ("four", 4)]) = Except.ok ()
Equations
- Batteries.HashMap.forM f self = Batteries.HashMap.Imp.forM f self.val
Instances For
Converts the map into a list of key-value pairs.
open List
(ofList [("one", 1), ("two", 2)]).toList ~ [("one", 1), ("two", 2)]
Equations
- self.toList = Batteries.HashMap.fold (fun (r : List (α × β)) (k : α) (v : β) => (k, v) :: r) [] self
Instances For
Converts the map into an array of key-value pairs.
open List
(ofList [("one", 1), ("two", 2)]).toArray.data ~ #[("one", 1), ("two", 2)].data
Equations
- self.toArray = Batteries.HashMap.fold (fun (r : Array (α × β)) (k : α) (v : β) => r.push (k, v)) #[] self
Instances For
The number of buckets in the hash map.
Equations
- self.numBuckets = self.val.buckets.val.size
Instances For
Builds a HashMap
from a list of key-value pairs.
Values of duplicated keys are replaced by their respective last occurrences.
ofList [("one", 1), ("one", 2)] = {"one" => 2}
Equations
- Batteries.HashMap.ofList l = List.foldl (fun (m : Batteries.HashMap α β) (x : α × β) => match x with | (k, v) => m.insert k v) Batteries.HashMap.empty l
Instances For
Variant of ofList
which accepts a function that combines values of duplicated keys.
ofListWith [("one", 1), ("one", 2)] (fun v₁ v₂ => v₁ + v₂) = {"one" => 3}
Equations
- One or more equations did not get rendered due to their size.