forked from leanprover/lean4
-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy pathSMap.lean
106 lines (80 loc) · 3.95 KB
/
SMap.lean
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
/-
Copyright (c) 2019 Microsoft Corporation. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Leonardo de Moura
-/
import Lean.Data.HashMap
import Lean.Data.PersistentHashMap
universe u v w w'
namespace Lean
/-- Staged map for implementing the Environment. The idea is to store
imported entries into a hashtable and local entries into a persistent hashtable.
Hypotheses:
- The number of entries (i.e., declarations) coming from imported files is much bigger than
the number of entries in the current file.
- HashMap is faster than PersistentHashMap.
- When we are reading imported files, we have exclusive access to the map, and efficient
destructive updates are performed.
Remarks:
- We never remove declarations from the Environment. In principle, we could support
deletion by using `(PHashMap α (Option β))` where the value `none` would indicate
that an entry was "removed" from the hashtable.
- We do not need additional bookkeeping for extracting the local entries.
-/
structure SMap (α : Type u) (β : Type v) [BEq α] [Hashable α] where
stage₁ : Bool := true
map₁ : HashMap α β := {}
map₂ : PHashMap α β := {}
namespace SMap
variable {α : Type u} {β : Type v} [BEq α] [Hashable α]
instance : Inhabited (SMap α β) := ⟨{}⟩
def empty : SMap α β := {}
@[inline] def fromHashMap (m : HashMap α β) (stage₁ := true) : SMap α β :=
{ map₁ := m, stage₁ := stage₁ }
@[specialize] def insert : SMap α β → α → β → SMap α β
| ⟨true, m₁, m₂⟩, k, v => ⟨true, m₁.insert k v, m₂⟩
| ⟨false, m₁, m₂⟩, k, v => ⟨false, m₁, m₂.insert k v⟩
@[specialize] def insert' : SMap α β → α → β → SMap α β
| ⟨true, m₁, m₂⟩, k, v => ⟨true, m₁.insert k v, m₂⟩
| ⟨false, m₁, m₂⟩, k, v => ⟨false, m₁, m₂.insert k v⟩
@[specialize] def find? : SMap α β → α → Option β
| ⟨true, m₁, _⟩, k => m₁.find? k
| ⟨false, m₁, m₂⟩, k => (m₂.find? k).orElse fun _ => m₁.find? k
@[inline] def findD (m : SMap α β) (a : α) (b₀ : β) : β :=
(m.find? a).getD b₀
@[inline] def find! [Inhabited β] (m : SMap α β) (a : α) : β :=
match m.find? a with
| some b => b
| none => panic! "key is not in the map"
@[specialize] def contains : SMap α β → α → Bool
| ⟨true, m₁, _⟩, k => m₁.contains k
| ⟨false, m₁, m₂⟩, k => m₁.contains k || m₂.contains k
/-- Similar to `find?`, but searches for result in the hashmap first.
So, the result is correct only if we never "overwrite" `map₁` entries using `map₂`. -/
@[specialize] def find?' : SMap α β → α → Option β
| ⟨true, m₁, _⟩, k => m₁.find? k
| ⟨false, m₁, m₂⟩, k => (m₁.find? k).orElse fun _ => m₂.find? k
def forM [Monad m] (s : SMap α β) (f : α → β → m PUnit) : m PUnit := do
s.map₁.forM f
s.map₂.forM f
/-- Move from stage 1 into stage 2. -/
def switch (m : SMap α β) : SMap α β :=
if m.stage₁ then { m with stage₁ := false } else m
@[inline] def foldStage2 {σ : Type w} (f : σ → α → β → σ) (s : σ) (m : SMap α β) : σ :=
m.map₂.foldl f s
def fold {σ : Type w} (f : σ → α → β → σ) (init : σ) (m : SMap α β) : σ :=
m.map₂.foldl f $ m.map₁.fold f init
def size (m : SMap α β) : Nat :=
m.map₁.size + m.map₂.size
def stageSizes (m : SMap α β) : Nat × Nat :=
(m.map₁.size, m.map₂.size)
def numBuckets (m : SMap α β) : Nat :=
m.map₁.numBuckets
def toList (m : SMap α β) : List (α × β) :=
m.fold (init := []) fun es a b => (a, b)::es
end SMap
def _root_.List.toSMap [BEq α] [Hashable α] (es : List (α × β)) : SMap α β :=
es.foldl (init := {}) fun s (a, b) => s.insert a b
instance {_ : BEq α} {_ : Hashable α} [Repr α] [Repr β] : Repr (SMap α β) where
reprPrec v prec := Repr.addAppParen (reprArg v.toList ++ ".toSMap") prec
end Lean