Skip to content

Commit 6665837

Browse files
feat: verify insertMany method for adding lists to HashMaps (#6211)
This PR verifies the `insertMany` method on `HashMap`s for the special case of inserting lists. --------- Co-authored-by: jt0202 <johannes.tantow@gmail.com> Co-authored-by: monsterkrampe <monsterkrampe@users.noreply.github.com> Co-authored-by: Johannes Tantow <44068763+jt0202@users.noreply.github.com>
1 parent c7fd873 commit 6665837

File tree

17 files changed

+5423
-86
lines changed

17 files changed

+5423
-86
lines changed

src/Std/Data/DHashMap/Basic.lean

Lines changed: 12 additions & 12 deletions
Original file line numberDiff line numberDiff line change
@@ -290,24 +290,12 @@ This function ensures that the value is used linearly.
290290
⟨(Raw₀.Const.insertManyIfNewUnit ⟨m.1, m.2.size_buckets_pos⟩ l).1,
291291
(Raw₀.Const.insertManyIfNewUnit ⟨m.1, m.2.size_buckets_pos⟩ l).2 _ Raw.WF.insertIfNew₀ m.2
292292

293-
@[inline, inherit_doc Raw.ofList] def ofList [BEq α] [Hashable α] (l : List ((a : α) × β a)) :
294-
DHashMap α β :=
295-
insertMany ∅ l
296-
297293
/-- Computes the union of the given hash maps, by traversing `m₂` and inserting its elements into `m₁`. -/
298294
@[inline] def union [BEq α] [Hashable α] (m₁ m₂ : DHashMap α β) : DHashMap α β :=
299295
m₂.fold (init := m₁) fun acc x => acc.insert x
300296

301297
instance [BEq α] [Hashable α] : Union (DHashMap α β) := ⟨union⟩
302298

303-
@[inline, inherit_doc Raw.Const.ofList] def Const.ofList {β : Type v} [BEq α] [Hashable α]
304-
(l : List (α × β)) : DHashMap α (fun _ => β) :=
305-
Const.insertMany ∅ l
306-
307-
@[inline, inherit_doc Raw.Const.unitOfList] def Const.unitOfList [BEq α] [Hashable α] (l : List α) :
308-
DHashMap α (fun _ => Unit) :=
309-
Const.insertManyIfNewUnit ∅ l
310-
311299
@[inline, inherit_doc Raw.Const.unitOfArray] def Const.unitOfArray [BEq α] [Hashable α] (l : Array α) :
312300
DHashMap α (fun _ => Unit) :=
313301
Const.insertManyIfNewUnit ∅ l
@@ -321,4 +309,16 @@ instance [BEq α] [Hashable α] [Repr α] [(a : α) → Repr (β a)] : Repr (DHa
321309

322310
end Unverified
323311

312+
@[inline, inherit_doc Raw.ofList] def ofList [BEq α] [Hashable α] (l : List ((a : α) × β a)) :
313+
DHashMap α β :=
314+
insertMany ∅ l
315+
316+
@[inline, inherit_doc Raw.Const.ofList] def Const.ofList {β : Type v} [BEq α] [Hashable α]
317+
(l : List (α × β)) : DHashMap α (fun _ => β) :=
318+
Const.insertMany ∅ l
319+
320+
@[inline, inherit_doc Raw.Const.unitOfList] def Const.unitOfList [BEq α] [Hashable α] (l : List α) :
321+
DHashMap α (fun _ => Unit) :=
322+
Const.insertManyIfNewUnit ∅ l
323+
324324
end Std.DHashMap

src/Std/Data/DHashMap/Internal/List/Associative.lean

Lines changed: 873 additions & 2 deletions
Large diffs are not rendered by default.

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

Lines changed: 64 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -201,7 +201,7 @@ theorem toListModel_updateAllBuckets {m : Raw₀ α β} {f : AssocList α β →
201201
omega
202202
rw [updateAllBuckets, toListModel, Array.toList_map, List.flatMap_eq_foldl, List.foldl_map,
203203
toListModel, List.flatMap_eq_foldl]
204-
suffices ∀ (l : List (AssocList α β)) (l' : List ((a: α) × δ a)) (l'' : List ((a : α) × β a)),
204+
suffices ∀ (l : List (AssocList α β)) (l' : List ((a : α) × δ a)) (l'' : List ((a : α) × β a)),
205205
Perm (g l'') l' →
206206
Perm (l.foldl (fun acc a => acc ++ (f a).toList) l')
207207
(g (l.foldl (fun acc a => acc ++ a.toList) l'')) by
@@ -378,6 +378,12 @@ def mapₘ (m : Raw₀ α β) (f : (a : α) → β a → δ a) : Raw₀ α δ :=
378378
def filterₘ (m : Raw₀ α β) (f : (a : α) → β a → Bool) : Raw₀ α β :=
379379
⟨withComputedSize (updateAllBuckets m.1.buckets fun l => l.filter f), by simpa using m.2
380380

381+
/-- Internal implementation detail of the hash map -/
382+
def insertListₘ [BEq α] [Hashable α] (m : Raw₀ α β) (l : List ((a : α) × β a)) : Raw₀ α β :=
383+
match l with
384+
| .nil => m
385+
| .cons hd tl => insertListₘ (m.insert hd.1 hd.2) tl
386+
381387
section
382388

383389
variable {β : Type v}
@@ -398,6 +404,20 @@ def Const.getDₘ [BEq α] [Hashable α] (m : Raw₀ α (fun _ => β)) (a : α)
398404
def Const.get!ₘ [BEq α] [Hashable α] [Inhabited β] (m : Raw₀ α (fun _ => β)) (a : α) : β :=
399405
(Const.get?ₘ m a).get!
400406

407+
/-- Internal implementation detail of the hash map -/
408+
def Const.insertListₘ [BEq α] [Hashable α] (m : Raw₀ α (fun _ => β)) (l: List (α × β)) :
409+
Raw₀ α (fun _ => β) :=
410+
match l with
411+
| .nil => m
412+
| .cons hd tl => insertListₘ (m.insert hd.1 hd.2) tl
413+
414+
/-- Internal implementation detail of the hash map -/
415+
def Const.insertListIfNewUnitₘ [BEq α] [Hashable α] (m : Raw₀ α (fun _ => Unit)) (l: List α) :
416+
Raw₀ α (fun _ => Unit) :=
417+
match l with
418+
| .nil => m
419+
| .cons hd tl => insertListIfNewUnitₘ (m.insertIfNew hd ()) tl
420+
401421
end
402422

403423
/-! # Equivalence between model functions and real implementations -/
@@ -569,6 +589,19 @@ theorem map_eq_mapₘ (m : Raw₀ α β) (f : (a : α) → β a → δ a) :
569589
theorem filter_eq_filterₘ (m : Raw₀ α β) (f : (a : α) → β a → Bool) :
570590
m.filter f = m.filterₘ f := rfl
571591

592+
theorem insertMany_eq_insertListₘ [BEq α] [Hashable α](m : Raw₀ α β) (l : List ((a : α) × β a)) : insertMany m l = insertListₘ m l := by
593+
simp only [insertMany, Id.run, Id.pure_eq, Id.bind_eq, List.forIn_yield_eq_foldl]
594+
suffices ∀ (t : { m' // ∀ (P : Raw₀ α β → Prop),
595+
(∀ {m'' : Raw₀ α β} {a : α} {b : β a}, P m'' → P (m''.insert a b)) → P m → P m' }),
596+
(List.foldl (fun m' p => ⟨m'.val.insert p.1 p.2, fun P h₁ h₂ => h₁ (m'.2 _ h₁ h₂)⟩) t l).val =
597+
t.val.insertListₘ l from this _
598+
intro t
599+
induction l generalizing m with
600+
| nil => simp [insertListₘ]
601+
| cons hd tl ih =>
602+
simp only [List.foldl_cons, insertListₘ]
603+
apply ih
604+
572605
section
573606

574607
variable {β : Type v}
@@ -599,6 +632,36 @@ theorem Const.getThenInsertIfNew?_eq_get?ₘ [BEq α] [Hashable α] (m : Raw₀
599632
dsimp only [Array.ugetElem_eq_getElem, Array.uset]
600633
split <;> simp_all [-getValue?_eq_none]
601634

635+
theorem Const.insertMany_eq_insertListₘ [BEq α] [Hashable α] (m : Raw₀ α (fun _ => β))
636+
(l: List (α × β)):
637+
(Const.insertMany m l).1 = Const.insertListₘ m l := by
638+
simp only [insertMany, Id.run, Id.pure_eq, Id.bind_eq, List.forIn_yield_eq_foldl]
639+
suffices ∀ (t : { m' // ∀ (P : Raw₀ α (fun _ => β) → Prop),
640+
(∀ {m'' : Raw₀ α (fun _ => β)} {a : α} {b : β}, P m'' → P (m''.insert a b)) → P m → P m' }),
641+
(List.foldl (fun m' p => ⟨m'.val.insert p.1 p.2, fun P h₁ h₂ => h₁ (m'.2 _ h₁ h₂)⟩) t l).val =
642+
Const.insertListₘ t.val l from this _
643+
intro t
644+
induction l generalizing m with
645+
| nil => simp [insertListₘ]
646+
| cons hd tl ih =>
647+
simp only [List.foldl_cons, insertListₘ]
648+
apply ih
649+
650+
theorem Const.insertManyIfNewUnit_eq_insertListIfNewUnitₘ [BEq α] [Hashable α]
651+
(m : Raw₀ α (fun _ => Unit)) (l: List α):
652+
(Const.insertManyIfNewUnit m l).1 = Const.insertListIfNewUnitₘ m l := by
653+
simp only [insertManyIfNewUnit, Id.run, Id.pure_eq, Id.bind_eq, List.forIn_yield_eq_foldl]
654+
suffices ∀ (t : { m' // ∀ (P : Raw₀ α (fun _ => Unit) → Prop),
655+
(∀ {m'' a b}, P m'' → P (m''.insertIfNew a b)) → P m → P m'}),
656+
(List.foldl (fun m' p => ⟨m'.val.insertIfNew p (), fun P h₁ h₂ => h₁ (m'.2 _ h₁ h₂)⟩) t l).val =
657+
Const.insertListIfNewUnitₘ t.val l from this _
658+
intro t
659+
induction l generalizing m with
660+
| nil => simp [insertListIfNewUnitₘ]
661+
| cons hd tl ih =>
662+
simp only [List.foldl_cons, insertListIfNewUnitₘ]
663+
apply ih
664+
602665
end
603666

604667
end Raw₀

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

Lines changed: 42 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -199,10 +199,52 @@ theorem filter_val [BEq α] [Hashable α] {m : Raw₀ α β} {f : (a : α) →
199199
m.val.filter f = m.filter f := by
200200
simp [Raw.filter, m.2]
201201

202+
theorem insertMany_eq [BEq α] [Hashable α] {m : Raw α β} (h : m.WF) {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] {l : ρ} :
203+
m.insertMany l = Raw₀.insertMany ⟨m, h.size_buckets_pos⟩ l := by
204+
simp [Raw.insertMany, h.size_buckets_pos]
205+
206+
theorem insertMany_val [BEq α][Hashable α] {m : Raw₀ α β} {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] {l : ρ} :
207+
m.val.insertMany l = m.insertMany l := by
208+
simp [Raw.insertMany, m.2]
209+
210+
theorem ofList_eq [BEq α] [Hashable α] {l : List ((a : α) × β a)} :
211+
Raw.ofList l = Raw₀.insertMany Raw₀.empty l := by
212+
simp only [Raw.ofList, Raw.insertMany, (Raw.WF.empty).size_buckets_pos ∅, ↓reduceDIte]
213+
congr
214+
202215
section
203216

204217
variable {β : Type v}
205218

219+
theorem Const.insertMany_eq [BEq α] [Hashable α] {m : Raw α (fun _ => β)} (h : m.WF) {ρ : Type w} [ForIn Id ρ (α × β)] {l : ρ} :
220+
Raw.Const.insertMany m l = Raw₀.Const.insertMany ⟨m, h.size_buckets_pos⟩ l := by
221+
simp [Raw.Const.insertMany, h.size_buckets_pos]
222+
223+
theorem Const.insertMany_val [BEq α][Hashable α] {m : Raw₀ α (fun _ => β)} {ρ : Type w} [ForIn Id ρ (α × β)] {l : ρ} :
224+
Raw.Const.insertMany m.val l = Raw₀.Const.insertMany m l := by
225+
simp [Raw.Const.insertMany, m.2]
226+
227+
theorem Const.ofList_eq [BEq α] [Hashable α] {l : List (α × β)} :
228+
Raw.Const.ofList l = Raw₀.Const.insertMany Raw₀.empty l := by
229+
simp only [Raw.Const.ofList, Raw.Const.insertMany, (Raw.WF.empty).size_buckets_pos ∅, ↓reduceDIte]
230+
congr
231+
232+
theorem Const.insertManyIfNewUnit_eq {ρ : Type w} [ForIn Id ρ α] [BEq α] [Hashable α]
233+
{m : Raw α (fun _ => Unit)} {l : ρ} (h : m.WF):
234+
Raw.Const.insertManyIfNewUnit m l = Raw₀.Const.insertManyIfNewUnit ⟨m, h.size_buckets_pos⟩ l := by
235+
simp [Raw.Const.insertManyIfNewUnit, h.size_buckets_pos]
236+
237+
theorem Const.insertManyIfNewUnit_val {ρ : Type w} [ForIn Id ρ α] [BEq α] [Hashable α]
238+
{m : Raw₀ α (fun _ => Unit)} {l : ρ} :
239+
Raw.Const.insertManyIfNewUnit m.val l = Raw₀.Const.insertManyIfNewUnit m l := by
240+
simp [Raw.Const.insertManyIfNewUnit, m.2]
241+
242+
theorem Const.unitOfList_eq [BEq α] [Hashable α] {l : List α} :
243+
Raw.Const.unitOfList l = Raw₀.Const.insertManyIfNewUnit Raw₀.empty l := by
244+
simp only [Raw.Const.unitOfList, Raw.Const.insertManyIfNewUnit, (Raw.WF.empty).size_buckets_pos ∅,
245+
↓reduceDIte]
246+
congr
247+
206248
theorem Const.get?_eq [BEq α] [Hashable α] {m : Raw α (fun _ => β)} (h : m.WF) {a : α} :
207249
Raw.Const.get? m a = Raw₀.Const.get? ⟨m, h.size_buckets_pos⟩ a := by
208250
simp [Raw.Const.get?, h.size_buckets_pos]

0 commit comments

Comments
 (0)