Skip to content

Commit fea5553

Browse files
wkrozowskiTwoFX
andauthored
feat: add ofArray to DHashMap/HashMap/HashSet (#11243)
This PR adds `ofArray` to `DHashMap`/`HashMap`/`HashSet` and proves a simp lemma allowing to rewrite `ofArray` to `ofList`. --------- Co-authored-by: Markus Himmel <[email protected]>
1 parent 70b4943 commit fea5553

File tree

12 files changed

+142
-2
lines changed

12 files changed

+142
-2
lines changed

src/Std/Data/DHashMap/Basic.lean

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -387,10 +387,18 @@ end Unverified
387387
DHashMap α β :=
388388
insertMany ∅ l
389389

390+
@[inline, inherit_doc Raw.ofArray] def ofArray [BEq α] [Hashable α] (l : Array ((a : α) × β a)) :
391+
DHashMap α β :=
392+
insertMany ∅ l
393+
390394
@[inline, inherit_doc Raw.Const.ofList] def Const.ofList {β : Type v} [BEq α] [Hashable α]
391395
(l : List (α × β)) : DHashMap α (fun _ => β) :=
392396
Const.insertMany ∅ l
393397

398+
@[inline, inherit_doc Raw.Const.ofArray] def Const.ofArray {β : Type v} [BEq α] [Hashable α]
399+
(l : Array (α × β)) : DHashMap α (fun _ => β) :=
400+
Const.insertMany ∅ l
401+
394402
@[inline, inherit_doc Raw.Const.unitOfList] def Const.unitOfList [BEq α] [Hashable α] (l : List α) :
395403
DHashMap α (fun _ => Unit) :=
396404
Const.insertManyIfNewUnit ∅ l

src/Std/Data/DHashMap/Internal/Raw.lean

Lines changed: 16 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -133,6 +133,11 @@ theorem ofList_eq [BEq α] [Hashable α] {l : List ((a : α) × β a)} :
133133
simp only [Raw.ofList, Raw.insertMany, (Raw.WF.empty).size_buckets_pos ∅, ↓reduceDIte]
134134
congr
135135

136+
theorem ofArray_eq [BEq α] [Hashable α] {a : Array ((a : α) × β a)} :
137+
Raw.ofArray a = Raw₀.insertMany Raw₀.emptyWithCapacity a := by
138+
simp only [Raw.ofArray, Raw.insertMany, (Raw.WF.empty).size_buckets_pos ∅, ↓reduceDIte]
139+
congr
140+
136141
theorem alter_eq [BEq α] [LawfulBEq α] [Hashable α] {m : Raw α β} (h : m.WF) {k : α} {f : Option (β k) → Option (β k)} :
137142
m.alter k f = Raw₀.alter ⟨m, h.size_buckets_pos⟩ k f := by
138143
simp [Raw.alter, h.size_buckets_pos]
@@ -162,6 +167,11 @@ theorem Const.ofList_eq [BEq α] [Hashable α] {l : List (α × β)} :
162167
simp only [Raw.Const.ofList, Raw.Const.insertMany, (Raw.WF.empty).size_buckets_pos ∅, ↓reduceDIte]
163168
congr
164169

170+
theorem Const.ofArray_eq [BEq α] [Hashable α] {a : Array (α × β)} :
171+
Raw.Const.ofArray a = Raw₀.Const.insertMany Raw₀.emptyWithCapacity a := by
172+
simp only [Raw.Const.ofArray, Raw.Const.insertMany, (Raw.WF.empty).size_buckets_pos ∅, ↓reduceDIte]
173+
congr
174+
165175
theorem Const.insertManyIfNewUnit_eq {ρ : Type w} [ForIn Id ρ α] [BEq α] [Hashable α]
166176
{m : Raw α (fun _ => Unit)} {l : ρ} (h : m.WF):
167177
Raw.Const.insertManyIfNewUnit m l = Raw₀.Const.insertManyIfNewUnit ⟨m, h.size_buckets_pos⟩ l := by
@@ -173,6 +183,12 @@ theorem Const.unitOfList_eq [BEq α] [Hashable α] {l : List α} :
173183
↓reduceDIte]
174184
congr
175185

186+
theorem Const.unitOfArray_eq [BEq α] [Hashable α] {a : Array α} :
187+
Raw.Const.unitOfArray a = Raw₀.Const.insertManyIfNewUnit Raw₀.emptyWithCapacity a := by
188+
simp only [Raw.Const.unitOfArray, Raw.Const.insertManyIfNewUnit, (Raw.WF.empty).size_buckets_pos ∅,
189+
↓reduceDIte]
190+
congr
191+
176192
theorem Const.get?_eq [BEq α] [Hashable α] {m : Raw α (fun _ => β)} (h : m.WF) {a : α} :
177193
Raw.Const.get? m a = Raw₀.Const.get? ⟨m, h.size_buckets_pos⟩ a := by
178194
simp [Raw.Const.get?, h.size_buckets_pos]

src/Std/Data/DHashMap/Internal/WF.lean

Lines changed: 12 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1124,6 +1124,18 @@ theorem insertMany_eq_insertListₘ_toListModel [BEq α] [Hashable α] (m m₂ :
11241124
simp only [List.foldl_cons, insertListₘ]
11251125
apply ih
11261126

1127+
theorem insertMany_array_eq_insertMany_toList [BEq α] [Hashable α] (m : Raw₀ α β) (a : Array ((k : α) × (β k))) :
1128+
insertMany m a = insertMany m a.toList := by
1129+
simp only [insertMany, bind_pure_comp, map_pure, bind_pure, ← Array.forIn_toList, forIn_pure_yield_eq_foldl, Array.foldl_toList, Id.run_pure]
1130+
1131+
theorem Const.insertMany_array_eq_insertMany_toList [BEq α] [Hashable α] {β : Type v} (m : Raw₀ α fun _ => β) (a : Array (α × β)) :
1132+
Const.insertMany m a = Const.insertMany m a.toList := by
1133+
simp only [insertMany, bind_pure_comp, map_pure, bind_pure, ← Array.forIn_toList, forIn_pure_yield_eq_foldl, Array.foldl_toList, Id.run_pure]
1134+
1135+
theorem Const.insertManyIfNewUnit_array_eq_insertManyIfNewUnit_toList [BEq α] [Hashable α] (m : Raw₀ α fun _ => Unit) (a : Array α) :
1136+
Const.insertManyIfNewUnit m a = Const.insertManyIfNewUnit m a.toList := by
1137+
simp only [insertManyIfNewUnit, bind_pure_comp, map_pure, bind_pure, ← Array.forIn_toList, forIn_pure_yield_eq_foldl, Array.foldl_toList, Id.run_pure]
1138+
11271139
theorem insertManyIfNew_eq_insertListIfNewₘ_toListModel [BEq α] [Hashable α] (m m₂ : Raw₀ α β) :
11281140
insertManyIfNew m m₂.1 = insertListIfNewₘ m (toListModel m₂.1.buckets) := by
11291141
simp only [insertManyIfNew, bind_pure_comp, map_pure, bind_pure]

src/Std/Data/DHashMap/Lemmas.lean

Lines changed: 15 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -2888,6 +2888,11 @@ end DHashMap
28882888

28892889
namespace DHashMap
28902890

2891+
@[simp, grind =]
2892+
theorem ofArray_eq_ofList (a : Array ((a : α) × (β a))) :
2893+
ofArray a = ofList a.toList :=
2894+
ext <| congrArg Subtype.val <| congrArg Subtype.val (Raw₀.insertMany_array_eq_insertMany_toList (α := α) _ a)
2895+
28912896
@[simp, grind =]
28922897
theorem ofList_nil :
28932898
ofList ([] : List ((a : α) × (β a))) = ∅ :=
@@ -3037,6 +3042,11 @@ namespace Const
30373042

30383043
variable {β : Type v}
30393044

3045+
@[simp, grind =]
3046+
theorem ofArray_eq_ofList (a : Array (α × β)) :
3047+
ofArray a = ofList a.toList :=
3048+
ext <| congrArg Subtype.val <| congrArg Subtype.val (Raw₀.Const.insertMany_array_eq_insertMany_toList (α := α) _ a)
3049+
30403050
@[simp, grind =]
30413051
theorem ofList_nil :
30423052
ofList ([] : List (α × β)) = ∅ :=
@@ -3182,6 +3192,11 @@ theorem isEmpty_ofList [EquivBEq α] [LawfulHashable α]
31823192
(ofList l).isEmpty = l.isEmpty :=
31833193
Raw₀.Const.isEmpty_insertMany_emptyWithCapacity_list
31843194

3195+
@[simp, grind =]
3196+
theorem unitOfArray_eq_unitOfList (a : Array α) :
3197+
unitOfArray a = unitOfList a.toList :=
3198+
ext <| congrArg Subtype.val <| congrArg Subtype.val (Raw₀.Const.insertManyIfNewUnit_array_eq_insertManyIfNewUnit_toList (α := α) _ a)
3199+
31853200
@[simp]
31863201
theorem unitOfList_nil :
31873202
unitOfList ([] : List α) = ∅ :=

src/Std/Data/DHashMap/Raw.lean

Lines changed: 21 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -607,10 +607,19 @@ occurrence takes precedence. -/
607607
@[inline] def ofList [BEq α] [Hashable α] (l : List ((a : α) × β a)) : Raw α β :=
608608
insertMany ∅ l
609609

610+
/-- Creates a hash map from an array of mappings. If the same key appears multiple times, the last
611+
occurrence takes precedence. -/
612+
@[inline] def ofArray [BEq α] [Hashable α] (l : Array ((a : α) × β a)) : Raw α β :=
613+
insertMany ∅ l
614+
610615
@[inline, inherit_doc Raw.ofList] def Const.ofList {β : Type v} [BEq α] [Hashable α]
611616
(l : List (α × β)) : Raw α (fun _ => β) :=
612617
Const.insertMany ∅ l
613618

619+
@[inline, inherit_doc Raw.ofArray] def Const.ofArray {β : Type v} [BEq α] [Hashable α]
620+
(l : Array (α × β)) : Raw α (fun _ => β) :=
621+
Const.insertMany ∅ l
622+
614623
/-- Creates a hash map from a list of keys, associating the value `()` with each key.
615624
616625
This is mainly useful to implement `HashSet.ofList`, so if you are considering using this,
@@ -760,6 +769,18 @@ theorem WF.Const.unitOfList [BEq α] [Hashable α] {l : List α} :
760769
(Const.unitOfList l : Raw α (fun _ => Unit)).WF :=
761770
Const.insertManyIfNewUnit WF.empty
762771

772+
theorem WF.ofArray [BEq α] [Hashable α] {a : Array ((a : α) × β a)} :
773+
(ofArray a : Raw α β).WF :=
774+
.insertMany WF.empty
775+
776+
theorem WF.Const.ofArray {β : Type v} [BEq α] [Hashable α] {a : Array (α × β)} :
777+
(Const.ofArray a : Raw α (fun _ => β)).WF :=
778+
Const.insertMany WF.empty
779+
780+
theorem WF.Const.unitOfArray [BEq α] [Hashable α] {a : Array α} :
781+
(Const.unitOfArray a : Raw α (fun _ => Unit)).WF :=
782+
Const.insertManyIfNewUnit WF.empty
783+
763784
theorem WF.union₀ [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (Raw₀.union ⟨m₁, h₁.size_buckets_pos⟩ ⟨m₂, h₂.size_buckets_pos⟩).val.WF := by
764785
simp only [Raw₀.union]
765786
split

src/Std/Data/DHashMap/RawLemmas.lean

Lines changed: 20 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -46,8 +46,8 @@ private meta def baseNames : Array Name :=
4646
``get?_eq, ``get_eq, ``get!_eq, ``getD_eq,
4747
``Const.get?_eq, ``Const.get_eq, ``Const.getD_eq, ``Const.get!_eq,
4848
``getKey?_eq, ``getKey_eq, ``getKey!_eq, ``getKeyD_eq,
49-
``insertMany_eq, ``Const.insertMany_eq, ``Const.insertManyIfNewUnit_eq,
50-
``ofList_eq, ``Const.ofList_eq, ``Const.unitOfList_eq,
49+
``insertMany_eq, ``Const.insertMany_eq, ``Const.insertManyIfNewUnit_eq, ``ofArray_eq, ``Const.ofArray_eq,
50+
``ofList_eq, ``Const.ofList_eq, ``Const.unitOfList_eq, ``Const.unitOfArray_eq,
5151
``alter_eq, ``Const.alter_eq, ``modify_eq, ``Const.modify_eq, ``union_eq, ``inter_eq]
5252

5353
/-- Internal implementation detail of the hash map -/
@@ -3340,6 +3340,12 @@ variable [BEq α] [Hashable α]
33403340

33413341
open Internal.Raw Internal.Raw₀
33423342

3343+
@[simp, grind =]
3344+
theorem ofArray_eq_ofList (a : Array ((a : α) × (β a))) :
3345+
ofArray a = ofList a.toList := by
3346+
simp_to_raw
3347+
rw [Raw₀.insertMany_array_eq_insertMany_toList]
3348+
33433349
@[simp, grind =]
33443350
theorem ofList_nil :
33453351
ofList ([] : List ((a : α) × (β a))) = ∅ := by
@@ -3492,6 +3498,12 @@ namespace Const
34923498

34933499
variable {β : Type v}
34943500

3501+
@[simp, grind =]
3502+
theorem ofArray_eq_ofList (a : Array (α × β)) :
3503+
ofArray a = ofList a.toList := by
3504+
simp_to_raw
3505+
rw [Const.insertMany_array_eq_insertMany_toList]
3506+
34953507
@[simp, grind =]
34963508
theorem ofList_nil :
34973509
ofList ([] : List (α × β)) = ∅ := by
@@ -3640,6 +3652,12 @@ theorem isEmpty_ofList [EquivBEq α] [LawfulHashable α]
36403652
(ofList l).isEmpty = l.isEmpty := by
36413653
simp_to_raw using Raw₀.Const.isEmpty_insertMany_emptyWithCapacity_list
36423654

3655+
@[simp, grind =]
3656+
theorem unitOfArray_eq_unitOfList (a : Array α) :
3657+
unitOfArray a = unitOfList a.toList := by
3658+
simp_to_raw
3659+
rw [Raw₀.Const.insertManyIfNewUnit_array_eq_insertManyIfNewUnit_toList]
3660+
36433661
@[simp]
36443662
theorem unitOfList_nil :
36453663
unitOfList ([] : List α) = ∅ := by

src/Std/Data/HashMap/Basic.lean

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -198,6 +198,10 @@ instance [BEq α] [Hashable α] : GetElem? (HashMap α β) α β (fun m a => a
198198
HashMap α Unit :=
199199
⟨DHashMap.Const.unitOfList l⟩
200200

201+
@[inline, inherit_doc DHashMap.Const.ofArray] def ofArray [BEq α] [Hashable α] (a : Array (α × β)) :
202+
HashMap α β :=
203+
⟨DHashMap.Const.ofArray a⟩
204+
201205
@[inline, inherit_doc DHashMap.Const.toList] def toList (m : HashMap α β) :
202206
List (α × β) :=
203207
DHashMap.Const.toList m.inner

src/Std/Data/HashMap/Lemmas.lean

Lines changed: 12 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -2105,6 +2105,12 @@ theorem size_ofList_le [EquivBEq α] [LawfulHashable α]
21052105

21062106
grind_pattern size_ofList_le => (ofList l).size
21072107

2108+
@[simp, grind =]
2109+
theorem ofArray_eq_ofList (a : Array (α × β)) :
2110+
ofArray a = ofList a.toList := by
2111+
apply ext
2112+
apply DHashMap.Const.ofArray_eq_ofList
2113+
21082114
@[simp, grind =]
21092115
theorem isEmpty_ofList [EquivBEq α] [LawfulHashable α]
21102116
{l : List (α × β)} :
@@ -2222,6 +2228,12 @@ theorem size_unitOfList_le [EquivBEq α] [LawfulHashable α]
22222228
(unitOfList l).size ≤ l.length :=
22232229
DHashMap.Const.size_unitOfList_le
22242230

2231+
@[simp, grind =]
2232+
theorem unitOfArray_eq_unitOfList (a : Array α) :
2233+
unitOfArray a = unitOfList a.toList := by
2234+
apply ext
2235+
apply DHashMap.Const.unitOfArray_eq_unitOfList
2236+
22252237
@[simp]
22262238
theorem isEmpty_unitOfList [EquivBEq α] [LawfulHashable α]
22272239
{l : List α} :

src/Std/Data/HashMap/Raw.lean

Lines changed: 10 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -192,6 +192,10 @@ instance [BEq α] [Hashable α] : GetElem? (Raw α β) α β (fun m a => a ∈ m
192192
(l : List α) : Raw α Unit :=
193193
⟨DHashMap.Raw.Const.unitOfList l⟩
194194

195+
@[inline, inherit_doc DHashMap.Raw.Const.ofList] def ofArray [BEq α] [Hashable α]
196+
(a : Array (α × β)) : Raw α β :=
197+
⟨DHashMap.Raw.Const.ofArray a⟩
198+
195199
@[inline, inherit_doc DHashMap.Raw.Const.alter] def alter [BEq α] [EquivBEq α] [Hashable α]
196200
(m : Raw α β) (a : α) (f : Option β → Option β) : Raw α β :=
197201
⟨DHashMap.Raw.Const.alter m.inner a f⟩
@@ -342,9 +346,15 @@ theorem WF.insertManyIfNewUnit [BEq α] [Hashable α] {ρ : Type w} [ForIn Id ρ
342346
theorem WF.ofList [BEq α] [Hashable α] {l : List (α × β)} : (ofList l).WF :=
343347
⟨DHashMap.Raw.WF.Const.ofList⟩
344348

349+
theorem WF.ofArray [BEq α] [Hashable α] {a : Array (α × β)} : (ofArray a).WF :=
350+
⟨DHashMap.Raw.WF.Const.ofArray⟩
351+
345352
theorem WF.unitOfList [BEq α] [Hashable α] {l : List α} : (unitOfList l).WF :=
346353
⟨DHashMap.Raw.WF.Const.unitOfList⟩
347354

355+
theorem WF.unitOfArray [BEq α] [Hashable α] {a : Array α} : (unitOfArray a).WF :=
356+
⟨DHashMap.Raw.WF.Const.unitOfArray⟩
357+
348358
theorem WF.union [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁ ∪ m₂).WF :=
349359
⟨DHashMap.Raw.WF.union h₁.out h₂.out⟩
350360

src/Std/Data/HashMap/RawLemmas.lean

Lines changed: 12 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -2156,6 +2156,12 @@ theorem size_ofList_le [EquivBEq α] [LawfulHashable α]
21562156

21572157
grind_pattern size_ofList_le => (ofList l).size
21582158

2159+
@[simp, grind =]
2160+
theorem ofArray_eq_ofList (a : Array (α × β)) :
2161+
ofArray a = ofList a.toList := by
2162+
apply ext
2163+
apply DHashMap.Raw.Const.ofArray_eq_ofList
2164+
21592165
@[simp, grind =]
21602166
theorem isEmpty_ofList [EquivBEq α] [LawfulHashable α]
21612167
{l : List (α × β)} :
@@ -2244,6 +2250,12 @@ theorem size_unitOfList_le [EquivBEq α] [LawfulHashable α]
22442250
(unitOfList l).size ≤ l.length :=
22452251
DHashMap.Raw.Const.size_unitOfList_le
22462252

2253+
@[simp, grind =]
2254+
theorem unitOfArray_eq_unitOfList (a : Array α) :
2255+
unitOfArray a = unitOfList a.toList := by
2256+
apply ext
2257+
apply DHashMap.Raw.Const.unitOfArray_eq_unitOfList
2258+
22472259
@[simp]
22482260
theorem isEmpty_unitOfList [EquivBEq α] [LawfulHashable α]
22492261
{l : List α} :

0 commit comments

Comments
 (0)